Skip to content

Commit b127017

Browse files
committed
Complete the proof
Signed-off-by: zeramorphic <zeramorphic@proton.me>
1 parent 7484104 commit b127017

File tree

3 files changed

+227
-0
lines changed

3 files changed

+227
-0
lines changed

ConNF/Model.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7,3 +7,4 @@ import ConNF.Model.Predicative
77
import ConNF.Model.FOA
88
import ConNF.Model.RaiseStrong
99
import ConNF.Model.Union
10+
import ConNF.Model.Result

ConNF/Model/Predicative.lean

Lines changed: 85 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -106,6 +106,53 @@ theorem smul_up (t₁ t₂ : TSet β) (ρ : Allowable α) :
106106
intro s
107107
simp only [mem_smul_iff, mem_up_iff, inv_smul_eq_iff]
108108

109+
theorem up_self (t : TSet β) : up hβ t t = .singleton hβ t := by
110+
refine ext hβ _ _ ?_
111+
simp only [mem_up_iff, or_self, mem_singleton_iff, implies_true]
112+
113+
theorem up_injective (t₁ t₂ t₁' t₂' : TSet β) (h : up hβ t₁ t₂ = up hβ t₁' t₂') :
114+
(t₁ = t₁' ∧ t₂ = t₂') ∨ (t₁ = t₂' ∧ t₂ = t₁') := by
115+
have h₁ : t₁ ∈[hβ] up hβ t₁ t₂
116+
· rw [mem_up_iff]
117+
exact Or.inl rfl
118+
have h₂ : t₂ ∈[hβ] up hβ t₁ t₂
119+
· rw [mem_up_iff]
120+
exact Or.inr rfl
121+
rw [h] at h₁ h₂
122+
rw [mem_up_iff] at h₁ h₂
123+
obtain (h₁ | h₁) := h₁ <;> obtain (h₂ | h₂) := h₂
124+
· refine Or.inl ⟨h₁, ?_⟩
125+
cases h₁.trans h₂.symm
126+
have : t₂' ∈[hβ] up hβ t₁' t₂'
127+
· rw [mem_up_iff]
128+
exact Or.inr rfl
129+
rw [← h, up_self, mem_singleton_iff] at this
130+
exact this.symm
131+
· refine Or.inl ⟨h₁, h₂⟩
132+
· refine Or.inr ⟨h₁, h₂⟩
133+
· refine Or.inr ⟨h₁, ?_⟩
134+
cases h₁.trans h₂.symm
135+
have : t₁' ∈[hβ] up hβ t₁' t₂'
136+
· rw [mem_up_iff]
137+
exact Or.inl rfl
138+
rw [← h, up_self, mem_singleton_iff] at this
139+
exact this.symm
140+
141+
theorem up_eq_singleton_iff (t t₁ t₂ : TSet β) :
142+
up hβ t₁ t₂ = .singleton hβ t ↔ t₁ = t ∧ t₂ = t := by
143+
constructor
144+
· intro h
145+
have h₁ : t₁ ∈[hβ] up hβ t₁ t₂
146+
· rw [mem_up_iff]
147+
exact Or.inl rfl
148+
have h₂ : t₂ ∈[hβ] up hβ t₁ t₂
149+
· rw [mem_up_iff]
150+
exact Or.inr rfl
151+
rw [h, mem_singleton_iff] at h₁ h₂
152+
exact ⟨h₁, h₂⟩
153+
· rintro ⟨rfl, rfl⟩
154+
rw [up_self]
155+
109156
@[simp]
110157
theorem TangleData.TSet.mem_op_iff (t₁ t₂ : TSet γ) (s : TSet β) :
111158
s ∈[hβ] op hβ hγ t₁ t₂ ↔ s = singleton hγ t₁ ∨ s = up hγ t₁ t₂ :=
@@ -119,6 +166,44 @@ theorem smul_op (t₁ t₂ : TSet γ) (ρ : Allowable α) :
119166
intro s
120167
simp only [mem_smul_iff, mem_op_iff, inv_smul_eq_iff, smul_singleton, smul_up]
121168

