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

Commit c0be8dc

Browse files
committed
feat(model_theory/basic): define quotient structures (#11747)
Defines prestructures and quotient structures
1 parent 792d3e5 commit c0be8dc

File tree

1 file changed

+43
-0
lines changed

1 file changed

+43
-0
lines changed

src/model_theory/basic.lean

Lines changed: 43 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1077,5 +1077,48 @@ instance : boolean_algebra (L.definable_set M α) :=
10771077
end definable_set
10781078
end definability
10791079

1080+
section quotients
1081+
1082+
variables (L) {M' : Type*}
1083+
1084+
/-- A prestructure is a first-order structure with a `setoid` equivalence relation on it,
1085+
such that quotienting by that equivalence relation is still a structure. -/
1086+
class prestructure (s : setoid M') :=
1087+
(to_structure : L.Structure M')
1088+
(fun_equiv : ∀{n} {f : L.functions n} (x y : fin n → M'),
1089+
x ≈ y → fun_map f x ≈ fun_map f y)
1090+
(rel_equiv : ∀{n} {r : L.relations n} (x y : fin n → M') (h : x ≈ y),
1091+
(rel_map r x = rel_map r y))
1092+
1093+
variables {L} {M'} {s : setoid M'} [ps : L.prestructure s]
1094+
1095+
instance quotient_structure :
1096+
L.Structure (quotient s) :=
1097+
{ fun_map := λ n f x, quotient.map (@fun_map L M' ps.to_structure n f) prestructure.fun_equiv
1098+
(quotient.fin_choice x),
1099+
rel_map := λ n r x, quotient.lift (@rel_map L M' ps.to_structure n r) prestructure.rel_equiv
1100+
(quotient.fin_choice x) }
1101+
1102+
variables [s]
1103+
include s
1104+
1105+
lemma fun_map_quotient_mk {n : ℕ} (f : L.functions n) (x : fin n → M') :
1106+
fun_map f (λ i, ⟦x i⟧) = ⟦@fun_map _ _ ps.to_structure _ f x⟧ :=
1107+
begin
1108+
change quotient.map (@fun_map L M' ps.to_structure n f) prestructure.fun_equiv
1109+
(quotient.fin_choice _) = _,
1110+
rw [quotient.fin_choice_eq, quotient.map_mk],
1111+
end
1112+
1113+
lemma realize_term_quotient_mk {β : Type*} (x : β → M') (t : L.term β) :
1114+
realize_term (λ i, ⟦x i⟧) t = ⟦@realize_term _ _ ps.to_structure _ x t⟧ :=
1115+
begin
1116+
induction t with a1 a2 a3 a4 ih a6 a7 a8 a9 a0,
1117+
{ refl },
1118+
simp only [ih, fun_map_quotient_mk, realize_term],
1119+
end
1120+
1121+
end quotients
1122+
10801123
end language
10811124
end first_order

0 commit comments

Comments
 (0)