Rename smulImage_subset to smulImage_mono

laurent-lost-commits
Shad Amethyst 11 months ago
parent 03cec8913a
commit 18912139ec

@ -134,7 +134,7 @@ lemma proposition_1_1_1 [h_lm : LocallyMoving G α] [T2Space α] (f g : G) (supp
_ ⊆ W (k •'' Support α f₁) := Set.union_subset_union_left _ supp_f₁_subset_W
_ ⊆ W (k •'' W) := by
apply Set.union_subset_union_right
exact (smulImage_subset k supp_f₁_subset_W)
exact (smulImage_mono k supp_f₁_subset_W)
_ = W (f₂ •'' W) := by rw [<-smulImage_eq_of_smul_eq h₂]
_ ⊆ V (f₂ •'' W) := Set.union_subset_union_left _ W_in_V
_ ⊆ V V := by
@ -217,7 +217,7 @@ by
have V_disjoint_smulImage: Disjoint V (f •'' V) := by
apply Set.disjoint_of_subset
· exact Set.inter_subset_left _ _
· apply smulImage_subset
· apply smulImage_mono
exact Set.inter_subset_left _ _
· exact disjoint_img_V₀
@ -253,7 +253,7 @@ by
exact rigidStabilizer_support.mp h_in_ristV
_ ⊆ V (f₂ •'' V) := by
apply Set.union_subset_union_right
apply smulImage_subset
apply smulImage_mono
exact rigidStabilizer_support.mp h_in_ristV
have support_h' : Support α h' ⊆ (i : Fin 2 × Fin 2), (f₁^(i.1.val) * f₂^(i.2.val)) •'' V := by
rw [rewrite_Union]
@ -266,7 +266,7 @@ by
exact support_f₂_h
_ ⊆ V (f₂ •'' V) (f₁ •'' V (f₂ •'' V)) := by
apply Set.union_subset_union_right
apply smulImage_subset f₁
apply smulImage_mono f₁
exact support_f₂_h
-- Since h' is nontrivial, it has at least one point p in its support
@ -367,9 +367,9 @@ by
have disj : ∀ (i j : Fin n), i ≠ j → Disjoint (g ^ (i : ) •'' V) (g ^ (j : ) •'' V) := by
intro i j i_ne_j
apply Set.disjoint_of_subset
· apply smulImage_subset
· apply smulImage_mono
apply Set.inter_subset_right
· apply smulImage_subset
· apply smulImage_mono
apply Set.inter_subset_right
exact disj' i j i_ne_j

@ -291,11 +291,7 @@ by
let U := ⋂(p : ι₂), smul_inj_nbhd p.prop f_smul_inj
use U
-- The notations provided afterwards tend to be quite ugly because we used `Exists.choose`,
-- but the idea is that this all boils down to applying `Exists.choose_spec`, except in the disjointness case,
-- where we transform `Disjoint (f i •'' U) (f j •'' U)` into `Disjoint U ((f i)⁻¹ * f j •'' U)`
-- and transform both instances of `U` into `N`, the neighborhood of the chosen `(i, j) ∈ ι₂`
repeat' constructor
repeat' apply And.intro
· apply isOpen_iInter_of_finite
intro ⟨⟨i, j⟩, i_ne_j⟩
apply smul_inj_nbhd_open
@ -303,6 +299,8 @@ by
intro ⟨⟨i, j⟩, i_ne_j⟩
apply smul_inj_nbhd_mem
· intro i j i_ne_j
-- We transform `Disjoint (f i •'' U) (f j •'' U)` into `Disjoint N ((f i)⁻¹ * f j •'' N)`
let N := smul_inj_nbhd i_ne_j f_smul_inj
have U_subset_N : U ⊆ N := Set.iInter_subset
(fun (⟨⟨i, j⟩, i_ne_j⟩ : ι₂) => (smul_inj_nbhd i_ne_j f_smul_inj))
@ -311,7 +309,7 @@ by
rw [disjoint_comm, smulImage_disjoint_mul]
apply Set.disjoint_of_subset U_subset_N
· apply smulImage_subset
· apply smulImage_mono
exact U_subset_N
· exact smul_inj_nbhd_disjoint i_ne_j f_smul_inj

@ -116,15 +116,11 @@ theorem smulImage_inv (g: G) (U V : Set α) : g •'' U = V ↔ U = g⁻¹ •''
apply SmulImage.inv_congr at h
exact h
-- TODO: rename to smulImage_mono
theorem smulImage_subset (g : G) {U V : Set α} : U ⊆ V → g •'' U ⊆ g •'' V :=
by
theorem smulImage_mono (g : G) {U V : Set α} : U ⊆ V → g •'' U ⊆ g •'' V := by
intro h1 x
rw [Rubin.mem_smulImage, Rubin.mem_smulImage]
exact fun h2 => h1 h2
#align smul''_subset Rubin.smulImage_subset
def smulImage_mono (g : G) {U V : Set α} : U ⊆ V → g •'' U ⊆ g •'' V := smulImage_subset g
#align smul''_subset Rubin.smulImage_mono
theorem smulImage_union (g : G) {U V : Set α} : g •'' U V = (g •'' U) (g •'' V) :=
by
@ -218,12 +214,12 @@ theorem smulImage_subset_inv {G α : Type _} [Group G] [MulAction G α]
by
constructor
· intro h
apply smulImage_subset f⁻¹ at h
apply smulImage_mono f⁻¹ at h
rw [smulImage_mul] at h
rw [mul_left_inv, one_smulImage] at h
exact h
· intro h
apply smulImage_subset f at h
apply smulImage_mono f at h
rw [smulImage_mul] at h
rw [mul_right_inv, one_smulImage] at h
exact h
@ -265,11 +261,10 @@ by
theorem smulImage_disjoint_subset {G α : Type _} [Group G] [MulAction G α]
{f g : G} {U V : Set α} (h_sub: U ⊆ V):
Disjoint (f •'' V) (g •'' V) → Disjoint (f •'' U) (g •'' U) :=
by
apply Set.disjoint_of_subset (smulImage_subset _ h_sub) (smulImage_subset _ h_sub)
Set.disjoint_of_subset (smulImage_mono _ h_sub) (smulImage_mono _ h_sub)
-- States that if `g^i •'' V` and `g^j •'' V` are disjoint for any `i ≠ j` and `x ∈ V`
-- then `g^i • x` will always lie outside of `V`.
-- then `g^i • x` will always lie outside of `V` if `x ∈ V`.
lemma smulImage_distinct_of_disjoint_pow {G α : Type _} [Group G] [MulAction G α] {g : G} {V : Set α} {n : }
(n_pos : 0 < n)
(h_disj : ∀ (i j : Fin n), i ≠ j → Disjoint (g ^ (i : ) •'' V) (g ^ (j : ) •'' V)) :

@ -76,7 +76,7 @@ theorem fixed_smulImage_in_support (g : G) {U : Set α} :
theorem smulImage_subset_in_support (g : G) (U V : Set α) :
U ⊆ V → Support α g ⊆ V → g •'' U ⊆ V := fun U_in_V support_in_V =>
Rubin.fixed_smulImage_in_support g support_in_V ▸
Rubin.smulImage_subset g U_in_V
smulImage_mono g U_in_V
#align moves_subset_within_support Rubin.smulImage_subset_in_support
theorem support_mul (g h : G) (α : Type _) [MulAction G α] :
@ -155,10 +155,11 @@ theorem disjoint_support_comm (f g : G) {U : Set α} :
Support α f ⊆ U → Disjoint U (g •'' U) → ∀ x ∈ U, ⁅f, g⁆ • x = f • x :=
by
intro support_in_U disjoint_U x x_in_U
have support_conj : Support α (g * f⁻¹ * g⁻¹) ⊆ g •'' U :=
((Rubin.support_conjugate α f⁻¹ g).trans
(Rubin.SmulImage.congr g (Rubin.support_inv α f))).symm ▸
Rubin.smulImage_subset g support_in_U
have support_conj : Support α (g * f⁻¹ * g⁻¹) ⊆ g •'' U := by
rw [support_conjugate]
apply smulImage_mono
rw [support_inv]
exact support_in_U
have goal :=
(congr_arg (f • ·)
(Rubin.fixed_of_disjoint x_in_U

@ -13,7 +13,6 @@ namespace Rubin
section Continuity
-- TODO: don't have this extend MulAction and move this somewhere else
class ContinuousMulAction (G α : Type _) [Group G] [TopologicalSpace α] extends
MulAction G α where
continuous : ∀ g : G, Continuous (fun x: α => g • x)

Loading…
Cancel
Save