|
| 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