Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 641a9d3

Browse files
committed
feat(model_theory/basic): Substructures (#7762)
Defines substructures of first-order structures
1 parent 456a6d5 commit 641a9d3

File tree

1 file changed

+262
-0
lines changed

1 file changed

+262
-0
lines changed

src/model_theory/basic.lean

Lines changed: 262 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Aaron Anderson, Jesse Michael Han, Floris van Doorn
55
-/
66
import data.nat.basic
7+
import data.set_like
8+
import data.set.lattice
9+
import order.closure
710

811
/-!
912
# Basics on First-Order Structures
@@ -114,6 +117,13 @@ protected structure equiv extends M ≃ N :=
114117
localized "notation A ` ≃[`:25 L `] ` B := L.equiv A B" in first_order
115118

116119
variables {L M N} {P : Type*} [L.Structure P] {Q : Type*} [L.Structure Q]
120+
121+
instance : has_coe_t L.const M :=
122+
⟨λ c, fun_map c fin.elim0⟩
123+
124+
lemma fun_map_eq_coe_const {c : L.const} {x : fin 0 → M} :
125+
fun_map c x = c := congr rfl (funext fin.elim0)
126+
117127
namespace hom
118128

119129
@[simps] instance has_coe_to_fun : has_coe_to_fun (M →[L] N) :=
@@ -134,6 +144,9 @@ lemma ext_iff {f g : M →[L] N} : f = g ↔ ∀ x, f x = g x :=
134144
@[simp] lemma map_fun (φ : M →[L] N) {n : ℕ} (f : L.functions n) (x : fin n → M) :
135145
φ (fun_map f x) = fun_map f (φ ∘ x) := φ.map_fun' f x
136146

147+
@[simp] lemma map_const (φ : M →[L] N) (c : L.const) : φ c = c :=
148+
(φ.map_fun c fin.elim0).trans (congr rfl (funext fin.elim0))
149+
137150
@[simp] lemma map_rel (φ : M →[L] N) {n : ℕ} (r : L.relations n) (x : fin n → M) :
138151
rel_map r x → rel_map r (φ ∘ x) := φ.map_rel' r x
139152

@@ -171,6 +184,9 @@ namespace embedding
171184
@[simp] lemma map_fun (φ : M ↪[L] N) {n : ℕ} (f : L.functions n) (x : fin n → M) :
172185
φ (fun_map f x) = fun_map f (φ ∘ x) := φ.map_fun' f x
173186

187+
@[simp] lemma map_const (φ : M ↪[L] N) (c : L.const) : φ c = c :=
188+
(φ.map_fun c fin.elim0).trans (congr rfl (funext fin.elim0))
189+
174190
@[simp] lemma map_rel (φ : M ↪[L] N) {n : ℕ} (r : L.relations n) (x : fin n → M) :
175191
rel_map r (φ ∘ x) ↔ rel_map r x := φ.map_rel' r x
176192

@@ -249,6 +265,9 @@ namespace equiv
249265
@[simp] lemma map_fun (φ : M ≃[L] N) {n : ℕ} (f : L.functions n) (x : fin n → M) :
250266
φ (fun_map f x) = fun_map f (φ ∘ x) := φ.map_fun' f x
251267

268+
@[simp] lemma map_const (φ : M ≃[L] N) (c : L.const) : φ c = c :=
269+
(φ.map_fun c fin.elim0).trans (congr rfl (funext fin.elim0))
270+
252271
@[simp] lemma map_rel (φ : M ≃[L] N) {n : ℕ} (r : L.relations n) (x : fin n → M) :
253272
rel_map r (φ ∘ x) ↔ rel_map r x := φ.map_rel' r x
254273

@@ -313,5 +332,248 @@ lemma comp_assoc (f : M ≃[L] N) (g : N ≃[L] P) (h : P ≃[L] Q) :
313332

314333
end equiv
315334

335+
section closed_under
336+
337+
open set
338+
339+
variables {n : ℕ} (f : L.functions n) (s : set M)
340+
341+
/-- Indicates that a set in a given structure is a closed under a function symbol. -/
342+
def closed_under : Prop :=
343+
∀ (x : fin n → M), (∀ i : fin n, x i ∈ s) → fun_map f x ∈ s
344+
345+
variables {f} {s} {t : set M}
346+
347+
namespace closed_under
348+
349+
lemma inter (hs : closed_under f s) (ht : closed_under f t) :
350+
closed_under f (s ∩ t) :=
351+
λ x h, mem_inter (hs x (λ i, mem_of_mem_inter_left (h i)))
352+
(ht x (λ i, mem_of_mem_inter_right (h i)))
353+
354+
lemma inf (hs : closed_under f s) (ht : closed_under f t) :
355+
closed_under f (s ⊓ t) := hs.inter ht
356+
357+
variables {S : set (set M)}
358+
359+
lemma Inf (hS : ∀ s, s ∈ S → closed_under f s) : closed_under f (Inf S) :=
360+
λ x h s hs, hS s hs x (λ i, h i s hs)
361+
362+
end closed_under
363+
end closed_under
364+
365+
variables (L) (M)
366+
367+
/-- A substructure of a structure `M` is a set closed under application of function symbols. -/
368+
structure substructure :=
369+
(carrier : set M)
370+
(fun_mem : ∀{n}, ∀ (f : L.functions n), closed_under f carrier)
371+
372+
variables {L} {M}
373+
374+
namespace substructure
375+
376+
instance : set_like (L.substructure M) M :=
377+
⟨substructure.carrier, λ p q h, by cases p; cases q; congr'⟩
378+
379+
/-- See Note [custom simps projection] -/
380+
def simps.coe (S : L.substructure M) : set M := S
381+
initialize_simps_projections substructure (carrier → coe)
382+
383+
@[simp]
384+
lemma mem_carrier {s : L.substructure M} {x : M} : x ∈ s.carrier ↔ x ∈ s := iff.rfl
385+
386+
/-- Two substructures are equal if they have the same elements. -/
387+
@[ext]
388+
theorem ext {S T : L.substructure M}
389+
(h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := set_like.ext h
390+
391+
/-- Copy a substructure replacing `carrier` with a set that is equal to it. -/
392+
protected def copy (S : L.substructure M) (s : set M) (hs : s = S) : L.substructure M :=
393+
{ carrier := s,
394+
fun_mem := λ n f, hs.symm ▸ (S.fun_mem f) }
395+
396+
variable {S : L.substructure M}
397+
398+
@[simp] lemma coe_copy {s : set M} (hs : s = S) :
399+
(S.copy s hs : set M) = s := rfl
400+
401+
lemma copy_eq {s : set M} (hs : s = S) : S.copy s hs = S :=
402+
set_like.coe_injective hs
403+
404+
lemma const_mem {c : L.const} : ↑c ∈ S :=
405+
mem_carrier.2 (S.fun_mem c _ fin.elim0)
406+
407+
/-- The substructure `M` of the structure `M`. -/
408+
instance : has_top (L.substructure M) :=
409+
⟨{ carrier := set.univ,
410+
fun_mem := λ n f x h, set.mem_univ _ }⟩
411+
412+
instance : inhabited (L.substructure M) := ⟨⊤⟩
413+
414+
@[simp] lemma mem_top (x : M) : x ∈ (⊤ : L.substructure M) := set.mem_univ x
415+
416+
@[simp] lemma coe_top : ((⊤ : L.substructure M) : set M) = set.univ := rfl
417+
418+
/-- The inf of two substructures is their intersection. -/
419+
instance : has_inf (L.substructure M) :=
420+
⟨λ S₁ S₂,
421+
{ carrier := S₁ ∩ S₂,
422+
fun_mem := λ n f, (S₁.fun_mem f).inf (S₂.fun_mem f) }⟩
423+
424+
@[simp]
425+
lemma coe_inf (p p' : L.substructure M) : ((p ⊓ p' : L.substructure M) : set M) = p ∩ p' := rfl
426+
427+
@[simp]
428+
lemma mem_inf {p p' : L.substructure M} {x : M} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' := iff.rfl
429+
430+
instance : has_Inf (L.substructure M) :=
431+
⟨λ s, { carrier := ⋂ t ∈ s, ↑t,
432+
fun_mem := λ n f, closed_under.Inf begin
433+
rintro _ ⟨t, rfl⟩,
434+
simp only,
435+
by_cases h : t ∈ s,
436+
{ rw [set.Inter_pos h],
437+
exact t.fun_mem f },
438+
{ rw [set.Inter_neg h],
439+
exact λ _ _, set.mem_univ _ }
440+
end }⟩
441+
442+
@[simp, norm_cast]
443+
lemma coe_Inf (S : set (L.substructure M)) :
444+
((Inf S : L.substructure M) : set M) = ⋂ s ∈ S, ↑s := rfl
445+
446+
lemma mem_Inf {S : set (L.substructure M)} {x : M} : x ∈ Inf S ↔ ∀ p ∈ S, x ∈ p :=
447+
set.mem_bInter_iff
448+
449+
lemma mem_infi {ι : Sort*} {S : ι → L.substructure M} {x : M} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i :=
450+
by simp only [infi, mem_Inf, set.forall_range_iff]
451+
452+
@[simp, norm_cast]
453+
lemma coe_infi {ι : Sort*} {S : ι → L.substructure M} : (↑(⨅ i, S i) : set M) = ⋂ i, S i :=
454+
by simp only [infi, coe_Inf, set.bInter_range]
455+
456+
/-- Substructures of a structure form a complete lattice. -/
457+
instance : complete_lattice (L.substructure M) :=
458+
{ le := (≤),
459+
lt := (<),
460+
top := (⊤),
461+
le_top := λ S x hx, mem_top x,
462+
inf := (⊓),
463+
Inf := has_Inf.Inf,
464+
le_inf := λ a b c ha hb x hx, ⟨ha hx, hb hx⟩,
465+
inf_le_left := λ a b x, and.left,
466+
inf_le_right := λ a b x, and.right,
467+
.. complete_lattice_of_Inf (L.substructure M) $ λ s,
468+
is_glb.of_image (λ S T,
469+
show (S : set M) ≤ T ↔ S ≤ T, from set_like.coe_subset_coe) is_glb_binfi }
470+
471+
variable (L)
472+
473+
/-- The `L.substructure` generated by a set. -/
474+
def closure : lower_adjoint (coe : L.substructure M → set M) := ⟨λ s, Inf {S | s ⊆ S},
475+
λ s S, ⟨set.subset.trans (λ x hx, mem_Inf.2 $ λ S hS, hS hx), λ h, Inf_le h⟩⟩
476+
477+
variables {L} {s : set M}
478+
479+
lemma mem_closure {x : M} : x ∈ closure L s ↔ ∀ S : L.substructure M, s ⊆ S → x ∈ S :=
480+
mem_Inf
481+
482+
/-- The substructure generated by a set includes the set. -/
483+
@[simp]
484+
lemma subset_closure : s ⊆ closure L s := (closure L).le_closure s
485+
486+
@[simp]
487+
lemma closed (S : L.substructure M) : (closure L).closed (S : set M) :=
488+
congr rfl ((closure L).eq_of_le set.subset.rfl (λ x xS, mem_closure.2 (λ T hT, hT xS)))
489+
490+
open set
491+
492+
/-- A substructure `S` includes `closure L s` if and only if it includes `s`. -/
493+
@[simp]
494+
lemma closure_le : closure L s ≤ S ↔ s ⊆ S := (closure L).closure_le_closed_iff_le s S.closed
495+
496+
/-- Substructure closure of a set is monotone in its argument: if `s ⊆ t`,
497+
then `closure L s ≤ closure L t`. -/
498+
lemma closure_mono ⦃s t : set M⦄ (h : s ⊆ t) : closure L s ≤ closure L t :=
499+
(closure L).monotone h
500+
501+
lemma closure_eq_of_le (h₁ : s ⊆ S) (h₂ : S ≤ closure L s) : closure L s = S :=
502+
(closure L).eq_of_le h₁ h₂
503+
504+
variable (S)
505+
506+
/-- An induction principle for closure membership. If `p` holds for all elements of `s`, and
507+
is preserved under function symbols, then `p` holds for all elements of the closure of `s`. -/
508+
@[elab_as_eliminator] lemma closure_induction {p : M → Prop} {x} (h : x ∈ closure L s)
509+
(Hs : ∀ x ∈ s, p x)
510+
(Hfun : ∀ {n : ℕ} (f : L.functions n), closed_under f (set_of p)) : p x :=
511+
(@closure_le L M _ ⟨set_of p, λ n, Hfun⟩ _).2 Hs h
512+
513+
/-- If `s` is a dense set in a structure `M`, `substructure.closure L s = ⊤`, then in order to prove
514+
that some predicate `p` holds for all `x : M` it suffices to verify `p x` for `x ∈ s`, and verify
515+
that `p` is preserved under function symbols. -/
516+
@[elab_as_eliminator] lemma dense_induction {p : M → Prop} (x : M) {s : set M}
517+
(hs : closure L s = ⊤) (Hs : ∀ x ∈ s, p x)
518+
(Hfun : ∀ {n : ℕ} (f : L.functions n), closed_under f (set_of p)) : p x :=
519+
have ∀ x ∈ closure L s, p x, from λ x hx, closure_induction hx Hs (λ n, Hfun),
520+
by simpa [hs] using this x
521+
522+
variables (L) (M)
523+
524+
/-- `closure` forms a Galois insertion with the coercion to set. -/
525+
protected def gi : galois_insertion (@closure L M _) coe :=
526+
{ choice := λ s _, closure L s,
527+
gc := λ s t, closure_le,
528+
le_l_u := λ s, subset_closure,
529+
choice_eq := λ s h, rfl }
530+
531+
variables {L} {M}
532+
533+
/-- Closure of a substructure `S` equals `S`. -/
534+
@[simp] lemma closure_eq : closure L (S : set M) = S := (substructure.gi L M).l_u_eq S
535+
536+
@[simp] lemma closure_empty : closure L (∅ : set M) = ⊥ :=
537+
(substructure.gi L M).gc.l_bot
538+
539+
@[simp] lemma closure_univ : closure L (univ : set M) = ⊤ :=
540+
@coe_top L M _ ▸ closure_eq ⊤
541+
542+
lemma closure_union (s t : set M) : closure L (s ∪ t) = closure L s ⊔ closure L t :=
543+
(substructure.gi L M).gc.l_sup
544+
545+
lemma closure_Union {ι} (s : ι → set M) : closure L (⋃ i, s i) = ⨆ i, closure L (s i) :=
546+
(substructure.gi L M).gc.l_supr
547+
548+
end substructure
549+
550+
namespace hom
551+
552+
open substructure
553+
554+
/-- The substructure of elements `x : M` such that `f x = g x` -/
555+
def eq_locus (f g : M →[L] N) : substructure L M :=
556+
{ carrier := {x : M | f x = g x},
557+
fun_mem := λ n fn x hx, by {
558+
have h : f ∘ x = g ∘ x := by { ext, repeat {rw function.comp_apply}, apply hx, },
559+
simp [h], } }
560+
561+
/-- If two `L.hom`s are equal on a set, then they are equal on its substructure closure. -/
562+
lemma eq_on_closure {f g : M →[L] N} {s : set M} (h : set.eq_on f g s) :
563+
set.eq_on f g (closure L s) :=
564+
show closure L s ≤ f.eq_locus g, from closure_le.2 h
565+
566+
lemma eq_of_eq_on_top {f g : M →[L] N} (h : set.eq_on f g (⊤ : substructure L M)) :
567+
f = g :=
568+
ext $ λ x, h trivial
569+
570+
variable {s : set M}
571+
572+
lemma eq_of_eq_on_dense (hs : closure L s = ⊤) {f g : M →[L] N} (h : s.eq_on f g) :
573+
f = g :=
574+
eq_of_eq_on_top $ hs ▸ eq_on_closure h
575+
576+
end hom
577+
316578
end language
317579
end first_order

0 commit comments

Comments
 (0)