@@ -842,6 +842,49 @@ lemma map_comap_le (f : M →ₗ[R] M₂) (q : submodule R M₂) : map f (comap
842
842
lemma le_comap_map (f : M →ₗ[R] M₂) (p : submodule R M) : p ≤ comap f (map f p) :=
843
843
(gc_map_comap f).le_u_l _
844
844
845
+ section galois_insertion
846
+ variables {f : M →ₗ[R] M₂} (hf : surjective f)
847
+ include hf
848
+
849
+ /-- `map f` and `comap f` form a `galois_insertion` when `f` is surjective. -/
850
+ def gi_map_comap : galois_insertion (map f) (comap f) :=
851
+ (gc_map_comap f).to_galois_insertion
852
+ (λ S x hx, begin
853
+ rcases hf x with ⟨y, rfl⟩,
854
+ simp only [mem_map, mem_comap],
855
+ exact ⟨y, hx, rfl⟩
856
+ end )
857
+
858
+ lemma map_comap_eq_of_surjective (p : submodule R M₂) : (p.comap f).map f = p :=
859
+ (gi_map_comap hf).l_u_eq _
860
+
861
+ lemma map_surjective_of_surjective : function.surjective (map f) :=
862
+ (gi_map_comap hf).l_surjective
863
+
864
+ lemma comap_injective_of_surjective : function.injective (comap f) :=
865
+ (gi_map_comap hf).u_injective
866
+
867
+ lemma map_sup_comap_of_surjective (p q : submodule R M₂) :
868
+ (p.comap f ⊔ q.comap f).map f = p ⊔ q :=
869
+ (gi_map_comap hf).l_sup_u _ _
870
+
871
+ lemma map_supr_comap_of_sujective (S : ι → submodule R M₂) : (⨆ i, (S i).comap f).map f = supr S :=
872
+ (gi_map_comap hf).l_supr_u _
873
+
874
+ lemma map_inf_comap_of_surjective (p q : submodule R M₂) : (p.comap f ⊓ q.comap f).map f = p ⊓ q :=
875
+ (gi_map_comap hf).l_inf_u _ _
876
+
877
+ lemma map_infi_comap_of_surjective (S : ι → submodule R M₂) : (⨅ i, (S i).comap f).map f = infi S :=
878
+ (gi_map_comap hf).l_infi_u _
879
+
880
+ lemma comap_le_comap_iff_of_surjective (p q : submodule R M₂) : p.comap f ≤ q.comap f ↔ p ≤ q :=
881
+ (gi_map_comap hf).u_le_u_iff
882
+
883
+ lemma comap_strict_mono_of_surjective : strict_mono (comap f) :=
884
+ (gi_map_comap hf).strict_mono_u
885
+
886
+ end galois_insertion
887
+
845
888
section galois_coinsertion
846
889
variables {f : M →ₗ[R] M₂} (hf : injective f)
847
890
include hf
0 commit comments