169+
theorem op_injective (t₁ t₂ t₁' t₂' : TSet γ) (h : op hβ hγ t₁ t₂ = op hβ hγ t₁' t₂') :
170+
t₁ = t₁' ∧ t₂ = t₂' := by
171+
have h₁ : .singleton hγ t₁ ∈[hβ] op hβ hγ t₁ t₂
172+
· rw [mem_op_iff]
173+
exact Or.inl rfl
174+
have h₂ : up hγ t₁ t₂ ∈[hβ] op hβ hγ t₁ t₂
175+
· rw [mem_op_iff]
176+
exact Or.inr rfl
177+
rw [h] at h₁ h₂
178+
rw [mem_op_iff] at h₁ h₂
179+
obtain (h₁ | h₁) := h₁ <;> obtain (h₂ | h₂) := h₂
180+
· cases singleton_injective' _ _ _ h₁
181+
rw [up_eq_singleton_iff] at h₂
182+
have h' : up hγ t₁ t₂' ∈[hβ] op hβ hγ t₁ t₂'
183+
· rw [mem_op_iff]
184+
exact Or.inr rfl
185+
rw [← h, mem_op_iff] at h'
186+
obtain (h' | h') := h'
187+
· rw [up_eq_singleton_iff] at h'
188+
exact ⟨rfl, h₂.2.trans h'.2.symm⟩
189+
· obtain (h' | ⟨rfl, rfl⟩) := up_injective _ _ _ _ _ h'
190+
· exact ⟨rfl, h'.2.symm⟩
191+
· exact ⟨rfl, rfl⟩
192+
· cases singleton_injective' _ _ _ h₁
193+
obtain (h' | ⟨rfl, rfl⟩) := up_injective _ _ _ _ _ h₂
194+
· exact h'
195+
· exact ⟨rfl, rfl⟩
196+
· symm at h₁
197+
rw [up_eq_singleton_iff] at h₁ h₂
198+
obtain ⟨rfl, rfl⟩ := h₁
199+
obtain ⟨_, rfl⟩ := h₂
200+
exact ⟨rfl, rfl⟩
201+
· symm at h₁
202+
rw [up_eq_singleton_iff] at h₁
203+
obtain ⟨rfl, rfl⟩ := h₁
204+
rw [up_self, up_eq_singleton_iff] at h₂
205+
exact h₂
206+
122207
theorem Symmetric.singletonImage {s : Set (TSet γ)} (h : Symmetric hγ s) :
123208
Symmetric hβ {p | ∃ a b : TSet ε, op hδ hε a b ∈ s ∧
124209
p = op hγ hδ (.singleton hε a) (.singleton hε b)} := by

ConNF/Model/Result.lean

Lines changed: 141 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,141 @@
1+
import ConNF.Model.Union
2+
3+
namespace ConNF
4+
5+
noncomputable local instance : Params := minimalParams
6+
7+
variable {α β γ δ ε ζ : Λ} (hβ : β < α) (hγ : γ < β) (hδ : δ < γ) (hε : ε < δ) (hζ : ζ < ε)
8+
9+
theorem ext (t₁ t₂ : TSet α) :
10+
(∀ s, s ∈[hβ] t₁ ↔ s ∈[hβ] t₂) → t₁ = t₂ :=
11+
TangleData.TSet.ext hβ t₁ t₂
12+
13+
theorem inter {t₁ t₂ : TSet α} :
14+
∃ t, ∀ s, s ∈[hβ] t ↔ s ∈[hβ] t₁ ∧ s ∈[hβ] t₂ := by
15+
refine ⟨t₁.inter hβ t₂, ?_⟩
16+
simp only [TangleData.TSet.inter, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq, implies_true]
17+
18+
theorem compl {t : TSet α} :
19+
∃ t', ∀ s, s ∈[hβ] t' ↔ ¬s ∈[hβ] t := by
20+
refine ⟨t.compl hβ, ?_⟩
21+
simp only [TangleData.TSet.compl, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq, implies_true]
22+
23+
theorem mem_singleton_iff (t : TSet β) (s : TSet β) :
24+
s ∈[hβ] .singleton hβ t ↔ s = t :=
25+
TangleData.TSet.mem_singleton_iff hβ t s
26+
27+
theorem mem_up_iff (t₁ t₂ : TSet β) (s : TSet β) :
28+
s ∈[hβ] .up hβ t₁ t₂ ↔ s = t₁ ∨ s = t₂ :=
29+
TangleData.TSet.mem_up_iff hβ t₁ t₂ s
30+
31+
theorem mem_op_iff (t₁ t₂ : TSet γ) (s : TSet β) :
32+
s ∈[hβ] .op hβ hγ t₁ t₂ ↔ s = .singleton hγ t₁ ∨ s = .up hγ t₁ t₂ :=
33+
TangleData.TSet.mem_op_iff hβ hγ t₁ t₂ s
34+
35+
theorem singletonImage (t : TSet β) :
36+
∃ t', ∀ a b,
37+
.op hγ hδ (.singleton hε a) (.singleton hε b) ∈[hβ] t' ↔ .op hδ hε a b ∈[hγ] t := by
38+
refine ⟨t.singletonImage hβ hγ hδ hε, ?_⟩
39+
intro a b
40+
simp only [TangleData.TSet.singletonImage, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
41+
constructor
42+
· rintro ⟨a', b', h₁, h₂⟩
43+
have := op_injective _ _ _ _ _ _ h₂
44+
cases singleton_injective' _ _ _ this.1
45+
cases singleton_injective' _ _ _ this.2
46+
exact h₁
47+
· intro h
48+
exact ⟨a, b, h, rfl⟩
49+
50+
theorem insertion2 (t : TSet γ) :
51+
∃ t', ∀ a b c, .op hγ hδ (.singleton hε (.singleton hζ a)) (.op hε hζ b c) ∈[hβ] t' ↔
52+
.op hε hζ a c ∈[hδ] t := by
53+
refine ⟨t.insertion2 hβ hγ hδ hε hζ, ?_⟩
54+
intro a b c
55+
simp only [TangleData.TSet.insertion2, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
56+
constructor
57+
· rintro ⟨a', b', c', h₁, h₂⟩
58+
have := op_injective _ _ _ _ _ _ h₂
59+
cases singleton_injective' _ _ _ (singleton_injective' _ _ _ this.1)
60+
obtain ⟨rfl, rfl⟩ := op_injective _ _ _ _ _ _ this.2
61+
exact h₁
62+
· intro h
63+
exact ⟨a, b, c, h, rfl⟩
64+
65+
theorem insertion3 (t : TSet γ) :
66+
∃ t', ∀ a b c, .op hγ hδ (.singleton hε (.singleton hζ a)) (.op hε hζ b c) ∈[hβ] t' ↔
67+
.op hε hζ a b ∈[hδ] t := by
68+
refine ⟨t.insertion3 hβ hγ hδ hε hζ, ?_⟩
69+
intro a b c
70+
simp only [TangleData.TSet.insertion3, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
71+
constructor
72+
· rintro ⟨a', b', c', h₁, h₂⟩
73+
have := op_injective _ _ _ _ _ _ h₂
74+
cases singleton_injective' _ _ _ (singleton_injective' _ _ _ this.1)
75+
obtain ⟨rfl, rfl⟩ := op_injective _ _ _ _ _ _ this.2
76+
exact h₁
77+
· intro h
78+
exact ⟨a, b, c, h, rfl⟩
79+
80+
theorem cross (t : TSet γ) :
81+
∃ t', ∀ a, a ∈[hβ] t' ↔ ∃ b c, a = .op hγ hδ b c ∧ c ∈[hδ] t := by
82+
refine ⟨t.cross hβ hγ hδ, ?_⟩
83+
intro a
84+
simp only [TangleData.TSet.cross, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
85+
constructor
86+
· rintro ⟨a', b', h, rfl⟩
87+
exact ⟨a', b', rfl, h⟩
88+
· rintro ⟨b, c, rfl, h⟩
89+
exact ⟨b, c, h, rfl⟩
90+
91+
theorem typeLower (t : TSet α) :
92+
∃ t', ∀ a, a ∈[hε] t' ↔ ∀ b, .op hγ hδ b (.singleton hε a) ∈[hβ] t := by
93+
refine ⟨.mk hε {a | ∀ b : TSet δ, .op hγ hδ b (.singleton hε a) ∈[hβ] t} ?_, ?_⟩
94+
· refine symmetric_of_image_singleton_symmetric hδ hε ?_ ?_
95+
refine symmetric_of_image_singleton_symmetric hγ hδ ?_ ?_
96+
refine symmetric_of_image_singleton_symmetric hβ hγ ?_ ?_
97+
convert t.typeLower'_supported hβ hγ hδ hε using 1
98+
aesop
99+
· simp only [TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq, implies_true]
100+
101+
theorem converse (t : TSet α) :
102+
∃ t', ∀ a b, .op hγ hδ a b ∈[hβ] t' ↔ .op hγ hδ b a ∈[hβ] t := by
103+
refine ⟨t.converse hβ hγ hδ, ?_⟩
104+
intro a b
105+
simp only [TangleData.TSet.converse, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
106+
constructor
107+
· rintro ⟨a', b', h₁, h₂⟩
108+
obtain ⟨rfl, rfl⟩ := op_injective _ _ _ _ _ _ h₂
109+
exact h₁
110+
· intro h
111+
exact ⟨b, a, h, rfl⟩
112+
113+
theorem cardinalOne :
114+
∃ t, ∀ a, a ∈[hβ] t ↔ ∃ b, ∀ c, c ∈[hγ] a ↔ c = b := by
115+
refine ⟨TangleData.TSet.cardinalOne hβ hγ, ?_⟩
116+
intro a
117+
simp only [TangleData.TSet.cardinalOne, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
118+
constructor
119+
· rintro ⟨a, rfl⟩
120+
refine ⟨a, ?_⟩
121+
intro b
122+
rw [TangleData.TSet.mem_singleton_iff]
123+
· rintro ⟨b, hb⟩
124+
refine ⟨b, ?_⟩
125+
refine ext hγ _ _ ?_
126+
intro c
127+
rw [hb, TangleData.TSet.mem_singleton_iff]
128+
129+
theorem subset :
130+
∃ t, ∀ a b, .op hγ hδ a b ∈[hβ] t ↔ ∀ c, c ∈[hε] a → c ∈[hε] b := by
131+
refine ⟨TangleData.TSet.subset hβ hγ hδ hε, ?_⟩
132+
intro a b
133+
simp only [TangleData.TSet.subset, TangleData.TSet.mem_mk_iff, Set.mem_setOf_eq]
134+
constructor
135+
· rintro ⟨a', b', h₁, h₂⟩
136+
obtain ⟨rfl, rfl⟩ := op_injective _ _ _ _ _ _ h₂
137+
exact h₁
138+
· intro h
139+
exact ⟨a, b, h, rfl⟩
140+
141+
end ConNF

0 commit comments

Comments
 (0)