@@ -80,6 +80,28 @@ instance [Preorder α] : OrderHomClass (ClosureOperator α) α α where
80
80
81
81
initialize_simps_projections ClosureOperator (toFun → apply, IsClosed → isClosed)
82
82
83
+
84
+ /-- If `c` is a closure operator on `α` and `e` an order-isomorphism
85
+ between `α` and `β` then `e ∘ c ∘ e⁻¹` is a closure operator on `β`. -/
86
+ @[simps apply]
87
+ def conjBy {α β} [Preorder α] [Preorder β] (c : ClosureOperator α)
88
+ (e : α ≃o β) : ClosureOperator β where
89
+ toFun := e.conj c
90
+ IsClosed b := c.IsClosed (e.symm b)
91
+ monotone' _ _ h :=
92
+ (map_le_map_iff e).mpr <| c.monotone <| (map_le_map_iff e.symm).mpr h
93
+ le_closure' _ := e.symm_apply_le.mp (c.le_closure' _)
94
+ idempotent' _ :=
95
+ congrArg e <| Eq.trans (congrArg c (e.symm_apply_apply _)) (c.idempotent' _)
96
+ isClosed_iff := Iff.trans c.isClosed_iff e.eq_symm_apply
97
+
98
+ lemma conjBy_refl {α} [Preorder α] (c : ClosureOperator α) :
99
+ c.conjBy (OrderIso.refl α) = c := rfl
100
+
101
+ lemma conjBy_trans {α β γ} [Preorder α] [Preorder β] [Preorder γ]
102
+ (e₁ : α ≃o β) (e₂ : β ≃o γ) (c : ClosureOperator α) :
103
+ c.conjBy (e₁.trans e₂) = (c.conjBy e₁).conjBy e₂ := rfl
104
+
83
105
section PartialOrder
84
106
85
107
variable [PartialOrder α]
@@ -102,9 +124,8 @@ instance : Inhabited (ClosureOperator α) :=
102
124
variable {α} [PartialOrder α] (c : ClosureOperator α)
103
125
104
126
@[ext]
105
- theorem ext : ∀ c₁ c₂ : ClosureOperator α, (c₁ : α → α) = (c₂ : α → α) → c₁ = c₂ :=
106
- DFunLike.coe_injective
107
- #align closure_operator.ext ClosureOperator.ext
127
+ theorem ext : ∀ c₁ c₂ : ClosureOperator α, (∀ x, c₁ x = c₂ x) → c₁ = c₂ :=
128
+ DFunLike.ext
108
129
109
130
/-- Constructor for a closure operator using the weaker idempotency axiom: `f (f x) ≤ f x`. -/
110
131
@[simps]
@@ -195,6 +216,15 @@ theorem IsClosed.closure_le_iff (hy : c.IsClosed y) : c x ≤ y ↔ x ≤ y := b
195
216
196
217
lemma closure_min (hxy : x ≤ y) (hy : c.IsClosed y) : c x ≤ y := hy.closure_le_iff.2 hxy
197
218
219
+ lemma closure_isGLB (x : α) : IsGLB { y | x ≤ y ∧ c.IsClosed y } (c x) where
220
+ left _ := and_imp.mpr closure_min
221
+ right _ h := h ⟨c.le_closure x, c.isClosed_closure x⟩
222
+
223
+ theorem ext_isClosed (c₁ c₂ : ClosureOperator α)
224
+ (h : ∀ x, c₁.IsClosed x ↔ c₂.IsClosed x) : c₁ = c₂ :=
225
+ ext c₁ c₂ <| fun x => IsGLB.unique (c₁.closure_isGLB x) <|
226
+ (Set.ext (and_congr_right' <| h ·)).substr (c₂.closure_isGLB x)
227
+
198
228
/-- A closure operator is equal to the closure operator obtained by feeding `c.closed` into the
199
229
`ofPred` constructor. -/
200
230
theorem eq_ofPred_closed (c : ClosureOperator α) :
@@ -262,6 +292,11 @@ def ofCompletePred (p : α → Prop) (hsinf : ∀ s, (∀ a ∈ s, p a) → p (s
262
292
(fun a ↦ hsinf _ <| forall_mem_range.2 fun b ↦ b.2 .2 )
263
293
(fun a b hab hb ↦ iInf_le_of_le ⟨b, hab, hb⟩ le_rfl)
264
294
295
+ theorem sInf_isClosed {c : ClosureOperator α} {S : Set α}
296
+ (H : ∀ x ∈ S, c.IsClosed x) : c.IsClosed (sInf S) :=
297
+ isClosed_iff_closure_le.mpr <| le_of_le_of_eq c.monotone.map_sInf_le <|
298
+ Eq.trans (biInf_congr (c.isClosed_iff.mp <| H · ·)) sInf_eq_iInf.symm
299
+
265
300
@[simp]
266
301
theorem closure_iSup_closure (f : ι → α) : c (⨆ i, c (f i)) = c (⨆ i, f i) :=
267
302
le_antisymm (le_closure_iff.1 <| iSup_le fun i => c.monotone <| le_iSup f i) <|
@@ -279,6 +314,18 @@ end CompleteLattice
279
314
280
315
end ClosureOperator
281
316
317
+ /-- Conjugating `ClosureOperators` on `α` and on `β` by a fixed isomorphism
318
+ `e : α ≃o β` gives an equivalence `ClosureOperator α ≃ ClosureOperator β`. -/
319
+ @[simps apply symm_apply]
320
+ def OrderIso.equivClosureOperator {α β} [Preorder α] [Preorder β] (e : α ≃o β) :
321
+ ClosureOperator α ≃ ClosureOperator β where
322
+ toFun c := c.conjBy e
323
+ invFun c := c.conjBy e.symm
324
+ left_inv c := Eq.trans (c.conjBy_trans _ _).symm
325
+ <| Eq.trans (congrArg _ e.self_trans_symm) c.conjBy_refl
326
+ right_inv c := Eq.trans (c.conjBy_trans _ _).symm
327
+ <| Eq.trans (congrArg _ e.symm_trans_self) c.conjBy_refl
328
+
282
329
/-! ### Lower adjoint -/
283
330
284
331
0 commit comments