diff --git a/Arithmetization/Basic/IOpen.lean b/Arithmetization/Basic/IOpen.lean index d3c9fe7..596d03e 100644 --- a/Arithmetization/Basic/IOpen.lean +++ b/Arithmetization/Basic/IOpen.lean @@ -846,6 +846,8 @@ lemma nat_cast_pair (n m : ℕ) : (⟪n, m⟫ : ℕ) = ⟪(n : V), (m : V)⟫ := lemma nat_pair_eq (m n : ℕ) : ⟪n, m⟫ = Nat.pair n m := by simp [Arith.pair, Nat.pair]; congr +lemma pair_coe_eq_coe_pair (m n : ℕ) : ⟪n, m⟫ = (Nat.pair n m : V) := by simp [nat_cast_pair, nat_pair_eq] + end end LO.Arith diff --git a/Arithmetization/Definability/Boldface.lean b/Arithmetization/Definability/Boldface.lean index cfccc89..ac79f3f 100644 --- a/Arithmetization/Definability/Boldface.lean +++ b/Arithmetization/Definability/Boldface.lean @@ -66,6 +66,9 @@ abbrev DefinedFunction {k} (f : (Fin k → V) → V) (p : ℌ.Semisentence (k + variable (ℌ) +abbrev DefinedFunction₀ (c : V) (p : ℌ.Semisentence 1) : Prop := + DefinedFunction (fun _ => c) p + abbrev DefinedFunction₁ (f : V → V) (p : ℌ.Semisentence 2) : Prop := DefinedFunction (fun v => f (v 0)) p @@ -95,6 +98,8 @@ abbrev BoldfaceRel₆ (P : V → V → V → V → V → V → Prop) : Prop := abbrev BoldfaceFunction (f : (Fin k → V) → V) : Prop := ℌ.Boldface (k := k + 1) (fun v ↦ v 0 = f (v ·.succ)) +abbrev BoldfaceFunction₀ (c : V) : Prop := ℌ.BoldfaceFunction (k := 0) (fun _ ↦ c) + abbrev BoldfaceFunction₁ (f : V → V) : Prop := ℌ.BoldfaceFunction (k := 1) (fun v ↦ f (v 0)) abbrev BoldfaceFunction₂ (f : V → V → V) : Prop := ℌ.BoldfaceFunction (k := 2) (fun v ↦ f (v 0) (v 1)) @@ -115,6 +120,8 @@ notation Γ "-Relation₃ " P " via " p => DefinedRel₃ Γ P p notation Γ "-Relation₄ " P " via " p => DefinedRel₄ Γ P p +notation Γ "-Function₀ " c " via " p => DefinedFunction₀ Γ c p + notation Γ "-Function₁ " f " via " p => DefinedFunction₁ Γ f p notation Γ "-Function₂ " f " via " p => DefinedFunction₂ Γ f p @@ -389,7 +396,7 @@ lemma of_sigma_of_pi (hσ : 𝚺-[m].Boldface P) (hπ : 𝚷-[m].Boldface P) : lemma of_zero (h : Γ'-[0].Boldface P) : ℌ.Boldface P := by rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable₀ -lemma of_deltaOne (h : 𝚫₁.Boldface P) (Γ m) : Γ-[m + 1].Boldface P := by +lemma of_deltaOne (h : 𝚫₁.Boldface P) {Γ m} : Γ-[m + 1].Boldface P := by rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable_deltaOne instance [𝚺₀.Boldface P] (ℌ : HierarchySymbol) : ℌ.Boldface P := Boldface.of_zero (Γ' := 𝚺) (ℌ := ℌ) inferInstance @@ -683,7 +690,7 @@ instance {k} {f : (Fin k → V) → V} [h : 𝚺-[m].BoldfaceFunction f] : 𝚫- instance {k} {f : (Fin k → V) → V} [𝚺₀.BoldfaceFunction f] : ℌ.BoldfaceFunction f := inferInstance lemma of_sigmaOne {k} {f : (Fin k → V) → V} - (h : 𝚺₁.BoldfaceFunction f) (Γ m) : Γ-[m + 1].BoldfaceFunction f := Boldface.of_deltaOne (graph_delta h) Γ m + (h : 𝚺₁.BoldfaceFunction f) {Γ m} : Γ-[m + 1].BoldfaceFunction f := Boldface.of_deltaOne (graph_delta h) @[simp] lemma var {k} (i : Fin k) : ℌ.BoldfaceFunction (fun v : Fin k → V ↦ v i) := .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!#i.succ” (by simp), by intro _; simp⟩ diff --git a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean index 9b18529..f8fade1 100644 --- a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean +++ b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean @@ -128,9 +128,7 @@ lemma termSet_defined : 𝚺₁.DefinedFunction (fun v ↦ c.limSeq (v ·.succ) instance limSeq_definable : 𝚺₁.BoldfaceFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) := c.termSet_defined.to_definable -@[simp, definability] instance limSeq_definable' (Γ) : - Γ-[m + 1].BoldfaceFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) := - .of_sigmaOne c.limSeq_definable _ _ +@[simp, definability] instance limSeq_definable' (Γ) : Γ-[m + 1].BoldfaceFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) := c.limSeq_definable.of_sigmaOne lemma mem_limSeq_succ_iff {x s : V} : x ∈ c.limSeq v (s + 1) ↔ x ≤ s ∧ c.Φ v {z | z ∈ c.limSeq v s} x := by simp [limSeq_succ, mem_succ_iff] diff --git a/Arithmetization/ISigmaOne/HFS/Seq.lean b/Arithmetization/ISigmaOne/HFS/Seq.lean index b2dbded..a2722ea 100644 --- a/Arithmetization/ISigmaOne/HFS/Seq.lean +++ b/Arithmetization/ISigmaOne/HFS/Seq.lean @@ -376,7 +376,7 @@ lemma mkSeq₂_defined : 𝚺₁-Function₂ (fun x y : V ↦ !⟦x, y⟧) via m instance mkSeq₂_definable : 𝚺₁-Function₂ (fun x y : V ↦ !⟦x, y⟧) := mkSeq₂_defined.to_definable -instance mkSeq₂_definable' (Γ m) : Γ-[m + 1]-Function₂ (fun x y : V ↦ !⟦x, y⟧) := .of_sigmaOne mkSeq₂_definable _ _ +instance mkSeq₂_definable' (Γ m) : Γ-[m + 1]-Function₂ (fun x y : V ↦ !⟦x, y⟧) := mkSeq₂_definable.of_sigmaOne end @@ -423,7 +423,7 @@ lemma mem_vectoSeq {n : ℕ} (v : Fin n → V) (i : Fin n) : ⟪(i : V), v i⟫ end seqToVec -lemma sigma₁_order_ball_induction {f : V → V → V} (hf : 𝚺₁-Function₂ f) {P : V → V → Prop} (hP : 𝚺₁-Relation P) +lemma order_ball_induction_sigma1 {f : V → V → V} (hf : 𝚺₁-Function₂ f) {P : V → V → Prop} (hP : 𝚺₁-Relation P) (ind : ∀ x y, (∀ x' < x, ∀ y' ≤ f x y, P x' y') → P x y) : ∀ x y, P x y := by have maxf : ∀ x y, ∃ m, ∀ x' ≤ x, ∀ y' ≤ y, f x' y' ≤ m := by intro x y; @@ -517,6 +517,73 @@ lemma sigma₁_order_ball_induction {f : V → V → V} (hf : 𝚺₁-Function exact ih m₁ (by simp [m₁]) (by simp [m₁]) x'' (lt_succ_iff_le.mp (lt_of_lt_of_le hx'' hx')) y'' (le_trans hy'' this) exact this x (by rfl) y (lt_of_mem_rng hW₀) (by simpa using hW₀) x (by rfl) y (by rfl) +lemma order_ball_induction_sigma1' {f : V → V} (hf : 𝚺₁-Function₁ f) {P : V → V → Prop} (hP : 𝚺₁-Relation P) + (ind : ∀ x y, (∀ x' < x, ∀ y' ≤ f y, P x' y') → P x y) : ∀ x y, P x y := + have : 𝚺₁-Function₂ (fun _ ↦ f) := FirstOrder.Arith.HierarchySymbol.BoldfaceFunction.comp₁ (by simp) + order_ball_induction_sigma1 this hP ind + +lemma order_ball_induction₂_sigma1 {fy fz : V → V → V → V} + (hfy : 𝚺₁-Function₃ fy) (hfz : 𝚺₁-Function₃ fz) {P : V → V → V → Prop} (hP : 𝚺₁-Relation₃ P) + (ind : ∀ x y z, (∀ x' < x, ∀ y' ≤ fy x y z, ∀ z' ≤ fz x y z, P x' y' z') → P x y z) : ∀ x y z, P x y z := by + let Q : V → V → Prop := fun x w ↦ P x (π₁ w) (π₂ w) + have hQ : 𝚺₁-Relation Q := by + simp [Q] + apply HierarchySymbol.Boldface.comp₃ (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) + (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) + let f : V → V → V := fun x w ↦ ⟪fy x (π₁ w) (π₂ w), fz x (π₁ w) (π₂ w)⟫ + have hf : 𝚺₁-Function₂ f := by + simp [f] + apply HierarchySymbol.BoldfaceFunction.comp₂ + · apply HierarchySymbol.BoldfaceFunction.comp₃ (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₃ (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _) + intro x y z + simpa [Q] using order_ball_induction_sigma1 hf hQ (fun x w ih ↦ + ind x (π₁ w) (π₂ w) (fun x' hx' y' hy' z' hz' ↦ by simpa [Q] using ih x' hx' ⟪y', z'⟫ (pair_le_pair hy' hz'))) + x ⟪y, z⟫ + +lemma order_ball_induction₃_sigma1 {fy fz fw : V → V → V → V → V} + (hfy : 𝚺₁-Function₄ fy) (hfz : 𝚺₁-Function₄ fz) (hfw : 𝚺₁-Function₄ fw) {P : V → V → V → V → Prop} (hP : 𝚺₁-Relation₄ P) + (ind : ∀ x y z w, (∀ x' < x, ∀ y' ≤ fy x y z w, ∀ z' ≤ fz x y z w, ∀ w' ≤ fw x y z w, P x' y' z' w') → P x y z w) : ∀ x y z w, P x y z w := by + let Q : V → V → Prop := fun x v ↦ P x (π₁ v) (π₁ (π₂ v)) (π₂ (π₂ v)) + have hQ : 𝚺₁-Relation Q := by + simp [Q] + apply HierarchySymbol.Boldface.comp₄ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + let f : V → V → V := fun x v ↦ ⟪fy x (π₁ v) (π₁ (π₂ v)) (π₂ (π₂ v)), fz x (π₁ v) (π₁ (π₂ v)) (π₂ (π₂ v)), fw x (π₁ v) (π₁ (π₂ v)) (π₂ (π₂ v))⟫ + have hf : 𝚺₁-Function₂ f := by + simp [f] + apply HierarchySymbol.BoldfaceFunction.comp₂ + · apply HierarchySymbol.BoldfaceFunction.comp₄ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₂ + · apply HierarchySymbol.BoldfaceFunction.comp₄ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.BoldfaceFunction.comp₄ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _) + intro x y z w + have := order_ball_induction_sigma1 hf hQ (fun x v ih ↦ + ind x (π₁ v) (π₁ (π₂ v)) (π₂ (π₂ v)) (fun x' hx' y' hy' z' hz' w' hw' ↦ by + simpa [Q] using ih x' hx' ⟪y', z', w'⟫ (pair_le_pair hy' <| pair_le_pair hz' hw'))) + x ⟪y, z, w⟫ + simpa [Q] using this + end LO.Arith end diff --git a/Arithmetization/ISigmaOne/HFS/Supplemental.lean b/Arithmetization/ISigmaOne/HFS/Supplemental.lean index e603f1f..a2ccae3 100644 --- a/Arithmetization/ISigmaOne/HFS/Supplemental.lean +++ b/Arithmetization/ISigmaOne/HFS/Supplemental.lean @@ -109,8 +109,7 @@ lemma seqExp_defined : 𝚺₁-Function₂ (seqExp : V → V → V) via seqExpDe instance seqExp_definable : 𝚺₁-Function₂ (seqExp : V → V → V) := seqExp_defined.to_definable -@[simp, definability] instance seqExp_definable' (Γ) : Γ-[m + 1]-Function₂ (seqExp : V → V → V) := - .of_sigmaOne seqExp_definable _ _ +instance seqExp_definable' (Γ) : Γ-[m + 1]-Function₂ (seqExp : V → V → V) := seqExp_definable.of_sigmaOne lemma mem_seqExp_iff {s a k : V} : s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ (∀ i z, ⟪i, z⟫ ∈ s → z ∈ a) := by induction k using induction_pi1 generalizing s diff --git a/Arithmetization/ISigmaOne/HFS/Vec.lean b/Arithmetization/ISigmaOne/HFS/Vec.lean index 716718c..f4269f9 100644 --- a/Arithmetization/ISigmaOne/HFS/Vec.lean +++ b/Arithmetization/ISigmaOne/HFS/Vec.lean @@ -104,7 +104,7 @@ lemma mkVec₂_defined : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) via mkVec instance mkVec₂_definable : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) := mkVec₂_defined.to_definable -instance mkVec₂_definable' (Γ m) : Γ-[m + 1]-Function₂ (fun x y : V ↦ ?[x, y]) := .of_sigmaOne mkVec₂_definable _ _ +instance mkVec₂_definable' (Γ m) : Γ-[m + 1]-Function₂ (fun x y : V ↦ ?[x, y]) := mkVec₂_definable.of_sigmaOne end @@ -271,12 +271,12 @@ lemma cons_induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) · exact cons _ _ (ih v (by simp))) @[elab_as_elim] -lemma cons_induction_sigma₁ {P : V → Prop} (hP : 𝚺₁-Predicate P) +lemma cons_induction_sigma1 {P : V → Prop} (hP : 𝚺₁-Predicate P) (nil : P 0) (cons : ∀ x v, P v → P (x ∷ v)) : ∀ v, P v := cons_induction 𝚺 hP nil cons @[elab_as_elim] -lemma cons_induction_pi₁ {P : V → Prop} (hP : 𝚷₁-Predicate P) +lemma cons_induction_pi1 {P : V → Prop} (hP : 𝚷₁-Predicate P) (nil : P 0) (cons : ∀ x v, P v → P (x ∷ v)) : ∀ v, P v := cons_induction 𝚷 hP nil cons @@ -296,7 +296,7 @@ lemma nth_defined : 𝚺₁-Function₂ (nth : V → V → V) via nthDef := by instance nth_definable : 𝚺₁-Function₂ (nth : V → V → V) := nth_defined.to_definable -instance nth_definable' (Γ m) : Γ-[m + 1]-Function₂ (nth : V → V → V) := .of_sigmaOne nth_definable _ _ +instance nth_definable' (Γ m) : Γ-[m + 1]-Function₂ (nth : V → V → V) := nth_definable.of_sigmaOne end @@ -469,7 +469,7 @@ lemma graph_cons {x xs y : V} : variable (param) lemma graph_exists (xs : V) : ∃ y, c.Graph param ⟪xs, y⟫ := by - induction xs using cons_induction_sigma₁ + induction xs using cons_induction_sigma1 · definability case nil => exact ⟨c.nil param, c.graph_nil.mpr rfl⟩ @@ -480,7 +480,7 @@ lemma graph_exists (xs : V) : ∃ y, c.Graph param ⟪xs, y⟫ := by variable {param} lemma graph_unique {xs y₁ y₂ : V} : c.Graph param ⟪xs, y₁⟫ → c.Graph param ⟪xs, y₂⟫ → y₁ = y₂ := by - induction xs using cons_induction_pi₁ generalizing y₁ y₂ + induction xs using cons_induction_pi1 generalizing y₁ y₂ · definability case nil => simp [graph_nil]; rintro rfl rfl; rfl @@ -524,7 +524,7 @@ instance result_definable : 𝚺₁.BoldfaceFunction (fun v ↦ c.result (v ·.s c.result_defined.to_definable instance result_definable' (Γ m) : - Γ-[m + 1].BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0)) := .of_sigmaOne c.result_definable _ _ + Γ-[m + 1].BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0)) := c.result_definable.of_sigmaOne end @@ -573,7 +573,7 @@ lemma len_defined : 𝚺₁-Function₁ (len : V → V) via lenDef := constructi instance len_definable : 𝚺₁-Function₁ (len : V → V) := len_defined.to_definable -instance len_definable' (Γ m) : Γ-[m + 1]-Function₁ (len : V → V) := .of_sigmaOne len_definable _ _ +instance len_definable' (Γ m) : Γ-[m + 1]-Function₁ (len : V → V) := len_definable.of_sigmaOne end @@ -581,7 +581,7 @@ end rcases nil_or_cons v with (rfl | ⟨x, v, rfl⟩) <;> simp lemma nth_lt_len {v i : V} (hl : len v ≤ i) : v.[i] = 0 := by - induction v using cons_induction_pi₁ generalizing i + induction v using cons_induction_pi1 generalizing i · definability case nil => simp case cons x v ih => @@ -590,7 +590,7 @@ lemma nth_lt_len {v i : V} (hl : len v ≤ i) : v.[i] = 0 := by simpa using ih (by simpa using hl) @[simp] lemma len_le (v : V) : len v ≤ v := by - induction v using cons_induction_pi₁ + induction v using cons_induction_pi1 · definability case nil => simp case cons x v ih => @@ -601,7 +601,7 @@ lemma nth_lt_len {v i : V} (hl : len v ≤ i) : v.[i] = 0 := by end len lemma nth_ext {v₁ v₂ : V} (hl : len v₁ = len v₂) (H : ∀ i < len v₁, v₁.[i] = v₂.[i]) : v₁ = v₂ := by - induction v₁ using cons_induction_pi₁ generalizing v₂ + induction v₁ using cons_induction_pi1 generalizing v₂ · definability case nil => exact Eq.symm <| len_zero_iff_eq_nil.mp (by simp [←hl]) @@ -616,7 +616,7 @@ lemma nth_ext' (l : V) {v₁ v₂ : V} (hl₁ : len v₁ = l) (hl₂ : len v₂ rcases hl₂; exact nth_ext hl₁ (by simpa [hl₁] using H) lemma le_of_nth_le_nth {v₁ v₂ : V} (hl : len v₁ = len v₂) (H : ∀ i < len v₁, v₁.[i] ≤ v₂.[i]) : v₁ ≤ v₂ := by - induction v₁ using cons_induction_pi₁ generalizing v₂ + induction v₁ using cons_induction_pi1 generalizing v₂ · definability case nil => simp case cons x₁ v₁ ih => @@ -654,6 +654,87 @@ lemma eq_doubleton_of_len_eq_two {v : V} : len v = 2 ↔ ∃ x y, v = ?[x, y] := nth_ext (by simp [h, one_add_one_eq_two]) (by simp [lt_two_iff_le_one, le_one_iff_eq_zero_or_one, h])⟩ · rintro ⟨x, y, rfl⟩; simp [one_add_one_eq_two] + +/-! + +### Maximum of List + +-/ + +namespace ListMax + +def blueprint : VecRec.Blueprint 0 where + nil := .mkSigma “y | y = 0” (by simp) + cons := .mkSigma “y x xs ih | !FirstOrder.Arith.max y x ih” (by simp) + +def construction : VecRec.Construction V blueprint where + nil _ := 0 + cons _ x _ ih := max x ih + nil_defined := by intro v; simp [blueprint] + cons_defined := by intro v; simp [blueprint]; rfl + +end ListMax + +section listMax + +open ListMax + +def listMax (v : V) : V := construction.result ![] v + +@[simp] lemma listMax_nil : listMax (0 : V) = 0 := by simp [listMax, construction] + +@[simp] lemma listMax_cons (x v : V) : listMax (x ∷ v) = max x (listMax v) := by simp [listMax, construction] + +section + +def _root_.LO.FirstOrder.Arith.listMaxDef : 𝚺₁.Semisentence 2 := blueprint.resultDef + +lemma listMax_defined : 𝚺₁-Function₁ (listMax : V → V) via listMaxDef := construction.result_defined + +@[simp] lemma eval_listMaxDef (v) : + Semiformula.Evalbm V v listMaxDef.val ↔ v 0 = listMax (v 1) := listMax_defined.df.iff v + +instance listMax_definable : 𝚺₁-Function₁ (listMax : V → V) := listMax_defined.to_definable + +instance listMax_definable' (Γ m) : Γ-[m + 1]-Function₁ (listMax : V → V) := listMax_definable.of_sigmaOne + +end + +lemma nth_le_listMax {i v : V} (h : i < len v) : v.[i] ≤ listMax v := by + induction v using cons_induction_pi1 generalizing i + · definability + case nil => simp + case cons x v ih => + rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) + · simp + · simp [ih (by simpa using h)] + +lemma listMaxss_le {v z : V} (h : ∀ i < len v, v.[i] ≤ z) : listMax v ≤ z := by + induction v using cons_induction_pi1 + · definability + case nil => simp + case cons x v ih => + simp only [listMax_cons, max_le_iff] + constructor + · simpa using h 0 (by simp) + · exact ih (fun i hi ↦ by simpa using h (i + 1) (by simp [hi])) + +lemma listMaxss_le_iff {v z : V} : listMax v ≤ z ↔ ∀ i < len v, v.[i] ≤ z := by + constructor + · intro h i hi; exact le_trans (nth_le_listMax hi) h + · exact listMaxss_le + +/- +lemma nth_le_listMaxs (v : V) (hv : v ≠ 0) : ∃ i < len v, v.[i] = listMax v := by + induction v using cons_induction_sigma1 + · definability + case nil => simp at hv + case cons x v ih => + simp +-/ + +end listMax + /-! ### Take Last k-Element @@ -705,12 +786,12 @@ lemma takeLast_defined : 𝚺₁-Function₂ (takeLast : V → V → V) via take instance takeLast_definable : 𝚺₁-Function₂ (takeLast : V → V → V) := takeLast_defined.to_definable -instance takeLast_definable' (Γ m) : Γ-[m + 1]-Function₂ (takeLast : V → V → V) := .of_sigmaOne takeLast_definable _ _ +instance takeLast_definable' (Γ m) : Γ-[m + 1]-Function₂ (takeLast : V → V → V) := takeLast_definable.of_sigmaOne end lemma len_takeLast {v k : V} (h : k ≤ len v) : len (takeLast v k) = k := by - induction v using cons_induction_sigma₁ + induction v using cons_induction_sigma1 · definability case nil => simp_all case cons x v ih => @@ -730,13 +811,13 @@ lemma len_takeLast {v k : V} (h : k ≤ len v) : len (takeLast v k) = k := by @[simp] lemma add_sub_add (a b c : V) : (a + c) - (b + c) = a - b := add_tsub_add_eq_tsub_right a c b @[simp] lemma takeLast_zero (v : V) : takeLast v 0 = 0 := by - induction v using cons_induction_sigma₁ + induction v using cons_induction_sigma1 · definability case nil => simp case cons x v ih => simp [takeLast_cons, ih] lemma takeLast_succ_of_lt {i v : V} (h : i < len v) : takeLast v (i + 1) = v.[len v - (i + 1)] ∷ takeLast v i := by - induction v using cons_induction_sigma₁ generalizing i + induction v using cons_induction_sigma1 generalizing i · definability case nil => simp at h case cons x v ih => @@ -792,18 +873,18 @@ lemma concat_defined : 𝚺₁-Function₂ (concat : V → V → V) via concatDe instance concat_definable : 𝚺₁-Function₂ (concat : V → V → V) := concat_defined.to_definable -instance concat_definable' (Γ m) : Γ-[m + 1]-Function₂ (concat : V → V → V) := .of_sigmaOne concat_definable _ _ +instance concat_definable' (Γ m) : Γ-[m + 1]-Function₂ (concat : V → V → V) := concat_definable.of_sigmaOne end @[simp] lemma len_concat (v z : V) : len (concat v z) = len v + 1 := by - induction v using cons_induction_sigma₁ + induction v using cons_induction_sigma1 · definability case nil => simp case cons x v ih => simp [ih] lemma concat_nth_lt (v z : V) {i} (hi : i < len v) : (concat v z).[i] = v.[i] := by - induction v using cons_induction_sigma₁ generalizing i + induction v using cons_induction_sigma1 generalizing i · definability case nil => simp at hi case cons x v ih => @@ -812,7 +893,7 @@ lemma concat_nth_lt (v z : V) {i} (hi : i < len v) : (concat v z).[i] = v.[i] := · simp [ih (by simpa using hi)] @[simp] lemma concat_nth_len (v z : V) : (concat v z).[len v] = z := by - induction v using cons_induction_sigma₁ + induction v using cons_induction_sigma1 · definability case nil => simp case cons x v ih => simp [ih] @@ -867,7 +948,7 @@ lemma memVec_defined : 𝚫₁-Relation (MemVec : V → V → Prop) via memVecDe instance memVec_definable : 𝚫₁-Relation (MemVec : V → V → Prop) := memVec_defined.to_definable -instance memVec_definable' (Γ m) : Γ-[m + 1]-Relation (MemVec : V → V → Prop) := .of_deltaOne memVec_definable _ _ +instance memVec_definable' (Γ m) : Γ-[m + 1]-Relation (MemVec : V → V → Prop) := memVec_definable.of_deltaOne end @@ -908,7 +989,7 @@ lemma subsetVec_defined : 𝚫₁-Relation (SubsetVec : V → V → Prop) via su instance subsetVec_definable : 𝚫₁-Relation (SubsetVec : V → V → Prop) := subsetVec_defined.to_definable -instance subsetVec_definable' (Γ m) : Γ-[m + 1]-Relation (SubsetVec : V → V → Prop) := .of_deltaOne subsetVec_definable _ _ +instance subsetVec_definable' (Γ m) : Γ-[m + 1]-Relation (SubsetVec : V → V → Prop) := subsetVec_definable.of_deltaOne end @@ -951,8 +1032,7 @@ lemma repeatVec_defined : 𝚺₁-Function₂ (repeatVec : V → V → V) via re instance repeatVec_definable : 𝚺₁-Function₂ (repeatVec : V → V → V) := repeatVec_defined.to_definable -@[simp] instance repeatVec_definable' (Γ) : Γ-[m + 1]-Function₂ (repeatVec : V → V → V) := - .of_sigmaOne repeatVec_definable _ _ +instance repeatVec_definable' (Γ) : Γ-[m + 1]-Function₂ (repeatVec : V → V → V) := repeatVec_definable.of_sigmaOne end @@ -1021,12 +1101,12 @@ lemma vecToSet_defined : 𝚺₁-Function₁ (vecToSet : V → V) via vecToSetDe instance vecToSet_definable : 𝚺₁-Function₁ (vecToSet : V → V) := vecToSet_defined.to_definable -instance vecToSet_definable' (Γ) : Γ-[m + 1]-Function₁ (vecToSet : V → V) := .of_sigmaOne vecToSet_definable _ _ +instance vecToSet_definable' (Γ) : Γ-[m + 1]-Function₁ (vecToSet : V → V) := vecToSet_definable.of_sigmaOne end lemma mem_vecToSet_iff {v x : V} : x ∈ vecToSet v ↔ ∃ i < len v, x = v.[i] := by - induction v using cons_induction_sigma₁ + induction v using cons_induction_sigma1 · definability case nil => simp case cons y v ih => diff --git a/Arithmetization/ISigmaOne/Metamath/Coding.lean b/Arithmetization/ISigmaOne/Metamath/Coding.lean index c3d23e1..9053cf0 100644 --- a/Arithmetization/ISigmaOne/Metamath/Coding.lean +++ b/Arithmetization/ISigmaOne/Metamath/Coding.lean @@ -4,9 +4,39 @@ import Mathlib.Combinatorics.Colex namespace LO.FirstOrder -namespace Semiformula.Operator +variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] + +variable {ξ : Type*} [Encodable ξ] + +open Encodable + +namespace Semiterm + +lemma encode_eq_toNat (t : Semiterm L ξ n) : Encodable.encode t = toNat t := rfl + +lemma toNat_func {k} (f : L.Func k) (v : Fin k → Semiterm L ξ n) : + toNat (func f v) = (Nat.pair 2 <| Nat.pair k <| Nat.pair (encode f) <| Matrix.vecToNat fun i ↦ toNat (v i)) + 1 := rfl + +@[simp] lemma encode_emb (t : Semiterm L Empty n) : Encodable.encode (Rew.emb t : Semiterm L ξ n) = Encodable.encode t := by + simp only [encode_eq_toNat] + induction t + · simp [toNat] + · contradiction + · simp [Rew.func, toNat_func, *] + +end Semiterm -variable {L : Language} +namespace Semiformula + +lemma encode_eq_toNat (p : Semiformula L ξ n) : Encodable.encode p = toNat p := rfl + +@[simp] lemma encode_emb (p : Semisentence L n) : Encodable.encode (Rew.emb.hom p : Semiformula L ξ n) = Encodable.encode p := by + simp [encode_eq_toNat] + induction p using rec' <;> simp [toNat, Rew.rel, Rew.nrel, *] + +end Semiformula + +namespace Semiformula.Operator lemma lt_eq [L.LT] (t u : Semiterm L ξ n) : LT.lt.operator ![t, u] = Semiformula.rel Language.LT.lt ![t, u] := by simp [operator, LT.sentence_eq, Rew.rel] @@ -73,6 +103,30 @@ lemma quote_matrix_succ (v : Fin (k + 1) → V) : · simp · simp [quote_matrix_succ, ih, cons_absolute] +lemma quote_eq_vecToNat (v : Fin k → ℕ) : ⌜v⌝ = Matrix.vecToNat v := by + induction k + case zero => simp + case succ k ih => + simp [quote_matrix_succ, Matrix.vecToNat, cons, nat_pair_eq, Function.comp, ih] + +section + +variable {α : Type*} [Encodable α] + +instance : GoedelQuote α V := ⟨fun x ↦ Encodable.encode x⟩ + +lemma quote_eq_coe_encode (x : α) : (⌜x⌝ : V) = Encodable.encode x := rfl + +@[simp] lemma quote_nat (n : ℕ) : (⌜n⌝ : V) = n := rfl + +lemma coe_quote (x : α) : ↑(⌜x⌝ : ℕ) = (⌜x⌝ : V) := by simp [quote_eq_coe_encode] + +@[simp] lemma quote_quote (x : α) : (⌜(⌜x⌝ : ℕ)⌝ : V) = ⌜x⌝ := by simp [quote_eq_coe_encode] + +lemma quote_eq_encode (x : α) : (⌜x⌝ : ℕ) = Encodable.encode x := by simp [quote_eq_coe_encode] + +end + end LO.Arith namespace LO.FirstOrder.Semiterm @@ -85,21 +139,17 @@ variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Enco variable (V) -def codeIn {n} : SyntacticSemiterm L n → V - | #z => ^#z - | &x => ^&x - | func (arity := k) f v => ^func (k : V) ⌜f⌝ ⌜fun i ↦ (v i).codeIn⌝ +lemma quote_eq_toNat (t : SyntacticSemiterm L n) : (⌜t⌝ : V) = toNat t := rfl -instance : GoedelQuote (SyntacticSemiterm L n) V := ⟨(·.codeIn V)⟩ +lemma quote_bvar (z : Fin n) : ⌜(#z : SyntacticSemiterm L n)⌝ = ^#(z : V) := by + simp [quote_eq_toNat, toNat, qqBvar, ←nat_pair_eq, nat_cast_pair] -lemma quote_syntacticSemiterm_def (t : SyntacticSemiterm L n) : ⌜t⌝ = t.codeIn V := rfl - -lemma quote_bvar (z : Fin n) : ⌜(#z : SyntacticSemiterm L n)⌝ = ^#(z : V) := rfl - -lemma quote_fvar (x : ℕ) : ⌜(&x : SyntacticSemiterm L n)⌝ = ^&(x : V) := rfl +lemma quote_fvar (x : ℕ) : ⌜(&x : SyntacticSemiterm L n)⌝ = ^&(x : V) := by + simp [quote_eq_toNat, toNat, qqFvar, ←nat_pair_eq, nat_cast_pair] lemma quote_func {k} (f : L.Func k) (v : Fin k → SyntacticSemiterm L n) : - ⌜func f v⌝ = ^func (k : V) ⌜f⌝ ⌜fun i ↦ ⌜v i⌝⌝ := rfl + ⌜func f v⌝ = ^func (k : V) ⌜f⌝ ⌜fun i ↦ ⌜v i⌝⌝ := by + simp [quote_eq_toNat, toNat, qqFunc, ←nat_pair_eq, nat_cast_pair, quote_func_def, quote_eq_vecToNat, ←quote_matrix_absolute] @[simp] lemma codeIn_inj {n} {t u : SyntacticSemiterm L n} : (⌜t⌝ : V) = ⌜u⌝ ↔ t = u := by induction t generalizing u @@ -130,6 +180,13 @@ lemma quote_func {k} (f : L.Func k) (v : Fin k → SyntacticSemiterm L n) : ((⌜t⌝ : ℕ) : V) = ⌜t⌝ := by induction t <;> simp [quote_bvar, quote_fvar, quote_func, qqBvar, qqFvar, qqFunc, Fin.val_inj, nat_cast_pair, *] +lemma quote_eq_encode (t : SyntacticSemiterm L n) : ⌜t⌝ = Encodable.encode t := by + induction t + case bvar z => simp [encode_eq_toNat, toNat, quote_bvar, qqBvar, nat_pair_eq] + case fvar z => simp [encode_eq_toNat, toNat, quote_fvar, qqFvar, nat_pair_eq] + case func k f v ih => + simp [encode_eq_toNat, toNat, quote_func, qqFunc, nat_pair_eq, quote_func_def, quote_eq_vecToNat, ih] + end LO.FirstOrder.Semiterm namespace LO.Arith @@ -146,82 +203,91 @@ lemma eq_fin_of_lt_nat {n : ℕ} {x : V} (hx : x < n) : ∃ i : Fin n, x = i := exact ⟨⟨x, by simpa using hx⟩, by simp⟩ @[simp] lemma semiterm_codeIn {n} (t : SyntacticSemiterm L n) : - (L.codeIn V).Semiterm n ⌜t⌝ := by + (L.codeIn V).IsSemiterm n ⌜t⌝ := by induction t <;> simp [quote_bvar, quote_fvar, quote_func] case func k f v ih => + apply Language.IsSemitermVec.iff.mpr exact ⟨by simp, by rintro i hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ simpa using ih i⟩ @[simp] lemma semitermVec_codeIn {k n} (v : Fin k → SyntacticSemiterm L n) : - (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := + (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := Language.IsSemitermVec.iff.mpr ⟨by simp, by intro i hi; rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩; simp⟩ +@[simp] lemma isUTermVec_codeIn {k n} (v : Fin k → SyntacticSemiterm L n) : + (L.codeIn V).IsUTermVec k ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v |>.isUTerm + lemma termSubst_quote {n m} (t : SyntacticSemiterm L n) (w : Fin n → SyntacticSemiterm L m) : - (L.codeIn V).termSubst ↑n ↑m ⌜fun i ↦ ⌜w i⌝⌝ ⌜t⌝ = ⌜Rew.substs w t⌝ := by + (L.codeIn V).termSubst ⌜fun i ↦ ⌜w i⌝⌝ ⌜t⌝ = ⌜Rew.substs w t⌝ := by induction t case bvar z => simp [quote_bvar, quote_fvar, quote_func] case fvar x => simp [quote_bvar, quote_fvar, quote_func] case func k f v ih => - have Hw : (L.codeIn V).SemitermVec n m ⌜fun i ↦ ⌜w i⌝⌝ := semitermVec_codeIn w - have Hv : (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v - simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, termSubst_func (codeIn_func_quote f) Hv] + have Hw : (L.codeIn V).IsSemitermVec n m ⌜fun i ↦ ⌜w i⌝⌝ := semitermVec_codeIn w + have Hv : (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v + simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, Language.termSubst_func (codeIn_func_quote f) Hv.isUTerm] congr - apply nth_ext (by simp [←Hw.termSubstVec Hv |>.lh]) + apply nth_ext (by simp [(Hw.termSubstVec Hv).lh]) intro i hi - have hi : i < k := by simpa [← Hw.termSubstVec Hv |>.lh] using hi + have hi : i < k := by simpa [Hw.termSubstVec Hv |>.lh] using hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simpa [nth_termSubstVec] using ih i + rw [nth_termSubstVec Hv.isUTerm hi] + simpa using ih i lemma termSubstVec_quote {k n m} (w : Fin n → SyntacticSemiterm L m) (v : Fin k → SyntacticSemiterm L n) : - (L.codeIn V).termSubstVec ↑k ↑n ↑m ⌜fun i ↦ ⌜w i⌝⌝ ⌜fun i => ⌜v i⌝⌝ = ⌜fun i ↦ ⌜(Rew.substs w) (v i)⌝⌝ := by - have Hw : (L.codeIn V).SemitermVec n m ⌜fun i ↦ ⌜w i⌝⌝ := semitermVec_codeIn w - have Hv : (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v - apply nth_ext (by simp [←Hw.termSubstVec Hv |>.lh]) + (L.codeIn V).termSubstVec ↑k ⌜fun i ↦ ⌜w i⌝⌝ ⌜fun i => ⌜v i⌝⌝ = ⌜fun i ↦ ⌜(Rew.substs w) (v i)⌝⌝ := by + have Hw : (L.codeIn V).IsSemitermVec n m ⌜fun i ↦ ⌜w i⌝⌝ := semitermVec_codeIn w + have Hv : (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v + apply nth_ext (by simp [Hw.termSubstVec Hv |>.lh]) intro i hi - have hi : i < k := by simpa [← Hw.termSubstVec Hv |>.lh] using hi + have hi : i < k := by simpa [Hw.termSubstVec Hv |>.lh] using hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simpa [nth_termSubstVec] using termSubst_quote (v i) w + rw [nth_termSubstVec Hv.isUTerm hi] + simpa using termSubst_quote (v i) w lemma termShift_quote {n} (t : SyntacticSemiterm L n) : - (L.codeIn V).termShift n ⌜t⌝ = ⌜Rew.shift t⌝ := by + (L.codeIn V).termShift ⌜t⌝ = ⌜Rew.shift t⌝ := by induction t case bvar => simp [quote_bvar, quote_fvar, quote_func] case fvar => simp [quote_bvar, quote_fvar, quote_func] case func k f v ih => - have Hv : (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v - simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, termShift_func (codeIn_func_quote f) Hv] + have Hv : (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v + simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, Language.termShift_func (codeIn_func_quote f) Hv.isUTerm] congr - apply nth_ext (by simp [←Hv.termShiftVec |>.lh]) + apply nth_ext (by simp [Hv.termShiftVec |>.lh]) intro i hi - have hi : i < k := by simpa [← Hv.termShiftVec |>.lh] using hi + have hi : i < k := by simpa [Hv.termShiftVec |>.lh] using hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simpa [nth_termShiftVec] using ih i + rw [nth_termShiftVec Hv.isUTerm hi] + simpa using ih i lemma termShiftVec_quote {k n} (v : Fin k → SyntacticSemiterm L n) : - (L.codeIn V).termShiftVec k n ⌜fun i ↦ ⌜v i⌝⌝ = ⌜fun i ↦ ⌜Rew.shift (v i)⌝⌝ := by - have Hv : (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v - apply nth_ext (by simp [←Hv.termShiftVec |>.lh]) + (L.codeIn V).termShiftVec k ⌜fun i ↦ ⌜v i⌝⌝ = ⌜fun i ↦ ⌜Rew.shift (v i)⌝⌝ := by + have Hv : (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v + apply nth_ext (by simp [Hv.termShiftVec |>.lh]) intro i hi - have hi : i < k := by simpa [← Hv.termShiftVec |>.lh] using hi + have hi : i < k := by simpa [Hv.termShiftVec |>.lh] using hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simpa [nth_termShiftVec] using termShift_quote (v i) + rw [nth_termShiftVec Hv.isUTerm hi] + simpa using termShift_quote (v i) lemma termBShift_quote {n} (t : SyntacticSemiterm L n) : - (L.codeIn V).termBShift n ⌜t⌝ = ⌜Rew.bShift t⌝ := by + (L.codeIn V).termBShift ⌜t⌝ = ⌜Rew.bShift t⌝ := by induction t case bvar => simp [quote_bvar, quote_fvar, quote_func] case fvar => simp [quote_bvar, quote_fvar, quote_func] case func k f v ih => - have Hv : (L.codeIn V).SemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v - simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, termBShift_func (codeIn_func_quote f) Hv] + have Hv : (L.codeIn V).IsSemitermVec k n ⌜fun i ↦ ⌜v i⌝⌝ := semitermVec_codeIn v + simp only [Rew.func, Semiterm.quote_func, codeIn_func_quote, Language.termBShift_func (codeIn_func_quote f) Hv.isUTerm] congr - apply nth_ext (by simp [←Hv.termBShiftVec |>.lh]) + apply nth_ext (by simp [Hv.termBShiftVec |>.lh]) intro i hi - have hi : i < k := by simpa [← Hv.termBShiftVec |>.lh] using hi + have hi : i < k := by simpa [Hv.termBShiftVec |>.lh] using hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simpa [nth_termBShiftVec] using ih i + rw [nth_termBShiftVec Hv.isUTerm hi] + simpa using ih i end LO.Arith @@ -233,44 +299,36 @@ variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] -variable (V) - -def codeIn : {n : ℕ} → SyntacticSemiformula L n → V - | n, rel (arity := k) R v => ^rel (n : V) (k : V) ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ - | n, nrel (arity := k) R v => ^nrel (n : V) (k : V) ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ - | n, ⊤ => ^⊤[n] - | n, ⊥ => ^⊥[n] - | n, p ⋏ q => p.codeIn ^⋏[n] q.codeIn - | n, p ⋎ q => p.codeIn ^⋎[n] q.codeIn - | n, ∀' p => ^∀[n] p.codeIn - | n, ∃' p => ^∃[n] p.codeIn - -instance : GoedelQuote (SyntacticSemiformula L n) V := ⟨codeIn V⟩ - -lemma quote_syntacticSemiformula_def (p : SyntacticSemiformula L n) : ⌜p⌝ = p.codeIn V := rfl - -lemma quote_rel {k} (R : L.Rel k) (v : Fin k → SyntacticSemiterm L n) : - (⌜rel R v⌝ : V) = ^rel ↑n ↑k ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ := rfl -lemma quote_nrel {k} (R : L.Rel k) (v : Fin k → SyntacticSemiterm L n) : - (⌜nrel R v⌝ : V) = ^nrel ↑n ↑k ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ := rfl -lemma quote_verum (n : ℕ) : ⌜(⊤ : SyntacticSemiformula L n)⌝ = ^⊤[(n : V)] := rfl -lemma quote_falsum (n : ℕ) : ⌜(⊥ : SyntacticSemiformula L n)⌝ = ^⊥[(n : V)] := rfl -lemma quote_and (p q : SyntacticSemiformula L n) : ⌜p ⋏ q⌝ = ⌜p⌝ ^⋏[(n : V)] ⌜q⌝ := rfl -lemma quote_or (p q : SyntacticSemiformula L n) : ⌜p ⋎ q⌝ = ⌜p⌝ ^⋎[(n : V)] ⌜q⌝ := rfl -lemma quote_all (p : SyntacticSemiformula L (n + 1)) : ⌜∀' p⌝ = ^∀[(n : V)] ⌜p⌝ := rfl -lemma quote_ex (p : SyntacticSemiformula L (n + 1)) : ⌜∃' p⌝ = ^∃[(n : V)] ⌜p⌝ := rfl +lemma quote_eq_toNat (p : SyntacticSemiformula L n) : (⌜p⌝ : V) = toNat p := rfl + +lemma quote_rel {k} (R : L.Rel k) (v : Fin k → SyntacticSemiterm L n) : (⌜rel R v⌝ : V) = ^rel ↑k ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ := by + simp [Semiterm.quote_eq_toNat, quote_eq_toNat, toNat, qqRel, ←nat_pair_eq, nat_cast_pair, quote_rel_def, ←quote_eq_vecToNat]; rfl +lemma quote_nrel {k} (R : L.Rel k) (v : Fin k → SyntacticSemiterm L n) : (⌜nrel R v⌝ : V) = ^nrel ↑k ⌜R⌝ ⌜fun i ↦ ⌜v i⌝⌝ := by + simp [Semiterm.quote_eq_toNat, quote_eq_toNat, toNat, qqRel, ←nat_pair_eq, nat_cast_pair, quote_rel_def, ←quote_eq_vecToNat]; rfl +lemma quote_verum (n : ℕ) : ⌜(⊤ : SyntacticSemiformula L n)⌝ = (^⊤ : V) := by + simp [quote_eq_toNat, toNat, qqVerum, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] +lemma quote_falsum (n : ℕ) : ⌜(⊥ : SyntacticSemiformula L n)⌝ = (^⊥ : V) := by + simp [quote_eq_toNat, toNat, qqFalsum, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] +lemma quote_and (p q : SyntacticSemiformula L n) : (⌜p ⋏ q⌝ : V) = ⌜p⌝ ^⋏ ⌜q⌝ := by + simp [quote_eq_toNat, toNat, qqAnd, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] +lemma quote_or (p q : SyntacticSemiformula L n) : (⌜p ⋎ q⌝ : V) = ⌜p⌝ ^⋎ ⌜q⌝ := by + simp [quote_eq_toNat, toNat, qqOr, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] +lemma quote_all (p : SyntacticSemiformula L (n + 1)) : (⌜∀' p⌝ : V) = ^∀ ⌜p⌝ := by + simp [quote_eq_toNat, toNat, qqAll, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] +lemma quote_ex (p : SyntacticSemiformula L (n + 1)) : (⌜∃' p⌝ : V) = ^∃ ⌜p⌝ := by + simp [quote_eq_toNat, toNat, qqEx, pair_coe_eq_coe_pair, ←pair_coe_eq_coe_pair, nat_cast_pair] @[simp] lemma quote_eq (t u : SyntacticSemiterm ℒₒᵣ n) : - (⌜Semiformula.rel Language.Eq.eq ![t, u]⌝ : V) = (⌜t⌝ ^=[(n : V)] ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_rel]; rfl + (⌜Semiformula.rel Language.Eq.eq ![t, u]⌝ : V) = (⌜t⌝ ^= ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_rel]; rfl @[simp] lemma quote_neq (t u : SyntacticSemiterm ℒₒᵣ n) : - (⌜Semiformula.nrel Language.Eq.eq ![t, u]⌝ : V) = (⌜t⌝ ^≠[(n : V)] ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_nrel]; rfl + (⌜Semiformula.nrel Language.Eq.eq ![t, u]⌝ : V) = (⌜t⌝ ^≠ ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_nrel]; rfl @[simp] lemma quote_lt (t u : SyntacticSemiterm ℒₒᵣ n) : - (⌜Semiformula.rel Language.LT.lt ![t, u]⌝ : V) = (⌜t⌝ ^<[(n : V)] ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_rel]; rfl + (⌜Semiformula.rel Language.LT.lt ![t, u]⌝ : V) = (⌜t⌝ ^< ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_rel]; rfl @[simp] lemma quote_nlt (t u : SyntacticSemiterm ℒₒᵣ n) : - (⌜Semiformula.nrel Language.LT.lt ![t, u]⌝ : V) = (⌜t⌝ ^≮[(n : V)] ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_nrel]; rfl + (⌜Semiformula.nrel Language.LT.lt ![t, u]⌝ : V) = (⌜t⌝ ^≮ ⌜u⌝) := by simp [FirstOrder.Semiformula.quote_nrel]; rfl @[simp] lemma codeIn_inj {n} {p q : SyntacticSemiformula L n} : (⌜p⌝ : V) = ⌜q⌝ ↔ p = q := by induction p using rec' @@ -319,27 +377,14 @@ lemma quote_ex (p : SyntacticSemiformula L (n + 1)) : ⌜∃' p⌝ = ^∃[(n : V qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] rw [ih] -@[simp] lemma quote_absolute (p : SyntacticSemiformula L n) : - ((⌜p⌝ : ℕ) : V) = ⌜p⌝ := by - induction p using rec' <;> simp [quote_rel, quote_nrel, quote_verum, quote_falsum, quote_and, quote_or, quote_all, quote_ex, - qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx, nat_cast_pair, *] - -instance : GoedelQuote (Semisentence L n) V := ⟨fun σ ↦ ⌜(Rew.emb.hom σ : SyntacticSemiformula L n)⌝⟩ - -lemma quote_semisentence_def (p : Semisentence L n) : (⌜p⌝ : V) = ⌜(Rew.emb.hom p : SyntacticSemiformula L n)⌝ := rfl - -@[simp] lemma quote_semisentence_absolute (p : Semisentence L n) : ((⌜p⌝ : ℕ) : V) = ⌜p⌝ := by - simp [quote_semisentence_def] - -instance : Semiterm.Operator.GoedelNumber ℒₒᵣ (SyntacticFormula L) := ⟨fun p ↦ Semiterm.Operator.numeral ℒₒᵣ ⌜p⌝⟩ - -instance : Semiterm.Operator.GoedelNumber ℒₒᵣ (Sentence L) := ⟨fun σ ↦ Semiterm.Operator.numeral ℒₒᵣ ⌜σ⌝⟩ +@[simp] lemma quote_semisentence_def (p : Semisentence L n) : ⌜(Rew.emb.hom p : SyntacticSemiformula L n)⌝ = (⌜p⌝ : V) := by + simp [quote_eq_coe_encode] lemma sentence_goedelNumber_def (σ : Sentence L) : - (⌜σ⌝ : Semiterm ℒₒᵣ ξ n) = Semiterm.Operator.numeral ℒₒᵣ ⌜σ⌝ := rfl + (⌜σ⌝ : Semiterm ℒₒᵣ ξ n) = Semiterm.Operator.numeral ℒₒᵣ ⌜σ⌝ := by simp [Arith.goedelNumber'_def, quote_eq_encode] lemma syntacticformula_goedelNumber_def (p : SyntacticFormula L) : - (⌜p⌝ : Semiterm ℒₒᵣ ξ n) = Semiterm.Operator.numeral ℒₒᵣ ⌜p⌝ := rfl + (⌜p⌝ : Semiterm ℒₒᵣ ξ n) = Semiterm.Operator.numeral ℒₒᵣ ⌜p⌝ := by simp [Arith.goedelNumber'_def, quote_eq_encode] end LO.FirstOrder.Semiformula @@ -352,7 +397,7 @@ variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @[simp] lemma semiformula_quote {n} (p : SyntacticSemiformula L n) : - (L.codeIn V).Semiformula n ⌜p⌝ := by + (L.codeIn V).IsSemiformula n ⌜p⌝ := by induction p using Semiformula.rec' case hrel n k r v => simp [Semiformula.quote_rel] case hnrel n k r v => simp [Semiformula.quote_nrel] @@ -363,8 +408,11 @@ variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Enco case hall n p ihp => simpa [Semiformula.quote_all] using ihp case hex n p ihp => simpa [Semiformula.quote_ex] using ihp +@[simp] lemma isUFormula_quote {n} (p : SyntacticSemiformula L n) : + (L.codeIn V).IsUFormula ⌜p⌝ := semiformula_quote p |>.isUFormula + @[simp] lemma semiformula_quote_succ {n} (p : SyntacticSemiformula L (n + 1)) : - (L.codeIn V).Semiformula (n + 1) ⌜p⌝ := by simpa using semiformula_quote p + (L.codeIn V).IsSemiformula (n + 1) ⌜p⌝ := by simpa using semiformula_quote p lemma neg_quote {n} (p : SyntacticSemiformula L n) : (L.codeIn V).neg ⌜p⌝ = ⌜~p⌝ := by @@ -378,21 +426,21 @@ lemma shift_quote {n} (p : SyntacticSemiformula L n) : Rew.rel, Rew.nrel, termShiftVec_quote] lemma qVec_quote (w : Fin n → SyntacticSemiterm L m) : - (L.codeIn V).qVec ↑n ↑m ⌜fun i => ⌜w i⌝⌝ = ⌜^#0 :> fun i ↦ (⌜Rew.bShift (w i)⌝ : V)⌝ := by - have Hw : (L.codeIn V).SemitermVec ↑n (↑m + 1) ((L.codeIn V).termBShiftVec ↑n ↑m ⌜fun i ↦ ⌜w i⌝⌝) := + (L.codeIn V).qVec ⌜fun i => ⌜w i⌝⌝ = ⌜^#0 :> fun i ↦ (⌜Rew.bShift (w i)⌝ : V)⌝ := by + have Hw : (L.codeIn V).IsSemitermVec ↑n (↑m + 1) ((L.codeIn V).termBShiftVec ↑n ⌜fun i ↦ ⌜w i⌝⌝) := (semitermVec_codeIn w).termBShiftVec - have HqVec : (L.codeIn V).SemitermVec (↑n + 1) (↑m + 1) ((L.codeIn V).qVec ↑n ↑m ⌜fun i ↦ ⌜w i⌝⌝) := + have HqVec : (L.codeIn V).IsSemitermVec (↑n + 1) (↑m + 1) ((L.codeIn V).qVec ⌜fun i ↦ ⌜w i⌝⌝) := (semitermVec_codeIn w).qVec apply nth_ext (by simp [←HqVec.lh]) intro i hi - have : i < ↑(n + 1) := by simpa [Language.qVec, ←Hw.lh] using hi + have : i < ↑(n + 1) := by simpa [Language.qVec, Hw.lh] using hi rcases eq_fin_of_lt_nat this with ⟨i, rfl⟩ cases' i using Fin.cases with i · simp [Language.qVec] · simp [Language.qVec, termBShift_quote] lemma substs_quote {n m} (w : Fin n → SyntacticSemiterm L m) (p : SyntacticSemiformula L n) : - (L.codeIn V).substs ↑m ⌜fun i ↦ ⌜w i⌝⌝ ⌜p⌝ = ⌜(Rew.substs w).hom p⌝ := by + (L.codeIn V).substs ⌜fun i ↦ ⌜w i⌝⌝ ⌜p⌝ = ⌜(Rew.substs w).hom p⌝ := by induction p using Semiformula.rec' generalizing m <;> simp [*, quote_rel, quote_nrel, quote_verum, quote_falsum, quote_and, quote_or, quote_all, quote_ex, Rew.rel, Rew.nrel, termSubstVec_quote, Rew.q_substs] @@ -442,77 +490,81 @@ variable {n : ℕ} namespace Semiterm -def codeIn' (t : SyntacticSemiterm L n) : (L.codeIn V).TSemiterm n := ⟨⌜t⌝, by simp⟩ +def codeIn' (t : SyntacticSemiterm L n) : (L.codeIn V).Semiterm n := ⟨⌜t⌝, by simp⟩ -instance : GoedelQuote (SyntacticSemiterm L n) ((L.codeIn V).TSemiterm n) := ⟨Semiterm.codeIn' V⟩ +instance : GoedelQuote (SyntacticSemiterm L n) ((L.codeIn V).Semiterm n) := ⟨Semiterm.codeIn' V⟩ -@[simp] lemma codeIn'_val (t : SyntacticSemiterm L n) : (⌜t⌝ : (L.codeIn V).TSemiterm n).val = ⌜t⌝ := rfl +@[simp] lemma codeIn'_val (t : SyntacticSemiterm L n) : (⌜t⌝ : (L.codeIn V).Semiterm n).val = ⌜t⌝ := rfl -def vCodeIn' {k n} (v : Fin k → SyntacticSemiterm L n) : (L.codeIn V).TSemitermVec k n := ⟨⌜fun i ↦ ⌜v i⌝⌝, by simp⟩ +def vCodeIn' {k n} (v : Fin k → SyntacticSemiterm L n) : (L.codeIn V).SemitermVec k n := ⟨⌜fun i ↦ ⌜v i⌝⌝, by simp⟩ -instance {k n} : GoedelQuote (Fin k → SyntacticSemiterm L n) ((L.codeIn V).TSemitermVec k n) := ⟨Semiterm.vCodeIn' V⟩ +instance {k n} : GoedelQuote (Fin k → SyntacticSemiterm L n) ((L.codeIn V).SemitermVec k n) := ⟨Semiterm.vCodeIn' V⟩ -@[simp] lemma vCodeIn'_val (v : Fin k → SyntacticSemiterm L n) : (⌜v⌝ : (L.codeIn V).TSemitermVec k n).val = ⌜fun i ↦ ⌜v i⌝⌝ := rfl +@[simp] lemma vCodeIn'_val (v : Fin k → SyntacticSemiterm L n) : (⌜v⌝ : (L.codeIn V).SemitermVec k n).val = ⌜fun i ↦ ⌜v i⌝⌝ := rfl -@[simp] lemma codeIn'_bvar (z : Fin n) : (⌜(#z : SyntacticSemiterm L n)⌝ : (L.codeIn V).TSemiterm n) = (L.codeIn V).bvar z := rfl -@[simp] lemma codeIn'_fvar (x : ℕ) : (⌜(&x : SyntacticSemiterm L n)⌝ : (L.codeIn V).TSemiterm n) = (L.codeIn V).fvar x := rfl +@[simp] lemma codeIn'_bvar (z : Fin n) : (⌜(#z : SyntacticSemiterm L n)⌝ : (L.codeIn V).Semiterm n) = (L.codeIn V).bvar z := by ext; simp [quote_bvar] +@[simp] lemma codeIn'_fvar (x : ℕ) : (⌜(&x : SyntacticSemiterm L n)⌝ : (L.codeIn V).Semiterm n) = (L.codeIn V).fvar x := by ext; simp [quote_fvar] lemma codeIn'_func {k} (f : L.Func k) (v : Fin k → SyntacticSemiterm L n) : - (⌜func f v⌝ : (L.codeIn V).TSemiterm n) = (L.codeIn V).func (k := k) (f := ⌜f⌝) (by simp) ⌜v⌝ := rfl - + (⌜func f v⌝ : (L.codeIn V).Semiterm n) = (L.codeIn V).func (k := k) (f := ⌜f⌝) (by simp) ⌜v⌝ := by ext; simp [quote_func, Language.func] @[simp] lemma codeIn'_zero (n : ℕ) : - (⌜(func Language.Zero.zero ![] : SyntacticSemiterm ℒₒᵣ n)⌝ : (Language.codeIn ℒₒᵣ V).TSemiterm n) = ↑(0 : V) := by ext; simp + (⌜(func Language.Zero.zero ![] : SyntacticSemiterm ℒₒᵣ n)⌝ : (Language.codeIn ℒₒᵣ V).Semiterm n) = ↑(0 : V) := by ext; simp @[simp] lemma codeIn'_one (n : ℕ) : - (⌜(func Language.One.one ![] : SyntacticSemiterm ℒₒᵣ n)⌝ : (Language.codeIn ℒₒᵣ V).TSemiterm n) = ↑(1 : V) := by ext; simp + (⌜(func Language.One.one ![] : SyntacticSemiterm ℒₒᵣ n)⌝ : (Language.codeIn ℒₒᵣ V).Semiterm n) = ↑(1 : V) := by ext; simp @[simp] lemma codeIn'_add (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜func Language.Add.add v⌝ : (Language.codeIn ℒₒᵣ V).TSemiterm n) = ⌜v 0⌝ + ⌜v 1⌝ := by ext; simp; rfl + (⌜func Language.Add.add v⌝ : (Language.codeIn ℒₒᵣ V).Semiterm n) = ⌜v 0⌝ + ⌜v 1⌝ := by ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [quote_add] @[simp] lemma codeIn'_mul (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜func Language.Mul.mul v⌝ : (Language.codeIn ℒₒᵣ V).TSemiterm n) = ⌜v 0⌝ * ⌜v 1⌝ := by ext; simp; rfl + (⌜func Language.Mul.mul v⌝ : (Language.codeIn ℒₒᵣ V).Semiterm n) = ⌜v 0⌝ * ⌜v 1⌝ := by ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [quote_add] end Semiterm namespace Semiformula -def codeIn' (p : SyntacticSemiformula L n) : (L.codeIn V).TSemiformula n := ⟨⌜p⌝, by simp⟩ +def codeIn' (p : SyntacticSemiformula L n) : (L.codeIn V).Semiformula n := ⟨⌜p⌝, by simp⟩ -instance : GoedelQuote (SyntacticSemiformula L n) ((L.codeIn V).TSemiformula n) := ⟨Semiformula.codeIn' V⟩ +instance : GoedelQuote (SyntacticSemiformula L n) ((L.codeIn V).Semiformula n) := ⟨Semiformula.codeIn' V⟩ -@[simp] lemma codeIn'_val (p : SyntacticSemiformula L n) : (⌜p⌝ : (L.codeIn V).TSemiformula n).val = ⌜p⌝ := rfl +@[simp] lemma codeIn'_val (p : SyntacticSemiformula L n) : (⌜p⌝ : (L.codeIn V).Semiformula n).val = ⌜p⌝ := rfl -@[simp] lemma codeIn'_verum (n : ℕ) : (⌜(⊤ : SyntacticSemiformula L n)⌝ : (L.codeIn V).TSemiformula n) = ⊤ := rfl -@[simp] lemma codeIn'_falsum (n : ℕ) : (⌜(⊥ : SyntacticSemiformula L n)⌝ : (L.codeIn V).TSemiformula n) = ⊥ := rfl -@[simp] lemma codeIn'_and (p q : SyntacticSemiformula L n) : (⌜p ⋏ q⌝ : (L.codeIn V).TSemiformula n) = ⌜p⌝ ⋏ ⌜q⌝ := rfl -@[simp] lemma codeIn'_or (p q : SyntacticSemiformula L n) : (⌜p ⋎ q⌝ : (L.codeIn V).TSemiformula n) = ⌜p⌝ ⋎ ⌜q⌝ := rfl -@[simp] lemma codeIn'_all (p : SyntacticSemiformula L (n + 1)) : (⌜∀' p⌝ : (L.codeIn V).TSemiformula n) = .all (.cast (n := ↑(n + 1)) ⌜p⌝) := rfl -@[simp] lemma codeIn'_ex (p : SyntacticSemiformula L (n + 1)) : (⌜∃' p⌝ : (L.codeIn V).TSemiformula n) = .ex (.cast (n := ↑(n + 1)) ⌜p⌝) := rfl -@[simp] lemma codeIn'_neg (p : SyntacticSemiformula L n) : (⌜~p⌝ : (L.codeIn V).TSemiformula n) = ~⌜p⌝ := by +@[simp] lemma codeIn'_verum (n : ℕ) : (⌜(⊤ : SyntacticSemiformula L n)⌝ : (L.codeIn V).Semiformula n) = ⊤ := by ext; simp [quote_verum] +@[simp] lemma codeIn'_falsum (n : ℕ) : (⌜(⊥ : SyntacticSemiformula L n)⌝ : (L.codeIn V).Semiformula n) = ⊥ := by ext; simp [quote_falsum] +@[simp] lemma codeIn'_and (p q : SyntacticSemiformula L n) : (⌜p ⋏ q⌝ : (L.codeIn V).Semiformula n) = ⌜p⌝ ⋏ ⌜q⌝ := by ext; simp [quote_and] +@[simp] lemma codeIn'_or (p q : SyntacticSemiformula L n) : (⌜p ⋎ q⌝ : (L.codeIn V).Semiformula n) = ⌜p⌝ ⋎ ⌜q⌝ := by ext; simp [quote_or] +@[simp] lemma codeIn'_all (p : SyntacticSemiformula L (n + 1)) : (⌜∀' p⌝ : (L.codeIn V).Semiformula n) = .all (.cast (n := ↑(n + 1)) ⌜p⌝) := by ext; simp [quote_all] +@[simp] lemma codeIn'_ex (p : SyntacticSemiformula L (n + 1)) : (⌜∃' p⌝ : (L.codeIn V).Semiformula n) = .ex (.cast (n := ↑(n + 1)) ⌜p⌝) := by ext; simp [quote_ex] +@[simp] lemma codeIn'_neg (p : SyntacticSemiformula L n) : (⌜~p⌝ : (L.codeIn V).Semiformula n) = ~⌜p⌝ := by ext; simp [neg_quote] -@[simp] lemma codeIn'_imp (p q : SyntacticSemiformula L n) : (⌜p ⟶ q⌝ : (L.codeIn V).TSemiformula n) = ⌜p⌝ ⟶ ⌜q⌝ := by - simp [Semiformula.imp_eq, Language.TSemiformula.imp_def] +@[simp] lemma codeIn'_imp (p q : SyntacticSemiformula L n) : (⌜p ⟶ q⌝ : (L.codeIn V).Semiformula n) = ⌜p⌝ ⟶ ⌜q⌝ := by + simp [Semiformula.imp_eq, Language.Semiformula.imp_def] open LO.Arith Formalized @[simp] lemma codeIn'_eq (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜rel Language.Eq.eq v⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = (⌜v 0⌝ =' ⌜v 1⌝) := by ext; simp [Language.TSemiterm.equals]; rfl + (⌜rel Language.Eq.eq v⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = (⌜v 0⌝ =' ⌜v 1⌝) := by + ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [Language.Semiterm.equals] @[simp] lemma codeIn'_neq (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜nrel Language.Eq.eq v⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = (⌜v 0⌝ ≠' ⌜v 1⌝) := by ext; simp [Language.TSemiterm.notEquals]; rfl + (⌜nrel Language.Eq.eq v⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = (⌜v 0⌝ ≠' ⌜v 1⌝) := by + ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [Language.Semiterm.notEquals] @[simp] lemma codeIn'_lt (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜rel Language.LT.lt v⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = (⌜v 0⌝ <' ⌜v 1⌝) := by ext; simp [Language.TSemiterm.lessThan]; rfl + (⌜rel Language.LT.lt v⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = (⌜v 0⌝ <' ⌜v 1⌝) := by + ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [Language.Semiterm.lessThan] @[simp] lemma codeIn'_nlt (v : Fin 2 → SyntacticSemiterm ℒₒᵣ n) : - (⌜nrel Language.LT.lt v⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = (⌜v 0⌝ ≮' ⌜v 1⌝) := by ext; simp [Language.TSemiterm.notLessThan]; rfl + (⌜nrel Language.LT.lt v⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = (⌜v 0⌝ ≮' ⌜v 1⌝) := by + ext; rw [Matrix.fun_eq_vec₂ (v := v)]; simp [Language.Semiterm.notLessThan] @[simp] lemma codeIn'_ball (t : SyntacticSemiterm ℒₒᵣ n) (p : SyntacticSemiformula ℒₒᵣ (n + 1)) : - (⌜∀[“#0 < !!(Rew.bShift t)”] p⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = Language.TSemiformula.ball ⌜t⌝ (.cast (n := ↑(n + 1)) ⌜p⌝) := by - ext; simp [LogicalConnective.ball, imp_eq, Language.TSemiformula.cast, - Language.TSemiformula.ball, Semiformula.Operator.lt_eq, termBShift_quote] + (⌜∀[“#0 < !!(Rew.bShift t)”] p⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = Language.Semiformula.ball ⌜t⌝ (.cast (n := ↑(n + 1)) ⌜p⌝) := by + ext; simp [LogicalConnective.ball, imp_eq, Language.Semiformula.cast, + Language.Semiformula.ball, Semiformula.Operator.lt_eq, termBShift_quote] @[simp] lemma codeIn'_bex (t : SyntacticSemiterm ℒₒᵣ n) (p : SyntacticSemiformula ℒₒᵣ (n + 1)) : - (⌜∃[“#0 < !!(Rew.bShift t)”] p⌝ : (Language.codeIn ℒₒᵣ V).TSemiformula n) = Language.TSemiformula.bex ⌜t⌝ (.cast (n := ↑(n + 1)) ⌜p⌝) := by - ext; simp [LogicalConnective.bex, imp_eq, Language.TSemiformula.cast, - Language.TSemiformula.ball, Semiformula.Operator.lt_eq, termBShift_quote] + (⌜∃[“#0 < !!(Rew.bShift t)”] p⌝ : (Language.codeIn ℒₒᵣ V).Semiformula n) = Language.Semiformula.bex ⌜t⌝ (.cast (n := ↑(n + 1)) ⌜p⌝) := by + ext; simp [LogicalConnective.bex, imp_eq, Language.Semiformula.cast, + Language.Semiformula.ball, Semiformula.Operator.lt_eq, termBShift_quote] -instance : GoedelQuote (Sentence L) ((L.codeIn V).TFormula) := ⟨fun σ ↦ (⌜Rew.embs.hom σ⌝ : (Language.codeIn L V).TSemiformula (0 : ℕ))⟩ +instance : GoedelQuote (Sentence L) ((L.codeIn V).Formula) := ⟨fun σ ↦ (⌜Rew.embs.hom σ⌝ : (Language.codeIn L V).Semiformula (0 : ℕ))⟩ -lemma quote_sentence_def' (σ : Sentence L) : (⌜σ⌝ : (L.codeIn V).TFormula) = (⌜Rew.embs.hom σ⌝ : (Language.codeIn L V).TSemiformula (0 : ℕ)) := rfl +lemma quote_sentence_def' (σ : Sentence L) : (⌜σ⌝ : (L.codeIn V).Formula) = (⌜Rew.embs.hom σ⌝ : (Language.codeIn L V).Semiformula (0 : ℕ)) := rfl -@[simp] lemma quote_sentence_val (σ : Sentence L) : (⌜σ⌝ : (L.codeIn V).TFormula).val = ⌜σ⌝ := rfl +@[simp] lemma quote_sentence_val (σ : Sentence L) : (⌜σ⌝ : (L.codeIn V).Formula).val = ⌜σ⌝ := by + simp [quote_sentence_def', quote_eq_coe_encode] end Semiformula diff --git a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean index 0f6d17c..80b615f 100644 --- a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean +++ b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean @@ -98,10 +98,10 @@ variable {T : SyntacticTheory L} [T.Δ₁Definable] open Classical -@[simp] lemma formulaSet_codeIn_finset (Γ : Finset (SyntacticFormula L)) : (L.codeIn V).FormulaSet ⌜Γ⌝ := by +@[simp] lemma formulaSet_codeIn_finset (Γ : Finset (SyntacticFormula L)) : (L.codeIn V).IsFormulaSet ⌜Γ⌝ := by intro x hx rcases Derivation3.Sequent.mem_codeIn hx with ⟨p, _, rfl⟩; - apply semiformula_quote + apply semiformula_quote (n := 0) open Derivation3 @@ -125,20 +125,20 @@ lemma quote_image_shift (Γ : Finset (SyntacticFormula L)) : (L.codeIn V).setShi (by simpa [quote_verum] using (Sequent.mem_codeIn_iff (V := V)).mpr h) case and Δ p q hpq dp dq ihp ihq => apply Language.Theory.Derivation.andIntro - (by simpa using (Sequent.mem_codeIn_iff (V := V)).mpr hpq) + (by simpa [quote_and] using (Sequent.mem_codeIn_iff (V := V)).mpr hpq) ⟨by simp [fstidx_quote], ihp⟩ ⟨by simp [fstidx_quote], ihq⟩ case or Δ p q hpq d ih => apply Language.Theory.Derivation.orIntro - (by simpa using (Sequent.mem_codeIn_iff (V := V)).mpr hpq) + (by simpa [quote_or] using (Sequent.mem_codeIn_iff (V := V)).mpr hpq) ⟨by simp [fstidx_quote], ih⟩ case all Δ p h d ih => apply Language.Theory.Derivation.allIntro - (by simpa using (Sequent.mem_codeIn_iff (V := V)).mpr h) + (by simpa [quote_all] using (Sequent.mem_codeIn_iff (V := V)).mpr h) ⟨by simp [fstidx_quote, quote_image_shift, free_quote], ih⟩ case ex Δ p h t d ih => apply Language.Theory.Derivation.exIntro - (by simpa using (Sequent.mem_codeIn_iff (V := V)).mpr h) + (by simpa [quote_ex] using (Sequent.mem_codeIn_iff (V := V)).mpr h) (semiterm_codeIn t) ⟨by simp [fstidx_quote, ←substs_quote, Language.substs₁], ih⟩ case wk Δ Γ d h ih => @@ -167,7 +167,7 @@ section variable [L.ConstantInhabited] {T : Theory L} [T.Δ₁Definable] theorem provable_of_provable : T ⊢! p → (T.codeIn V).Provable ⌜p⌝ := fun h ↦ by - simpa [quote_semisentence_def] using derivable_of_quote (V := V) (provable_iff_derivable3'.mp h).some + simpa using derivable_of_quote (V := V) (provable_iff_derivable3'.mp h).some /-- Hilbert–Bernays provability condition D1 -/ theorem tprovable_of_provable : T ⊢! σ → T.tCodeIn V ⊢! ⌜σ⌝ := fun h ↦ by diff --git a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean index 069ecf3..69cfefd 100644 --- a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean +++ b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean @@ -21,11 +21,12 @@ namespace Formalized variable {T : LOR.TTheory (V := V)} [R₀Theory T] -def toNumVec {n} (e : Fin n → V) : (Language.codeIn ℒₒᵣ V).TSemitermVec n 0 := - ⟨⌜fun i ↦ numeral (e i)⌝, by simp, by +def toNumVec {n} (e : Fin n → V) : (Language.codeIn ℒₒᵣ V).SemitermVec n 0 := + ⟨⌜fun i ↦ numeral (e i)⌝, + Language.IsSemitermVec.iff.mpr <| ⟨by simp, by intro i hi rcases eq_fin_of_lt_nat hi with ⟨i, rfl⟩ - simp [quote_nth_fin (fun i ↦ numeral (e i)) i]⟩ + simp [quote_nth_fin (fun i ↦ numeral (e i)) i]⟩⟩ @[simp] lemma toNumVec_nil : (toNumVec (![] : Fin 0 → V)) = .nil _ _ := by ext; simp [toNumVec] @@ -38,26 +39,29 @@ def toNumVec {n} (e : Fin n → V) : (Language.codeIn ℒₒᵣ V).TSemitermVec calc (i : V) < (i : V) + (n - i : V) := by simp _ = (n : V) := by simp +@[simp] lemma len_semitermvec {L : Arith.Language V} {pL} [L.Defined pL] (v : L.SemitermVec k n) : len v.val = k := v.prop.lh + @[simp] lemma cast_substs_numVec (p : Semisentence ℒₒᵣ (n + 1)) : ((.cast (V := V) (n := ↑(n + 1)) (n' := ↑n + 1) ⌜Rew.embs.hom p⌝ (by simp)) ^/[(toNumVec e).q.substs (typedNumeral 0 x).sing]) = ⌜Rew.embs.hom p⌝ ^/[toNumVec (x :> e)] := by have : (toNumVec e).q.substs (typedNumeral 0 x).sing = x ∷ᵗ toNumVec e := by ext; simp apply nth_ext' ((↑n : V) + 1) - (by rw [len_termSubstVec]; simpa using (toNumVec e).prop.qVec) - (by simp [←(toNumVec e).prop.1]) + (by rw [len_termSubstVec]; simpa using (toNumVec e).prop.qVec.isUTerm) + (by simp [(toNumVec e).prop.lh]) intro i hi - rw [nth_termSubstVec (by simpa using (toNumVec e).prop.qVec) hi] + rw [nth_termSubstVec (by simpa using (toNumVec e).prop.qVec.isUTerm) hi] rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp [Language.qVec] - · simp only [Language.qVec, nth_cons_succ, Language.TSemitermVec.prop] + · simp only [Language.qVec, nth_cons_succ, Language.SemitermVec.prop] rcases eq_fin_of_lt_nat (by simpa using hi) with ⟨i, rfl⟩ - rw [nth_termBShiftVec (by simp)] - simp; exact coe_coe_lt (V := V) i + rw [nth_termBShiftVec (by simp), + toNumVec_val_nth, numeral_bShift, + numeral_substs (n := 1) (m := 0) (by simp)] + simp rw [this] ext; simp [toNumVec] - namespace TProof open Language.Theory.TProof System @@ -75,13 +79,13 @@ noncomputable def termEqComplete {n : ℕ} (e : Fin n → V) : Structure.add_eq_of_lang] have ih : T ⊢ (⌜Rew.embs (v 0)⌝^ᵗ/[toNumVec e] + ⌜Rew.embs (v 1)⌝^ᵗ/[toNumVec e]) =' (↑((v 0).valbm V e) + ↑((v 1).valbm V e)) := addExt T _ _ _ _ ⨀ termEqComplete e (v 0) ⨀ termEqComplete e (v 1) - have : T ⊢ ((v 0).valbm V e + (v 1).valbm V e : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑((v 0).valbm V e + (v 1).valbm V e) := addComplete T _ _ + have : T ⊢ ((v 0).valbm V e + (v 1).valbm V e : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑((v 0).valbm V e + (v 1).valbm V e) := addComplete T _ _ exact eqTrans T _ _ _ ⨀ ih ⨀ this | Semiterm.func Language.Mul.mul v => by simp [Rew.func, Semiterm.val_func] have ih : T ⊢ (⌜Rew.embs (v 0)⌝^ᵗ/[toNumVec e] * ⌜Rew.embs (v 1)⌝^ᵗ/[toNumVec e]) =' (↑((v 0).valbm V e) * ↑((v 1).valbm V e)) := mulExt T _ _ _ _ ⨀ termEqComplete e (v 0) ⨀ termEqComplete e (v 1) - have : T ⊢ ((v 0).valbm V e * (v 1).valbm V e : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑((v 0).valbm V e * (v 1).valbm V e) := mulComplete T _ _ + have : T ⊢ ((v 0).valbm V e * (v 1).valbm V e : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑((v 0).valbm V e * (v 1).valbm V e) := mulComplete T _ _ exact eqTrans T _ _ _ ⨀ ih ⨀ this lemma termEq_complete! {n : ℕ} (e : Fin n → V) (t : Semiterm ℒₒᵣ Empty n) : @@ -97,7 +101,7 @@ theorem bold_sigma₁_complete {n} {p : Semisentence ℒₒᵣ n} (hp : Hierarch case hFalsum => intro n simp only [LogicalConnective.HomClass.map_bot, Prop.bot_eq_false, - Semiformula.codeIn'_falsum, Language.TSemiformula.substs_falsum, false_implies, implies_true] + Semiformula.codeIn'_falsum, Language.Semiformula.substs_falsum, false_implies, implies_true] case hEQ => intro n t₁ t₂ e h have : t₁.valbm V e = t₂.valbm V e := by simpa using h @@ -122,15 +126,15 @@ theorem bold_sigma₁_complete {n} {p : Semisentence ℒₒᵣ n} (hp : Hierarch intro n p q hp hq ihp ihq e h have h : Semiformula.Evalbm V e p ∧ Semiformula.Evalbm V e q := by simpa using h simpa only [LogicalConnective.HomClass.map_and, Semiformula.codeIn'_and, - Language.TSemiformula.substs_and] using and_intro! (ihp h.1) (ihq h.2) + Language.Semiformula.substs_and] using and_intro! (ihp h.1) (ihq h.2) case hOr => intro n p q hp hq ihp ihq e h have : Semiformula.Evalbm V e p ∨ Semiformula.Evalbm V e q := by simpa using h rcases this with (h | h) · simpa only [LogicalConnective.HomClass.map_or, Semiformula.codeIn'_or, - Language.TSemiformula.substs_or] using or₁'! (ihp h) + Language.Semiformula.substs_or] using or₁'! (ihp h) · simpa only [LogicalConnective.HomClass.map_or, Semiformula.codeIn'_or, - Language.TSemiformula.substs_or] using or₂'! (ihq h) + Language.Semiformula.substs_or] using or₂'! (ihq h) case hBall => intro n t p hp ihp e h simp only [Rew.ball, Rew.q_emb, Rew.hom_finitary2, Rew.emb_bvar, ← Rew.emb_bShift_term, @@ -145,7 +149,7 @@ theorem bold_sigma₁_complete {n} {p : Semisentence ℒₒᵣ n} (hp : Hierarch exact ihp this case hEx => intro n p hp ihp e h - simp only [Rew.ex, Rew.q_emb, Semiformula.codeIn'_ex, Language.TSemiformula.substs_ex] + simp only [Rew.ex, Rew.q_emb, Semiformula.codeIn'_ex, Language.Semiformula.substs_ex] have : ∃ x, Semiformula.Evalbm V (x :> e) p := by simpa using h rcases this with ⟨x, hx⟩ apply ex! x diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index 61462ec..2dd0ef0 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -10,333 +10,299 @@ variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] -def qqRel (n k r v : V) : V := ⟪n, 0, k, r, v⟫ + 1 +def qqRel (k r v : V) : V := ⟪0, k, r, v⟫ + 1 -def qqNRel (n k r v : V) : V := ⟪n, 1, k, r, v⟫ + 1 +def qqNRel (k r v : V) : V := ⟪1, k, r, v⟫ + 1 -def qqVerum (n : V) : V := ⟪n, 2, 0⟫ + 1 +def qqVerum : V := ⟪2, 0⟫ + 1 -def qqFalsum (n : V) : V := ⟪n, 3, 0⟫ + 1 +def qqFalsum : V := ⟪3, 0⟫ + 1 -def qqAnd (n p q : V) : V := ⟪n, 4, p, q⟫ + 1 +def qqAnd (p q : V) : V := ⟪4, p, q⟫ + 1 -def qqOr (n p q : V) : V := ⟪n, 5, p, q⟫ + 1 +def qqOr (p q : V) : V := ⟪5, p, q⟫ + 1 -def qqAll (n p : V) : V := ⟪n, 6, p⟫ + 1 +def qqAll (p : V) : V := ⟪6, p⟫ + 1 -def qqEx (n p : V) : V := ⟪n, 7, p⟫ + 1 +def qqEx (p : V) : V := ⟪7, p⟫ + 1 scoped prefix:max "^rel " => qqRel scoped prefix:max "^nrel " => qqNRel -scoped notation "^⊤[" n "]" => qqVerum n +scoped notation "^⊤" => qqVerum -scoped notation "^⊥[" n "]" => qqFalsum n +scoped notation "^⊥" => qqFalsum -scoped notation p:69 " ^⋏[" n "] " q:70 => qqAnd n p q +scoped notation p:69 " ^⋏ " q:70 => qqAnd p q -scoped notation p:68 " ^⋎[" n "] " q:69 => qqOr n p q +scoped notation p:68 " ^⋎ " q:69 => qqOr p q -scoped notation "^∀[" n "] " p:64 => qqAll n p +scoped notation "^∀ " p:64 => qqAll p -scoped notation "^∃[" n "] " p:64 => qqEx n p - -scoped notation "^⊤" => qqVerum 0 - -scoped notation "^⊥" => qqFalsum 0 - -scoped notation p:69 " ^⋏ " q:70 => qqAnd 0 p q - -scoped notation p:68 " ^⋎ " q:69 => qqOr 0 p q - -scoped notation "^∀ " p:64 => qqAll 0 p - -scoped notation "^∃ " p:64 => qqEx 0 p +scoped notation "^∃ " p:64 => qqEx p section -def _root_.LO.FirstOrder.Arith.qqRelDef : 𝚺₀.Semisentence 5 := - .mkSigma “p n k r v | ∃ p' < p, !pair₅Def p' n 0 k r v ∧ p = p' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqRelDef : 𝚺₀.Semisentence 4 := + .mkSigma “p k r v | ∃ p' < p, !pair₄Def p' 0 k r v ∧ p = p' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqNRelDef : 𝚺₀.Semisentence 5 := - .mkSigma “p n k r v | ∃ p' < p, !pair₅Def p' n 1 k r v ∧ p = p' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqNRelDef : 𝚺₀.Semisentence 4 := + .mkSigma “p k r v | ∃ p' < p, !pair₄Def p' 1 k r v ∧ p = p' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqVerumDef : 𝚺₀.Semisentence 2 := - .mkSigma “p n | ∃ p' < p, !pair₃Def p' n 2 0 ∧ p = p' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqVerumDef : 𝚺₀.Semisentence 1 := + .mkSigma “p | ∃ p' < p, !pairDef p' 2 0 ∧ p = p' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqFalsumDef : 𝚺₀.Semisentence 2 := - .mkSigma “p n | ∃ p' < p, !pair₃Def p' n 3 0 ∧ p = p' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqFalsumDef : 𝚺₀.Semisentence 1 := + .mkSigma “p | ∃ p' < p, !pairDef p' 3 0 ∧ p = p' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqAndDef : 𝚺₀.Semisentence 4 := - .mkSigma “r n p q | ∃ r' < r, !pair₄Def r' n 4 p q ∧ r = r' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqAndDef : 𝚺₀.Semisentence 3 := + .mkSigma “r p q | ∃ r' < r, !pair₃Def r' 4 p q ∧ r = r' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqOrDef : 𝚺₀.Semisentence 4 := - .mkSigma “r n p q | ∃ r' < r, !pair₄Def r' n 5 p q ∧ r = r' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqOrDef : 𝚺₀.Semisentence 3 := + .mkSigma “r p q | ∃ r' < r, !pair₃Def r' 5 p q ∧ r = r' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqAllDef : 𝚺₀.Semisentence 3 := - .mkSigma “r n p | ∃ r' < r, !pair₃Def r' n 6 p ∧ r = r' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqAllDef : 𝚺₀.Semisentence 2 := + .mkSigma “r p | ∃ r' < r, !pairDef r' 6 p ∧ r = r' + 1” (by simp) -def _root_.LO.FirstOrder.Arith.qqExDef : 𝚺₀.Semisentence 3 := - .mkSigma “r n p | ∃ r' < r, !pair₃Def r' n 7 p ∧ r = r' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqExDef : 𝚺₀.Semisentence 2 := + .mkSigma “r p | ∃ r' < r, !pairDef r' 7 p ∧ r = r' + 1” (by simp) -lemma ss (x : V) : x < x + 1 := by exact lt_add_one x - -lemma qqRel_defined : 𝚺₀-Function₄ (qqRel : V → V → V → V → V) via qqRelDef := by +lemma qqRel_defined : 𝚺₀-Function₃ (qqRel : V → V → V → V) via qqRelDef := by intro v; simp [qqRelDef] constructor - · intro h; exact ⟨_, by simpa [h, qqRel] using lt_add_one _, rfl, h⟩ + · intro h; exact ⟨_, by simp [h, qqRel], rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqNRel_defined : 𝚺₀-Function₄ (qqNRel : V → V → V → V → V) via qqNRelDef := by +lemma qqNRel_defined : 𝚺₀-Function₃ (qqNRel : V → V → V → V) via qqNRelDef := by intro v; simp [qqNRelDef] constructor · intro h; exact ⟨_, by simpa [h, qqRel] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqVerum_defined : 𝚺₀-Function₁ (qqVerum : V → V) via qqVerumDef := by +lemma qqVerum_defined : 𝚺₀-Function₀ (qqVerum : V) via qqVerumDef := by intro v; simp [qqVerumDef] constructor · intro h; exact ⟨_, by simpa [h, qqRel] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqFalsum_defined : 𝚺₀-Function₁ (qqFalsum : V → V) via qqFalsumDef := by +lemma qqFalsum_defined : 𝚺₀-Function₀ (qqFalsum : V) via qqFalsumDef := by intro v; simp [qqFalsumDef] constructor · intro h; exact ⟨_, by simpa [h, qqRel] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqAnd_defined : 𝚺₀-Function₃ (qqAnd : V → V → V → V) via qqAndDef := by +lemma qqAnd_defined : 𝚺₀-Function₂ (qqAnd : V → V → V) via qqAndDef := by intro v; simp [qqAndDef] constructor · intro h; exact ⟨_, by simpa [h, qqRel] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqOr_defined : 𝚺₀-Function₃ (qqOr : V → V → V → V) via qqOrDef := by +lemma qqOr_defined : 𝚺₀-Function₂ (qqOr : V → V → V) via qqOrDef := by intro v; simp [qqOrDef, numeral_eq_natCast] constructor · intro h; exact ⟨_, by simpa [h] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqForall_defined : 𝚺₀-Function₂ (qqAll : V → V → V) via qqAllDef := by +lemma qqForall_defined : 𝚺₀-Function₁ (qqAll : V → V) via qqAllDef := by intro v; simp [qqAllDef, numeral_eq_natCast] constructor · intro h; exact ⟨_, by simpa [h] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma qqExists_defined : 𝚺₀-Function₂ (qqEx : V → V → V) via qqExDef := by +lemma qqExists_defined : 𝚺₀-Function₁ (qqEx : V → V) via qqExDef := by intro v; simp [qqExDef, numeral_eq_natCast] constructor · intro h; exact ⟨_, by simpa [h] using lt_add_one _, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h @[simp] lemma eval_qqRelDef (v) : - Semiformula.Evalbm V v qqRelDef.val ↔ v 0 = ^rel (v 1) (v 2) (v 3) (v 4) := qqRel_defined.df.iff v + Semiformula.Evalbm V v qqRelDef.val ↔ v 0 = ^rel (v 1) (v 2) (v 3) := qqRel_defined.df.iff v @[simp] lemma eval_qqNRelDef (v) : - Semiformula.Evalbm V v qqNRelDef.val ↔ v 0 = ^nrel (v 1) (v 2) (v 3) (v 4) := qqNRel_defined.df.iff v + Semiformula.Evalbm V v qqNRelDef.val ↔ v 0 = ^nrel (v 1) (v 2) (v 3) := qqNRel_defined.df.iff v @[simp] lemma eval_qqVerumDef (v) : - Semiformula.Evalbm V v qqVerumDef.val ↔ v 0 = ^⊤[v 1] := qqVerum_defined.df.iff v + Semiformula.Evalbm V v qqVerumDef.val ↔ v 0 = ^⊤ := qqVerum_defined.df.iff v @[simp] lemma eval_qqFalsumDef (v) : - Semiformula.Evalbm V v qqFalsumDef.val ↔ v 0 = ^⊥[v 1] := qqFalsum_defined.df.iff v + Semiformula.Evalbm V v qqFalsumDef.val ↔ v 0 = ^⊥ := qqFalsum_defined.df.iff v @[simp] lemma eval_qqAndDef (v) : - Semiformula.Evalbm V v qqAndDef.val ↔ v 0 = (v 2) ^⋏[v 1] (v 3) := qqAnd_defined.df.iff v + Semiformula.Evalbm V v qqAndDef.val ↔ v 0 = (v 1) ^⋏ (v 2) := qqAnd_defined.df.iff v @[simp] lemma eval_qqOrDef (v) : - Semiformula.Evalbm V v qqOrDef.val ↔ v 0 = (v 2) ^⋎[v 1] (v 3) := qqOr_defined.df.iff v + Semiformula.Evalbm V v qqOrDef.val ↔ v 0 = (v 1) ^⋎ (v 2) := qqOr_defined.df.iff v @[simp] lemma eval_qqAllDef (v) : - Semiformula.Evalbm V v qqAllDef.val ↔ v 0 = ^∀[v 1] (v 2) := qqForall_defined.df.iff v + Semiformula.Evalbm V v qqAllDef.val ↔ v 0 = ^∀ (v 1) := qqForall_defined.df.iff v @[simp] lemma eval_qqExDef (v) : - Semiformula.Evalbm V v qqExDef.val ↔ v 0 = ^∃[v 1] (v 2) := qqExists_defined.df.iff v + Semiformula.Evalbm V v qqExDef.val ↔ v 0 = ^∃ (v 1) := qqExists_defined.df.iff v -instance (ℌ : HierarchySymbol) : ℌ-Function₄ (qqRel : V → V → V → V → V) := .of_zero qqRel_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₃ (qqRel : V → V → V → V) := .of_zero qqRel_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₄ (qqNRel : V → V → V → V → V) := .of_zero qqNRel_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₃ (qqNRel : V → V → V → V) := .of_zero qqNRel_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₁ (qqVerum : V → V) := .of_zero qqVerum_defined.to_definable +-- instance (ℌ : HierarchySymbol) : ℌ-Function₀ (qqVerum : V) := .of_zero qqVerum_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₁ (qqFalsum : V → V) := .of_zero qqFalsum_defined.to_definable +-- instance (ℌ : HierarchySymbol) : ℌ-Function₁ (qqFalsum : V → V) := .of_zero qqFalsum_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₃ (qqAnd : V → V → V → V) := .of_zero qqAnd_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₂ (qqAnd : V → V → V) := .of_zero qqAnd_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₃ (qqOr : V → V → V → V) := .of_zero qqOr_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₂ (qqOr : V → V → V) := .of_zero qqOr_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₂ (qqAll : V → V → V) := .of_zero qqForall_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₁ (qqAll : V → V) := .of_zero qqForall_defined.to_definable -instance (ℌ : HierarchySymbol) : ℌ-Function₂ (qqEx : V → V → V) := .of_zero qqExists_defined.to_definable +instance (ℌ : HierarchySymbol) : ℌ-Function₁ (qqEx : V → V) := .of_zero qqExists_defined.to_definable end -@[simp] lemma qqRel_inj (n₁ k₁ r₁ v₁ n₂ k₂ r₂ v₂ : V) : - ^rel n₁ k₁ r₁ v₁ = ^rel n₂ k₂ r₂ v₂ ↔ n₁ = n₂ ∧ k₁ = k₂ ∧ r₁ = r₂ ∧ v₁ = v₂ := by simp [qqRel] -@[simp] lemma qqNRel_inj (n₁ k₁ r₁ v₁ n₂ k₂ r₂ v₂ : V) : - ^nrel n₁ k₁ r₁ v₁ = ^nrel n₂ k₂ r₂ v₂ ↔ n₁ = n₂ ∧ k₁ = k₂ ∧ r₁ = r₂ ∧ v₁ = v₂ := by simp [qqNRel] -@[simp] lemma qqVerum_inj (n₁ n₂ : V) : ^⊤[n₁] = ^⊤[n₂] ↔ n₁ = n₂ := by simp [qqVerum] -@[simp] lemma qqFalsum_inj (n₁ n₂ : V) : ^⊥[n₁] = ^⊥[n₂] ↔ n₁ = n₂ := by simp [qqFalsum] -@[simp] lemma qqAnd_inj (n₁ p₁ q₁ n₂ p₂ q₂ : V) : p₁ ^⋏[n₁] q₁ = p₂ ^⋏[n₂] q₂ ↔ n₁ = n₂ ∧ p₁ = p₂ ∧ q₁ = q₂ := by simp [qqAnd] -@[simp] lemma qqOr_inj (n₁ p₁ q₁ n₂ p₂ q₂ : V) : p₁ ^⋎[n₁] q₁ = p₂ ^⋎[n₂] q₂ ↔ n₁ = n₂ ∧ p₁ = p₂ ∧ q₁ = q₂ := by simp [qqOr] -@[simp] lemma qqAll_inj (n₁ p₁ n₂ p₂ : V) : ^∀[n₁] p₁ = ^∀[n₂] p₂ ↔ n₁ = n₂ ∧ p₁ = p₂ := by simp [qqAll] -@[simp] lemma qqEx_inj (n₁ p₁ n₂ p₂ : V) : ^∃[n₁] p₁ = ^∃[n₂] p₂ ↔ n₁ = n₂ ∧ p₁ = p₂ := by simp [qqEx] - -@[simp] lemma fstIdx_lt_rel (n k r v : V) : n < ^rel n k r v := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma arity_lt_rel (n k r v : V) : k < ^rel n k r v := +@[simp] lemma qqRel_inj (k₁ r₁ v₁ k₂ r₂ v₂ : V) : + ^rel k₁ r₁ v₁ = ^rel k₂ r₂ v₂ ↔ k₁ = k₂ ∧ r₁ = r₂ ∧ v₁ = v₂ := by simp [qqRel] +@[simp] lemma qqNRel_inj (k₁ r₁ v₁ k₂ r₂ v₂ : V) : + ^nrel k₁ r₁ v₁ = ^nrel k₂ r₂ v₂ ↔ k₁ = k₂ ∧ r₁ = r₂ ∧ v₁ = v₂ := by simp [qqNRel] +@[simp] lemma qqAnd_inj (p₁ q₁ p₂ q₂ : V) : p₁ ^⋏ q₁ = p₂ ^⋏ q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := by simp [qqAnd] +@[simp] lemma qqOr_inj (p₁ q₁ p₂ q₂ : V) : p₁ ^⋎ q₁ = p₂ ^⋎ q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := by simp [qqOr] +@[simp] lemma qqAll_inj (p₁ p₂ : V) : ^∀ p₁ = ^∀ p₂ ↔ p₁ = p₂ := by simp [qqAll] +@[simp] lemma qqEx_inj (p₁ p₂ : V) : ^∃ p₁ = ^∃ p₂ ↔ p₁ = p₂ := by simp [qqEx] + +@[simp] lemma arity_lt_rel (k r v : V) : k < ^rel k r v := + le_iff_lt_succ.mp <| le_trans (le_pair_left k ⟪r, v⟫) <| le_pair_right _ _ +@[simp] lemma r_lt_rel (k r v : V) : r < ^rel k r v := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma r_lt_rel (n k r v : V) : r < ^rel n k r v := - le_iff_lt_succ.mp <| le_trans (le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma v_lt_rel (n k r v : V) : v < ^rel n k r v := - le_iff_lt_succ.mp <| le_trans (le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ +@[simp] lemma v_lt_rel (k r v : V) : v < ^rel k r v := + le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma fstIdx_lt_nrel (n k r v : V) : n < ^nrel n k r v := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma arity_lt_nrel (n k r v : V) : k < ^nrel n k r v := +@[simp] lemma arity_lt_nrel (k r v : V) : k < ^nrel k r v := le_iff_lt_succ.mp <| le_trans (le_pair_left _ _) <| le_pair_right _ _ +@[simp] lemma r_lt_nrel (k r v : V) : r < ^nrel k r v := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma r_lt_nrel (n k r v : V) : r < ^nrel n k r v := - le_iff_lt_succ.mp <| le_trans (le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma v_lt_nrel (n k r v : V) : v < ^nrel n k r v := - le_iff_lt_succ.mp <| le_trans (le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ +@[simp] lemma v_lt_nrel (k r v : V) : v < ^nrel k r v := + le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma fstIdx_lt_verum (n : V) : n < ^⊤[n] := le_iff_lt_succ.mp <| le_pair_left _ _ +@[simp] lemma lt_and_left (p q : V) : p < p ^⋏ q := le_iff_lt_succ.mp <| le_trans (le_pair_left _ _) <| le_pair_right _ _ +@[simp] lemma lt_and_right (p q : V) : q < p ^⋏ q := le_iff_lt_succ.mp <| le_trans (le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma fstIdx_lt_falsum (n : V) : n < ^⊥[n] := le_iff_lt_succ.mp <| le_pair_left _ _ +@[simp] lemma lt_or_left (p q : V) : p < p ^⋎ q := le_iff_lt_succ.mp <| le_trans (le_pair_left _ _) <| le_pair_right _ _ +@[simp] lemma lt_or_right (p q : V) : q < p ^⋎ q := le_iff_lt_succ.mp <| le_trans (le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma fstIdx_lt_and (n p q : V) : n < p ^⋏[n] q := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma lt_and_left (n p q : V) : p < p ^⋏[n] q := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma lt_and_right (n p q : V) : q < p ^⋏[n] q := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ +@[simp] lemma lt_forall (p : V) : p < ^∀ p := le_iff_lt_succ.mp <| le_pair_right _ _ -@[simp] lemma fstIdx_lt_or (n p q : V) : n < p ^⋎[n] q := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma lt_or_left (n p q : V) : p < p ^⋎[n] q := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_left _ _) <| le_pair_right _ _) <| le_pair_right _ _ -@[simp] lemma lt_or_right (n p q : V) : q < p ^⋎[n] q := le_iff_lt_succ.mp <| le_trans (le_trans (le_pair_right _ _) <| le_pair_right _ _) <| le_pair_right _ _ - -@[simp] lemma fstIdx_lt_forall (n p : V) : n < ^∀[n] p := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma lt_forall (n p : V) : p < ^∀[n] p := le_iff_lt_succ.mp <| le_trans (le_pair_right _ _) <| le_pair_right _ _ - -@[simp] lemma fstIdx_lt_exists (n p : V) : n < ^∃[n] p := le_iff_lt_succ.mp <| le_pair_left _ _ -@[simp] lemma lt_exists (n p : V) : p < ^∃[n] p := le_iff_lt_succ.mp <| le_trans (le_pair_right _ _) <| le_pair_right _ _ - -@[simp] lemma fstIdx_rel (n k r v : V) : fstIdx (^rel n k r v) = n := by simp [fstIdx, qqRel] -@[simp] lemma fstIdx_nrel (n k r v : V) : fstIdx (^nrel n k r v) = n := by simp [fstIdx, qqNRel] -@[simp] lemma fstIdx_verum (n : V) : fstIdx ^⊤[n] = n := by simp [fstIdx, qqVerum] -@[simp] lemma fstIdx_falsum (n : V) : fstIdx ^⊥[n] = n := by simp [fstIdx, qqFalsum] -@[simp] lemma fstIdx_and (n p q : V) : fstIdx (p ^⋏[n] q) = n := by simp [fstIdx, qqAnd] -@[simp] lemma fstIdx_or (n p q : V) : fstIdx (p ^⋎[n] q) = n := by simp [fstIdx, qqOr] -@[simp] lemma fstIdx_all (n p : V) : fstIdx (^∀[n] p) = n := by simp [fstIdx, qqAll] -@[simp] lemma fstIdx_ex (n p : V) : fstIdx (^∃[n] p) = n := by simp [fstIdx, qqEx] +@[simp] lemma lt_exists (p : V) : p < ^∃ p := le_iff_lt_succ.mp <| le_pair_right _ _ namespace FormalizedFormula variable (L) def Phi (C : Set V) (p : V) : Prop := - (∃ n k r v, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^rel n k r v) ∨ - (∃ n k r v, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^nrel n k r v) ∨ - (∃ n, p = ^⊤[n]) ∨ - (∃ n, p = ^⊥[n]) ∨ - (∃ n q r, (q ∈ C ∧ n = fstIdx q) ∧ (r ∈ C ∧ n = fstIdx r) ∧ p = q ^⋏[n] r) ∨ - (∃ n q r, (q ∈ C ∧ n = fstIdx q) ∧ (r ∈ C ∧ n = fstIdx r) ∧ p = q ^⋎[n] r) ∨ - (∃ n q, (q ∈ C ∧ n + 1 = fstIdx q) ∧ p = ^∀[n] q) ∨ - (∃ n q, (q ∈ C ∧ n + 1 = fstIdx q) ∧ p = ^∃[n] q) + (∃ k R v, L.Rel k R ∧ L.IsUTermVec k v ∧ p = ^rel k R v) ∨ + (∃ k R v, L.Rel k R ∧ L.IsUTermVec k v ∧ p = ^nrel k R v) ∨ + (p = ^⊤) ∨ + (p = ^⊥) ∨ + (∃ p₁ p₂, p₁ ∈ C ∧ p₂ ∈ C ∧ p = p₁ ^⋏ p₂) ∨ + (∃ p₁ p₂, p₁ ∈ C ∧ p₂ ∈ C ∧ p = p₁ ^⋎ p₂) ∨ + (∃ p₁, p₁ ∈ C ∧ p = ^∀ p₁) ∨ + (∃ p₁, p₁ ∈ C ∧ p = ^∃ p₁) private lemma phi_iff (C p : V) : Phi L {x | x ∈ C} p ↔ - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^rel n k r v) ∨ - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^nrel n k r v) ∨ - (∃ n < p, p = ^⊤[n]) ∨ - (∃ n < p, p = ^⊥[n]) ∨ - (∃ n < p, ∃ q < p, ∃ r < p, (q ∈ C ∧ n = fstIdx q) ∧ (r ∈ C ∧ n = fstIdx r) ∧ p = q ^⋏[n] r) ∨ - (∃ n < p, ∃ q < p, ∃ r < p, (q ∈ C ∧ n = fstIdx q) ∧ (r ∈ C ∧ n = fstIdx r) ∧ p = q ^⋎[n] r) ∨ - (∃ n < p, ∃ q < p, (q ∈ C ∧ n + 1 = fstIdx q) ∧ p = ^∀[n] q) ∨ - (∃ n < p, ∃ q < p, (q ∈ C ∧ n + 1 = fstIdx q) ∧ p = ^∃[n] q) where + (∃ k < p, ∃ r < p, ∃ v < p, L.Rel k r ∧ L.IsUTermVec k v ∧ p = ^rel k r v) ∨ + (∃ k < p, ∃ r < p, ∃ v < p, L.Rel k r ∧ L.IsUTermVec k v ∧ p = ^nrel k r v) ∨ + (p = ^⊤) ∨ + (p = ^⊥) ∨ + (∃ p₁ < p, ∃ p₂ < p, p₁ ∈ C ∧ p₂ ∈ C ∧ p = p₁ ^⋏ p₂) ∨ + (∃ p₁ < p, ∃ p₂ < p, p₁ ∈ C ∧ p₂ ∈ C ∧ p = p₁ ^⋎ p₂) ∨ + (∃ p₁ < p, p₁ ∈ C ∧ p = ^∀ p₁) ∨ + (∃ p₁ < p, p₁ ∈ C ∧ p = ^∃ p₁) where mp := by - rintro (⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, k, r, v, hkr, hv, rfl⟩ | H) - · left; refine ⟨n, ?_, k, ?_, r, ?_, v, ?_, hkr, hv, rfl⟩ <;> simp - · right; left; refine ⟨n, ?_, k, ?_, r, ?_, v, ?_, hkr, hv, rfl⟩ <;> simp + rintro (⟨k, r, v, hkr, hv, rfl⟩ | ⟨k, r, v, hkr, hv, rfl⟩ | H) + · left; refine ⟨k, ?_, r, ?_, v, ?_, hkr, hv, rfl⟩ <;> simp + · right; left; refine ⟨k, ?_, r, ?_, v, ?_, hkr, hv, rfl⟩ <;> simp right; right - rcases H with (⟨n, rfl⟩ | ⟨n, rfl⟩ | H) - · left; exact ⟨n, by simp, rfl⟩ - · right; left; exact ⟨n, by simp, rfl⟩ + rcases H with (rfl | rfl | H) + · left; rfl + · right; left; rfl right; right - rcases H with (⟨n, q, r, hp, hq, rfl⟩ | ⟨n, q, r, hp, hq, rfl⟩ | H) - · left; refine ⟨n, ?_, q, ?_, r, ?_, hp, hq, rfl⟩ <;> simp - · right; left; refine ⟨n, ?_, q, ?_, r, ?_, hp, hq, rfl⟩ <;> simp + rcases H with (⟨q, r, hp, hq, rfl⟩ | ⟨q, r, hp, hq, rfl⟩ | H) + · left; refine ⟨q, ?_, r, ?_, hp, hq, rfl⟩ <;> simp + · right; left; refine ⟨q, ?_, r, ?_, hp, hq, rfl⟩ <;> simp right; right - rcases H with (⟨n, q, h, rfl⟩ | ⟨n, q, h, rfl⟩) - · left; refine ⟨n, ?_, q, ?_, h, rfl⟩ <;> simp - · right; refine ⟨n, ?_, q, ?_, h, rfl⟩ <;> simp + rcases H with (⟨q, h, rfl⟩ | ⟨q, h, rfl⟩) + · left; refine ⟨q, ?_, h, rfl⟩; simp + · right; refine ⟨q, ?_, h, rfl⟩; simp mpr := by unfold Phi - rintro (⟨n, _, k, _, r, _, v, _, hkr, hv, rfl⟩ | ⟨n, _, k, _, r, _, v, _, hkr, hv, rfl⟩ | H) - · left; exact ⟨n, k, r, v, hkr, hv, rfl⟩ - · right; left; exact ⟨n, k, r, v, hkr, hv, rfl⟩ + rintro (⟨k, _, r, _, v, _, hkr, hv, rfl⟩ | ⟨k, _, r, _, v, _, hkr, hv, rfl⟩ | H) + · left; exact ⟨k, r, v, hkr, hv, rfl⟩ + · right; left; exact ⟨k, r, v, hkr, hv, rfl⟩ right; right - rcases H with (⟨n, _, rfl⟩ | ⟨n, _, rfl⟩ | H) - · left; exact ⟨n, rfl⟩ - · right; left; exact ⟨n, rfl⟩ + rcases H with (rfl | rfl | H) + · left; rfl + · right; left; rfl right; right - rcases H with (⟨n, _, q, _, r, _, hq, hr, rfl⟩ | ⟨n, _, q, _, r, _, hq, hr, rfl⟩ | H) - · left; exact ⟨n, q, r, hq, hr, rfl⟩ - · right; left; exact ⟨n, q, r, hq, hr, rfl⟩ + rcases H with (⟨q, _, r, _, hq, hr, rfl⟩ | ⟨q, _, r, _, hq, hr, rfl⟩ | H) + · left; exact ⟨q, r, hq, hr, rfl⟩ + · right; left; exact ⟨q, r, hq, hr, rfl⟩ right; right - rcases H with (⟨n, _, q, _, hq, rfl⟩ | ⟨n, _, q, _, hq, rfl⟩) - · left; exact ⟨n, q, hq, rfl⟩ - · right; exact ⟨n, q, hq, rfl⟩ + rcases H with (⟨q, _, hq, rfl⟩ | ⟨q, _, hq, rfl⟩) + · left; exact ⟨q, hq, rfl⟩ + · right; exact ⟨q, hq, rfl⟩ def formulaAux : 𝚺₀.Semisentence 2 := .mkSigma “p C | - (∃ n < p, !qqVerumDef p n) ∨ - (∃ n < p, !qqFalsumDef p n) ∨ - (∃ n < p, ∃ q < p, ∃ r < p, (q ∈ C ∧ !fstIdxDef n q) ∧ (r ∈ C ∧ !fstIdxDef n r) ∧ !qqAndDef p n q r) ∨ - (∃ n < p, ∃ q < p, ∃ r < p, (q ∈ C ∧ !fstIdxDef n q) ∧ (r ∈ C ∧ !fstIdxDef n r) ∧ !qqOrDef p n q r) ∨ - (∃ n < p, ∃ q < p, (q ∈ C ∧ !fstIdxDef (n + 1) q) ∧ !qqAllDef p n q) ∨ - (∃ n < p, ∃ q < p, (q ∈ C ∧ !fstIdxDef (n + 1) q) ∧ !qqExDef p n q)” + !qqVerumDef p ∨ + !qqFalsumDef p ∨ + (∃ p₁ < p, ∃ p₂ < p, p₁ ∈ C ∧ p₂ ∈ C ∧ !qqAndDef p p₁ p₂) ∨ + (∃ p₁ < p, ∃ p₂ < p, p₁ ∈ C ∧ p₂ ∈ C ∧ !qqOrDef p p₁ p₂) ∨ + (∃ p₁ < p, p₁ ∈ C ∧ !qqAllDef p p₁) ∨ + (∃ p₁ < p, p₁ ∈ C ∧ !qqExDef p p₁)” (by simp) def blueprint (pL : LDef) : Fixpoint.Blueprint 0 := ⟨.mkDelta (.mkSigma “p C | - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.semitermVecDef.sigma k n v ∧ !qqRelDef p n k r v) ∨ - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.semitermVecDef.sigma k n v ∧ !qqNRelDef p n k r v) ∨ + (∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.isUTermVecDef.sigma k v ∧ !qqRelDef p k r v) ∨ + (∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.isUTermVecDef.sigma k v ∧ !qqNRelDef p k r v) ∨ !formulaAux p C” (by simp)) (.mkPi “p C | - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.semitermVecDef.pi k n v ∧ !qqRelDef p n k r v) ∨ - (∃ n < p, ∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.semitermVecDef.pi k n v ∧ !qqNRelDef p n k r v) ∨ + (∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.isUTermVecDef.pi k v ∧ !qqRelDef p k r v) ∨ + (∃ k < p, ∃ r < p, ∃ v < p, !pL.rel k r ∧ !pL.isUTermVecDef.pi k v ∧ !qqNRelDef p k r v) ∨ !formulaAux p C” (by simp))⟩ def construction : Fixpoint.Construction V (blueprint pL) where Φ := fun _ ↦ Phi L defined := ⟨ by intro v - -- simp [blueprint, HierarchySymbol.Semiformula.val_sigma, (termSeq_defined L).proper.iff'] - simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, HierarchySymbol.Semiformula.val_sigma, - HierarchySymbol.Semiformula.sigma_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_or, + -- simp [blueprint, HierarchySymbol.Semiformula.val_sigma, L.isUTermVec_defined.proper.iff'] + simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, + HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.sigma_mkDelta, + HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, - Matrix.cons_val_three, Fin.succ_one_eq_two, LogicalConnective.HomClass.map_and, - Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_zero, - Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, Matrix.cons_val_four, - Matrix.cons_val_succ, eval_qqRelDef, LogicalConnective.Prop.and_eq, eval_qqNRelDef, - LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, - (semitermVec_defined L).proper.iff'], + LogicalConnective.HomClass.map_and, Semiformula.eval_substs, Matrix.comp_vecCons', + Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, + Matrix.cons_val_three, Fin.succ_one_eq_two, eval_qqRelDef, LogicalConnective.Prop.and_eq, + eval_qqNRelDef, LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, + HierarchySymbol.Semiformula.val_mkPi, L.isUTermVec_defined.proper.iff'] + , by intro v - -- simpa [blueprint, Language.Defined.eval_rel_iff (L := L), eval_termSeq L, HierarchySymbol.Semiformula.val_sigma, formulaAux] using phi_iff L _ _ + -- simpa [blueprint, Language.Defined.eval_rel_iff (L := L), L.isUTermVec_defined.df.iff, + -- HierarchySymbol.Semiformula.val_sigma, formulaAux] using phi_iff L _ _ simpa only [Fin.isValue, Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, - HierarchySymbol.Semiformula.val_sigma, formulaAux, HierarchySymbol.Semiformula.val_mkSigma, - LogicalConnective.HomClass.map_or, HierarchySymbol.Semiformula.val_mkDelta, Semiformula.eval_bexLT, - Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, Matrix.cons_val_two, - Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, Matrix.cons_val_three, - Fin.succ_one_eq_two, LogicalConnective.HomClass.map_and, Semiformula.eval_substs, - Matrix.comp_vecCons', Matrix.cons_val_zero, Matrix.cons_val_fin_one, - Matrix.constant_eq_singleton, Language.Defined.eval_rel_iff (L := L), (semitermVec_defined L).df.iff, - Matrix.cons_val_four, Matrix.cons_val_succ, eval_qqRelDef, LogicalConnective.Prop.and_eq, + HierarchySymbol.Semiformula.val_sigma, formulaAux, + HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_or, + HierarchySymbol.Semiformula.val_mkDelta, Semiformula.eval_bexLT, Semiterm.val_bvar, + Matrix.cons_val_one, Matrix.vecHead, Matrix.cons_val_two, Matrix.vecTail, + Function.comp_apply, Fin.succ_zero_eq_one, LogicalConnective.HomClass.map_and, + Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_zero, + Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, + Language.Defined.eval_rel_iff (L := L), L.isUTermVec_defined.df.iff, + Matrix.cons_val_three, Fin.succ_one_eq_two, eval_qqRelDef, LogicalConnective.Prop.and_eq, eval_qqNRelDef, eval_qqVerumDef, eval_qqFalsumDef, Semiformula.eval_operator₂, - Structure.Mem.mem, eval_fstIdxDef, eval_qqAndDef, eval_qqOrDef, Semiterm.val_operator₂, - Semiterm.val_operator₀, Structure.numeral_eq_numeral, ORingSymbol.one_eq_one, - Structure.Add.add, eval_qqAllDef, eval_qqExDef, LogicalConnective.Prop.or_eq] using - phi_iff L _ _⟩ + Structure.Mem.mem, eval_qqAndDef, eval_qqOrDef, eval_qqAllDef, eval_qqExDef, + LogicalConnective.Prop.or_eq] using phi_iff L _ _⟩ + monotone := by unfold Phi rintro C C' hC _ x (h | h | h | h | H) @@ -345,13 +311,13 @@ def construction : Fixpoint.Construction V (blueprint pL) where · right; right; left; exact h · right; right; right; left; exact h right; right; right; right - rcases H with (⟨n, q, r, ⟨hqC, hq⟩, ⟨hrC, hr⟩, rfl⟩ | ⟨n, q, r, ⟨hqC, hq⟩, ⟨hrC, hr⟩, rfl⟩ | H) - · left; exact ⟨n, q, r, ⟨hC hqC, hq⟩, ⟨hC hrC, hr⟩, rfl⟩ - · right; left; exact ⟨n, q, r, ⟨hC hqC, hq⟩, ⟨hC hrC, hr⟩, rfl⟩ + rcases H with (⟨q, r, hqC, hrC, rfl⟩ | ⟨q, r, hqC, hrC, rfl⟩ | H) + · left; exact ⟨q, r, hC hqC, hC hrC, rfl⟩ + · right; left; exact ⟨q, r, hC hqC, hC hrC, rfl⟩ right; right - rcases H with (⟨n, q, ⟨hqC, hq⟩, rfl⟩ | ⟨n, q, ⟨hqC, hq⟩, rfl⟩) - · left; exact ⟨n, q, ⟨hC hqC, hq⟩, rfl⟩ - · right; exact ⟨n, q, ⟨hC hqC, hq⟩, rfl⟩ + rcases H with (⟨q, hqC, rfl⟩ | ⟨q, hqC, rfl⟩) + · left; exact ⟨q, hC hqC, rfl⟩ + · right; exact ⟨q, hC hqC, rfl⟩ instance : (construction L).StrongFinite V where strong_finite := by @@ -362,13 +328,13 @@ instance : (construction L).StrongFinite V where · right; right; left; exact h · right; right; right; left; exact h right; right; right; right - rcases H with (⟨n, q, r, ⟨hqC, hq⟩, ⟨hrC, hr⟩, rfl⟩ | ⟨n, q, r, ⟨hqC, hq⟩, ⟨hrC, hr⟩, rfl⟩ | H) - · left; exact ⟨n, q, r, ⟨by simp [hqC], hq⟩, ⟨by simp [hrC], hr⟩, rfl⟩ - · right; left; exact ⟨n, q, r, ⟨by simp [hqC], hq⟩, ⟨by simp [hrC], hr⟩, rfl⟩ + rcases H with (⟨q, r, hqC, hrC, rfl⟩ | ⟨q, r, hqC, hrC, rfl⟩ | H) + · left; exact ⟨q, r, by simp [hqC], by simp [hrC], rfl⟩ + · right; left; exact ⟨q, r, by simp [hqC], by simp [hrC], rfl⟩ right; right - rcases H with (⟨n, q, ⟨hqC, hq⟩, rfl⟩ | ⟨n, q, ⟨hqC, hq⟩, rfl⟩) - · left; exact ⟨n, q, ⟨by simp [hqC], hq⟩, rfl⟩ - · right; exact ⟨n, q, ⟨by simp [hqC], hq⟩, rfl⟩ + rcases H with (⟨q, hqC, rfl⟩ | ⟨q, hqC, rfl⟩) + · left; exact ⟨q, by simp [hqC], rfl⟩ + · right; exact ⟨q, by simp [hqC], rfl⟩ end FormalizedFormula @@ -378,187 +344,189 @@ open FormalizedFormula variable (L) -def Language.UFormula : V → Prop := (construction L).Fixpoint ![] +def Language.IsUFormula : V → Prop := (construction L).Fixpoint ![] -def _root_.LO.FirstOrder.Arith.LDef.uformulaDef (pL : LDef) : 𝚫₁.Semisentence 1 := - (blueprint pL).fixpointDefΔ₁ - -lemma uformula_defined : 𝚫₁-Predicate L.UFormula via pL.uformulaDef := - (construction L).fixpoint_definedΔ₁ - -@[simp] lemma eval_uformulaDef (v) : - Semiformula.Evalbm V v pL.uformulaDef.val ↔ L.UFormula (v 0) := (uformula_defined L).df.iff v - -instance uformulaDef_definable : 𝚫₁-Predicate L.UFormula := (uformula_defined L).to_definable - -@[simp, definability] instance uformulaDef_definable' (Γ) : Γ-[m + 1]-Predicate L.UFormula := - .of_deltaOne (uformulaDef_definable L) _ _ - -def Language.Semiformula (n p : V) : Prop := L.UFormula p ∧ n = fstIdx p - -abbrev Language.Formula (p : V) : Prop := L.Semiformula 0 p +section -lemma Language.UFormula.toSemiformula {p} (h : L.UFormula p) : L.Semiformula (fstIdx p) p := - ⟨h, by rfl⟩ +def _root_.LO.FirstOrder.Arith.LDef.isUFormulaDef (pL : LDef) : 𝚫₁.Semisentence 1 := + (blueprint pL).fixpointDefΔ₁ -def _root_.LO.FirstOrder.Arith.LDef.isSemiformulaDef (pL : LDef) : 𝚫₁.Semisentence 2 := .mkDelta - (.mkSigma “n p | !pL.uformulaDef.sigma p ∧ !fstIdxDef n p” (by simp)) - (.mkPi “n p | !pL.uformulaDef.pi p ∧ !fstIdxDef n p” (by simp)) +lemma Language.isUFormula_defined : 𝚫₁-Predicate L.IsUFormula via pL.isUFormulaDef := (construction L).fixpoint_definedΔ₁ -lemma semiformula_defined : 𝚫₁-Relation L.Semiformula via pL.isSemiformulaDef where - left := by intro v; simp [LDef.isSemiformulaDef, HierarchySymbol.Semiformula.val_sigma, (uformula_defined L).proper.iff'] - right := by intro v; simp [LDef.isSemiformulaDef, HierarchySymbol.Semiformula.val_sigma, eval_uformulaDef L, Language.Semiformula, eq_comm] +instance Language.isUFormulaDef_definable : 𝚫₁-Predicate L.IsUFormula := L.isUFormula_defined.to_definable -instance semiformula_definable : 𝚫₁-Relation L.Semiformula := (semiformula_defined L).to_definable +instance Language.isUFormulaDef_definable' : Γ-[m + 1]-Predicate L.IsUFormula := L.isUFormulaDef_definable.of_deltaOne -@[simp, definability] instance semiformula_defined' (Γ) : Γ-[m + 1]-Relation L.Semiformula := - .of_deltaOne (semiformula_definable L) _ _ +end variable {L} -local prefix:80 "𝐔 " => L.UFormula - -lemma Language.UFormula.case_iff {p : V} : - 𝐔 p ↔ - (∃ n k r v, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^rel n k r v) ∨ - (∃ n k r v, L.Rel k r ∧ L.SemitermVec k n v ∧ p = ^nrel n k r v) ∨ - (∃ n, p = ^⊤[n]) ∨ - (∃ n, p = ^⊥[n]) ∨ - (∃ n q r, L.Semiformula n q ∧ L.Semiformula n r ∧ p = q ^⋏[n] r) ∨ - (∃ n q r, L.Semiformula n q ∧ L.Semiformula n r ∧ p = q ^⋎[n] r) ∨ - (∃ n q, L.Semiformula (n + 1) q ∧ p = ^∀[n] q) ∨ - (∃ n q, L.Semiformula (n + 1) q ∧ p = ^∃[n] q) := +lemma Language.IsUFormula.case_iff {p : V} : + L.IsUFormula p ↔ + (∃ k R v, L.Rel k R ∧ L.IsUTermVec k v ∧ p = ^rel k R v) ∨ + (∃ k R v, L.Rel k R ∧ L.IsUTermVec k v ∧ p = ^nrel k R v) ∨ + (p = ^⊤) ∨ + (p = ^⊥) ∨ + (∃ p₁ p₂, L.IsUFormula p₁ ∧ L.IsUFormula p₂ ∧ p = p₁ ^⋏ p₂) ∨ + (∃ p₁ p₂, L.IsUFormula p₁ ∧ L.IsUFormula p₂ ∧ p = p₁ ^⋎ p₂) ∨ + (∃ p₁, L.IsUFormula p₁ ∧ p = ^∀ p₁) ∨ + (∃ p₁, L.IsUFormula p₁ ∧ p = ^∃ p₁) := (construction L).case -alias ⟨Language.UFormula.case, Language.UFormula.mk⟩ := Language.UFormula.case_iff +alias ⟨Language.IsUFormula.case, Language.IsUFormula.mk⟩ := Language.IsUFormula.case_iff -@[simp] lemma Language.UFormula.rel {n k r v : V} : - 𝐔 (^rel n k r v) ↔ L.Rel k r ∧ L.SemitermVec k n v := +@[simp] lemma Language.IsUFormula.rel {k r v : V} : + L.IsUFormula (^rel k r v) ↔ L.Rel k r ∧ L.IsUTermVec k v := ⟨by intro h - rcases h.case with (⟨n, k, r, v, hkr, hv, h⟩ | ⟨_, _, _, _, _, _, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, h⟩ | ⟨_, _, _, h⟩) <;> + rcases h.case with (⟨k, r, v, hkr, hv, h⟩ | ⟨_, _, _, _, _, h⟩ | h | h | + ⟨_, _, _, _, h⟩ | ⟨_, _, _, _, h⟩ | ⟨_, _, h⟩ | ⟨_, _, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact ⟨hkr, hv⟩, by rintro ⟨hkr, hv⟩ - exact Language.UFormula.mk (Or.inl ⟨n, k, r, v, hkr, hv, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inl ⟨k, r, v, hkr, hv, rfl⟩)⟩ -@[simp] lemma Language.UFormula.nrel {n k r v : V} : - 𝐔 (^nrel n k r v) ↔ L.Rel k r ∧ L.SemitermVec k n v := +@[simp] lemma Language.IsUFormula.nrel {k r v : V} : + L.IsUFormula (^nrel k r v) ↔ L.Rel k r ∧ L.IsUTermVec k v := ⟨by intro h - rcases h.case with (⟨_, _, _, _, _, _, h⟩ | ⟨n, k, r, v, hkr, hv, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, h⟩ | ⟨_, _, _, h⟩) <;> + rcases h.case with (⟨_, _, _, _, _, h⟩ | ⟨k, r, v, hkr, hv, h⟩ | h | h | + ⟨_, _, _, _, h⟩ | ⟨_, _, _, _, h⟩ | ⟨_, _, h⟩ | ⟨_, _, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact ⟨hkr, hv⟩, by rintro ⟨hkr, hv⟩ - exact Language.UFormula.mk (Or.inr <| Or.inl ⟨n, k, r, v, hkr, hv, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inr <| Or.inl ⟨k, r, v, hkr, hv, rfl⟩)⟩ -@[simp] lemma Language.UFormula.verum (n : V) : 𝐔 ^⊤[n] := - Language.UFormula.mk (Or.inr <| Or.inr <| Or.inl ⟨n, rfl⟩) +@[simp] lemma Language.IsUFormula.verum : L.IsUFormula ^⊤ := + Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inl rfl) -@[simp] lemma Language.UFormula.falsum (n : V) : 𝐔 ^⊥[n] := - Language.UFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, rfl⟩) +@[simp] lemma Language.IsUFormula.falsum : L.IsUFormula ^⊥ := + Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inl rfl) -@[simp] lemma Language.UFormula.and {n p q : V} : - 𝐔 (p ^⋏[n] q) ↔ L.Semiformula n p ∧ L.Semiformula n q := +@[simp] lemma Language.IsUFormula.and {p q : V} : + L.IsUFormula (p ^⋏ q) ↔ L.IsUFormula p ∧ L.IsUFormula q := ⟨by intro h - rcases h.case with (⟨_, _, _, _, _, _, h⟩ | ⟨_, _, _, _, _, _, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, hp, hq, h⟩ | ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, h⟩ | ⟨_, _, _, h⟩) <;> + rcases h.case with (⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | h | h | + ⟨_, _, hp, hq, h⟩ | ⟨_, _, _, _, h⟩ | ⟨_, _, h⟩ | ⟨_, _, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact ⟨hp, hq⟩, by rintro ⟨hp, hq⟩ - exact Language.UFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl - ⟨n, p, q, hp, hq, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p, q, hp, hq, rfl⟩)⟩ -@[simp] lemma Language.UFormula.or {n p q : V} : - 𝐔 (p ^⋎[n] q) ↔ L.Semiformula n p ∧ L.Semiformula n q := +@[simp] lemma Language.IsUFormula.or {p q : V} : + L.IsUFormula (p ^⋎ q) ↔ L.IsUFormula p ∧ L.IsUFormula q := ⟨by intro h - rcases h.case with (⟨_, _, _, _, _, _, h⟩ | ⟨_, _, _, _, _, _, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, hp, hq, h⟩ | ⟨_, _, _, h⟩ | ⟨_, _, _, h⟩) <;> + rcases h.case with (⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | h | h | + ⟨_, _, _, _, h⟩ | ⟨_, _, hp, hq, h⟩ | ⟨_, _, h⟩ | ⟨_, _, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact ⟨hp, hq⟩, by rintro ⟨hp, hq⟩ - exact Language.UFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl - ⟨n, p, q, hp, hq, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p, q, hp, hq, rfl⟩)⟩ -@[simp] lemma Language.UFormula.all {n p : V} : - 𝐔 (^∀[n] p) ↔ L.Semiformula (n + 1) p := +@[simp] lemma Language.IsUFormula.all {p : V} : + L.IsUFormula (^∀ p) ↔ L.IsUFormula p := ⟨by intro h - rcases h.case with (⟨_, _, _, _, _, _, h⟩ | ⟨_, _, _, _, _, _, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | ⟨_, _, hp, h⟩ | ⟨_, _, _, h⟩) <;> + rcases h.case with (⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | h | h | + ⟨_, _, _, _, h⟩ | ⟨_, _, _, _, h⟩ | ⟨_, hp, h⟩ | ⟨_, _, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact hp, by rintro hp - exact Language.UFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, p, hp, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p, hp, rfl⟩)⟩ -@[simp] lemma Language.UFormula.ex {n p : V} : - 𝐔 (^∃[n] p) ↔ L.Semiformula (n + 1) p := +@[simp] lemma Language.IsUFormula.ex {p : V} : + L.IsUFormula (^∃ p) ↔ L.IsUFormula p := ⟨by intro h - rcases h.case with (⟨_, _, _, _, _, _, h⟩ | ⟨_, _, _, _, _, _, h⟩ | ⟨_, h⟩ | ⟨_, h⟩ | - ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | ⟨_, _, _, h⟩ | ⟨_, _, hp, h⟩) <;> + rcases h.case with (⟨_, _, _, _, _, h⟩ | ⟨_, _, _, _, _, h⟩ | h | h | + ⟨_, _, _, _, h⟩ | ⟨_, _, _, _, h⟩ | ⟨_, _, h⟩ | ⟨_, hp, h⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] at h · rcases h with ⟨rfl, rfl, rfl, rfl⟩; exact hp, by rintro hp - exact Language.UFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr ⟨n, p, hp, rfl⟩)⟩ + exact Language.IsUFormula.mk (Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr ⟨p, hp, rfl⟩)⟩ -lemma Language.UFormula.pos {p : V} (h : L.UFormula p) : 0 < p := by +lemma Language.IsUFormula.pos {p : V} (h : L.IsUFormula p) : 0 < p := by rcases h.case with (⟨_, _, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, _, rfl⟩ | ⟨_, rfl⟩ | ⟨_, rfl⟩ | ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, rfl⟩ | ⟨_, _, _, rfl⟩) <;> simp [qqRel, qqNRel, qqVerum, qqFalsum, qqAnd, qqOr, qqAll, qqEx] -lemma Language.Semiformula.pos {n p : V} (h : L.Semiformula n p) : 0 < p := h.1.pos - -@[simp] lemma Language.UFormula.not_zero : ¬L.UFormula (0 : V) := by intro h; simpa using h.pos - -@[simp] lemma Language.Semiformula.not_zero (m : V) : ¬L.Semiformula m (0 : V) := by intro h; simpa using h.pos - -@[simp] lemma Language.Semiformula.rel {n k r v : V} : - L.Semiformula n (^rel n k r v) ↔ L.Rel k r ∧ L.SemitermVec k n v := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.nrel {n k r v : V} : - L.Semiformula n (^nrel n k r v) ↔ L.Rel k r ∧ L.SemitermVec k n v := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.verum (n : V) : L.Semiformula n ^⊤[n] := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.falsum (n : V) : L.Semiformula n ^⊥[n] := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.and {n p q : V} : - L.Semiformula n (p ^⋏[n] q) ↔ L.Semiformula n p ∧ L.Semiformula n q := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.or {n p q : V} : - L.Semiformula n (p ^⋎[n] q) ↔ L.Semiformula n p ∧ L.Semiformula n q := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.all {n p : V} : L.Semiformula n (^∀[n] p) ↔ L.Semiformula (n + 1) p := by simp [Language.Semiformula] -@[simp] lemma Language.Semiformula.ex {n p : V} : L.Semiformula n (^∃[n] p) ↔ L.Semiformula (n + 1) p := by simp [Language.Semiformula] - -lemma Language.UFormula.induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) - (hrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P (^rel n k r v)) - (hnrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P (^nrel n k r v)) - (hverum : ∀ n, P ^⊤[n]) - (hfalsum : ∀ n, P ^⊥[n]) - (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P p → P q → P (p ^⋏[n] q)) - (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P p → P q → P (p ^⋎[n] q)) - (hall : ∀ n p, L.Semiformula (n + 1) p → P p → P (^∀[n] p)) - (hex : ∀ n p, L.Semiformula (n + 1) p → P p → P (^∃[n] p)) : - ∀ p, 𝐔 p → P p := +--lemma Language.IsSemiformula.pos {n p : V} (h : L.Semiformula n p) : 0 < p := h.1.pos + +@[simp] lemma Language.IsUFormula.not_zero : ¬L.IsUFormula (0 : V) := by intro h; simpa using h.pos + +-- @[simp] lemma Language.IsSemiformula.not_zero (m : V) : ¬L.Semiformula m (0 : V) := by intro h; simpa using h.pos + +/- +@[simp] lemma Language.IsSemiformula.rel {k r v : V} : + L.IsUFormula (^rel k r v) ↔ L.Rel k r ∧ L.IsUTermVec k v := by simp +@[simp] lemma Language.IsSemiformula.nrel {n k r v : V} : + L.Semiformula n (^nrel n k r v) ↔ L.Rel k r ∧ L.SemitermVec k n v := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.verum (n : V) : L.Semiformula n ^⊤[n] := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.falsum (n : V) : L.Semiformula n ^⊥[n] := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.and {n p q : V} : + L.Semiformula n (p ^⋏ q) ↔ L.Semiformula n p ∧ L.Semiformula n q := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.or {n p q : V} : + L.Semiformula n (p ^⋎ q) ↔ L.Semiformula n p ∧ L.Semiformula n q := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.all {n p : V} : L.Semiformula n (^∀ p) ↔ L.Semiformula (n + 1) p := by simp [Language.IsSemiformula] +@[simp] lemma Language.IsSemiformula.ex {n p : V} : L.Semiformula n (^∃ p) ↔ L.Semiformula (n + 1) p := by simp [Language.IsSemiformula] +-/ + +lemma Language.IsUFormula.induction1 (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) + (hrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^rel k r v)) + (hnrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^nrel k r v)) + (hverum : P ^⊤) + (hfalsum : P ^⊥) + (hand : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋏ q)) + (hor : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋎ q)) + (hall : ∀ p, L.IsUFormula p → P p → P (^∀ p)) + (hex : ∀ p, L.IsUFormula p → P p → P (^∃ p)) : + ∀ p, L.IsUFormula p → P p := (construction L).induction (v := ![]) hP (by - rintro C hC x (⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, rfl⟩ | ⟨n, rfl⟩ | - ⟨n, p, q, ⟨hp, hnp⟩, ⟨hq, hnq⟩, rfl⟩ | ⟨n, p, q, ⟨hp, hnp⟩, ⟨hq, hnq⟩, rfl⟩ | ⟨n, p, ⟨hp, hnp⟩, rfl⟩ | ⟨n, p, ⟨hp, hnp⟩, rfl⟩) - · exact hrel n k r v hkr hv - · exact hnrel n k r v hkr hv - · exact hverum n - · exact hfalsum n - · exact hand n p q ⟨(hC p hp).1, hnp⟩ ⟨(hC q hq).1, hnq⟩ (hC p hp).2 (hC q hq).2 - · exact hor n p q ⟨(hC p hp).1, hnp⟩ ⟨(hC q hq).1, hnq⟩ (hC p hp).2 (hC q hq).2 - · exact hall n p ⟨(hC p hp).1, hnp⟩ (hC p hp).2 - · exact hex n p ⟨(hC p hp).1, hnp⟩ (hC p hp).2) - -lemma Language.Semiformula.induction (Γ) {P : V → V → Prop} (hP : Γ-[1]-Relation P) + rintro C hC x (⟨k, r, v, hkr, hv, rfl⟩ | ⟨k, r, v, hkr, hv, rfl⟩ | ⟨n, rfl⟩ | ⟨n, rfl⟩ | + ⟨p, q, hp, hq, rfl⟩ | ⟨p, q, hp, hq, rfl⟩ | ⟨p, hp, rfl⟩ | ⟨p, hp, rfl⟩) + · exact hrel k r v hkr hv + · exact hnrel k r v hkr hv + · exact hverum + · exact hfalsum + · exact hand p q (hC p hp).1 (hC q hq).1 (hC p hp).2 (hC q hq).2 + · exact hor p q (hC p hp).1 (hC q hq).1 (hC p hp).2 (hC q hq).2 + · exact hall p (hC p hp).1 (hC p hp).2 + · exact hex p (hC p hp).1 (hC p hp).2) + +lemma Language.IsUFormula.induction_sigma1 {P : V → Prop} (hP : 𝚺₁-Predicate P) + (hrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^rel k r v)) + (hnrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^nrel k r v)) + (hverum : P ^⊤) + (hfalsum : P ^⊥) + (hand : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋏ q)) + (hor : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋎ q)) + (hall : ∀ p, L.IsUFormula p → P p → P (^∀ p)) + (hex : ∀ p, L.IsUFormula p → P p → P (^∃ p)) : + ∀ p, L.IsUFormula p → P p := + Language.IsUFormula.induction1 𝚺 hP hrel hnrel hverum hfalsum hand hor hall hex + +lemma Language.IsUFormula.induction_pi1 {P : V → Prop} (hP : 𝚷₁-Predicate P) + (hrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^rel k r v)) + (hnrel : ∀ k r v, L.Rel k r → L.IsUTermVec k v → P (^nrel k r v)) + (hverum : P ^⊤) + (hfalsum : P ^⊥) + (hand : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋏ q)) + (hor : ∀ p q, L.IsUFormula p → L.IsUFormula q → P p → P q → P (p ^⋎ q)) + (hall : ∀ p, L.IsUFormula p → P p → P (^∀ p)) + (hex : ∀ p, L.IsUFormula p → P p → P (^∃ p)) : + ∀ p, L.IsUFormula p → P p := + Language.IsUFormula.induction1 𝚷 hP hrel hnrel hverum hfalsum hand hor hall hex + +/- +lemma Language.IsSemiformula.induction (Γ) {P : V → V → Prop} (hP : Γ-[1]-Relation P) (hrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^rel n k r v)) (hnrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^nrel n k r v)) (hverum : ∀ n, P n ^⊤[n]) (hfalsum : ∀ n, P n ^⊥[n]) - (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏[n] q)) - (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎[n] q)) - (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀[n] p)) - (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃[n] p)) : + (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) : ∀ n p, L.Semiformula n p → P n p := by - suffices ∀ p, 𝐔 p → ∀ n ≤ p, fstIdx p = n → P n p + suffices ∀ p, L.IsUFormula p → ∀ n ≤ p, fstIdx p = n → P n p by rintro n p ⟨h, rfl⟩; exact this p h (fstIdx p) (by simp) rfl - apply Language.UFormula.induction (P := fun p ↦ ∀ n ≤ p, fstIdx p = n → P n p) Γ + apply Language.IsUFormula.induction (P := fun p ↦ ∀ n ≤ p, fstIdx p = n → P n p) Γ · apply HierarchySymbol.Boldface.ball_le (by definability) apply HierarchySymbol.Boldface.imp (by definability) simp; exact hP @@ -577,45 +545,46 @@ lemma Language.Semiformula.induction (Γ) {P : V → V → Prop} (hP : Γ-[1]-Re · rintro n p hp ih _ _ rfl simpa using hex n p hp (by simpa [hp.2] using ih (fstIdx p) (by simp) rfl) -lemma Language.Semiformula.induction_sigma₁ {P : V → V → Prop} (hP : 𝚺₁-Relation P) +lemma Language.IsSemiformula.induction_sigma₁ {P : V → V → Prop} (hP : 𝚺₁-Relation P) (hrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^rel n k r v)) (hnrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^nrel n k r v)) (hverum : ∀ n, P n ^⊤[n]) (hfalsum : ∀ n, P n ^⊥[n]) - (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏[n] q)) - (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎[n] q)) - (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀[n] p)) - (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃[n] p)) : + (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) : ∀ n p, L.Semiformula n p → P n p := - Language.Semiformula.induction 𝚺 hP hrel hnrel hverum hfalsum hand hor hall hex + Language.IsSemiformula.induction 𝚺 hP hrel hnrel hverum hfalsum hand hor hall hex -lemma Language.Semiformula.induction_pi1 {P : V → V → Prop} (hP : 𝚷₁-Relation P) +lemma Language.IsSemiformula.induction_pi1 {P : V → V → Prop} (hP : 𝚷₁-Relation P) (hrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^rel n k r v)) (hnrel : ∀ n k r v, L.Rel k r → L.SemitermVec k n v → P n (^nrel n k r v)) (hverum : ∀ n, P n ^⊤[n]) (hfalsum : ∀ n, P n ^⊥[n]) - (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏[n] q)) - (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎[n] q)) - (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀[n] p)) - (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃[n] p)) : + (hand : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.Semiformula n p → L.Semiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.Semiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) : ∀ n p, L.Semiformula n p → P n p := - Language.Semiformula.induction 𝚷 hP hrel hnrel hverum hfalsum hand hor hall hex + Language.IsSemiformula.induction 𝚷 hP hrel hnrel hverum hfalsum hand hor hall hex +-/ end formula namespace Language.UformulaRec1 structure Blueprint (pL : LDef) where - rel : 𝚺₁.Semisentence 6 - nrel : 𝚺₁.Semisentence 6 - verum : 𝚺₁.Semisentence 3 - falsum : 𝚺₁.Semisentence 3 - and : 𝚺₁.Semisentence 7 - or : 𝚺₁.Semisentence 7 - all : 𝚺₁.Semisentence 5 - ex : 𝚺₁.Semisentence 5 - allChanges : 𝚺₁.Semisentence 3 - exChanges : 𝚺₁.Semisentence 3 + rel : 𝚺₁.Semisentence 5 + nrel : 𝚺₁.Semisentence 5 + verum : 𝚺₁.Semisentence 2 + falsum : 𝚺₁.Semisentence 2 + and : 𝚺₁.Semisentence 6 + or : 𝚺₁.Semisentence 6 + all : 𝚺₁.Semisentence 4 + ex : 𝚺₁.Semisentence 4 + allChanges : 𝚺₁.Semisentence 2 + exChanges : 𝚺₁.Semisentence 2 namespace Blueprint @@ -623,67 +592,67 @@ variable {pL : LDef} (β : Blueprint pL) def blueprint (β : Blueprint pL) : Fixpoint.Blueprint 0 := ⟨.mkDelta (.mkSigma “pr C | - ∃ param <⁺ pr, ∃ p <⁺ pr, ∃ y <⁺ pr, !pair₃Def pr param p y ∧ !pL.uformulaDef.sigma p ∧ - ((∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, !qqRelDef p n k R v ∧ !β.rel y param n k R v) ∨ - (∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, !qqNRelDef p n k R v ∧ !β.nrel y param n k R v) ∨ - (∃ n < p, !qqVerumDef p n ∧ !β.verum y param n) ∨ - (∃ n < p, !qqFalsumDef p n ∧ !β.falsum y param n) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqAndDef p n p₁ p₂ ∧ !β.and y param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqOrDef p n p₁ p₂ ∧ !β.or y param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - (∃ param', !β.allChanges param' param n ∧ :⟪param', p₁, y₁⟫:∈ C) ∧ !qqAllDef p n p₁ ∧ !β.all y param n p₁ y₁) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - (∃ param', !β.exChanges param' param n ∧ :⟪param', p₁, y₁⟫:∈ C) ∧ !qqExDef p n p₁ ∧ !β.ex y param n p₁ y₁)) + ∃ param <⁺ pr, ∃ p <⁺ pr, ∃ y <⁺ pr, !pair₃Def pr param p y ∧ !pL.isUFormulaDef.sigma p ∧ + ((∃ k < p, ∃ R < p, ∃ v < p, !qqRelDef p k R v ∧ !β.rel y param k R v) ∨ + (∃ k < p, ∃ R < p, ∃ v < p, !qqNRelDef p k R v ∧ !β.nrel y param k R v) ∨ + (!qqVerumDef p ∧ !β.verum y param) ∨ + (!qqFalsumDef p ∧ !β.falsum y param) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqAndDef p p₁ p₂ ∧ !β.and y param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqOrDef p p₁ p₂ ∧ !β.or y param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ y₁ < C, + (∃ param', !β.allChanges param' param ∧ :⟪param', p₁, y₁⟫:∈ C) ∧ !qqAllDef p p₁ ∧ !β.all y param p₁ y₁) ∨ + (∃ p₁ < p, ∃ y₁ < C, + (∃ param', !β.exChanges param' param ∧ :⟪param', p₁, y₁⟫:∈ C) ∧ !qqExDef p p₁ ∧ !β.ex y param p₁ y₁)) ” (by simp)) (.mkPi “pr C | - ∃ param <⁺ pr, ∃ p <⁺ pr, ∃ y <⁺ pr, !pair₃Def pr param p y ∧ !pL.uformulaDef.pi p ∧ - ((∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, !qqRelDef p n k R v ∧ !β.rel.graphDelta.pi.val y param n k R v) ∨ - (∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, !qqNRelDef p n k R v ∧ !β.nrel.graphDelta.pi.val y param n k R v) ∨ - (∃ n < p, !qqVerumDef p n ∧ !β.verum.graphDelta.pi.val y param n) ∨ - (∃ n < p, !qqFalsumDef p n ∧ !β.falsum.graphDelta.pi.val y param n) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqAndDef p n p₁ p₂ ∧ !β.and.graphDelta.pi.val y param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqOrDef p n p₁ p₂ ∧ !β.or.graphDelta.pi.val y param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - (∀ param', !β.allChanges param' param n → :⟪param', p₁, y₁⟫:∈ C) ∧ !qqAllDef p n p₁ ∧ !β.all.graphDelta.pi.val y param n p₁ y₁) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - (∀ param', !β.exChanges param' param n → :⟪param', p₁, y₁⟫:∈ C) ∧ !qqExDef p n p₁ ∧ !β.ex.graphDelta.pi.val y param n p₁ y₁)) + ∃ param <⁺ pr, ∃ p <⁺ pr, ∃ y <⁺ pr, !pair₃Def pr param p y ∧ !pL.isUFormulaDef.pi p ∧ + ((∃ k < p, ∃ R < p, ∃ v < p, !qqRelDef p k R v ∧ !β.rel.graphDelta.pi.val y param k R v) ∨ + (∃ k < p, ∃ R < p, ∃ v < p, !qqNRelDef p k R v ∧ !β.nrel.graphDelta.pi.val y param k R v) ∨ + (!qqVerumDef p ∧ !β.verum.graphDelta.pi.val y param) ∨ + (!qqFalsumDef p ∧ !β.falsum.graphDelta.pi.val y param) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqAndDef p p₁ p₂ ∧ !β.and.graphDelta.pi.val y param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + :⟪param, p₁, y₁⟫:∈ C ∧ :⟪param, p₂, y₂⟫:∈ C ∧ !qqOrDef p p₁ p₂ ∧ !β.or.graphDelta.pi.val y param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ y₁ < C, + (∀ param', !β.allChanges param' param → :⟪param', p₁, y₁⟫:∈ C) ∧ !qqAllDef p p₁ ∧ !β.all.graphDelta.pi.val y param p₁ y₁) ∨ + (∃ p₁ < p, ∃ y₁ < C, + (∀ param', !β.exChanges param' param → :⟪param', p₁, y₁⟫:∈ C) ∧ !qqExDef p p₁ ∧ !β.ex.graphDelta.pi.val y param p₁ y₁)) ” (by simp))⟩ def graph : 𝚺₁.Semisentence 3 := .mkSigma “param p y | ∃ pr, !pair₃Def pr param p y ∧ !β.blueprint.fixpointDef pr” (by simp) def result : 𝚺₁.Semisentence 3 := .mkSigma - “y param p | (!pL.uformulaDef.pi p → !β.graph param p y) ∧ (¬!pL.uformulaDef.sigma p → y = 0)” (by simp) + “y param p | (!pL.isUFormulaDef.pi p → !β.graph param p y) ∧ (¬!pL.isUFormulaDef.sigma p → y = 0)” (by simp) end Blueprint variable (V) structure Construction (L : Arith.Language V) (φ : Blueprint pL) where - rel (param n k R v : V) : V - nrel (param n k R v : V) : V - verum (param n : V) : V - falsum (param n : V) : V - and (param n p₁ p₂ y₁ y₂ : V) : V - or (param n p₁ p₂ y₁ y₂ : V) : V - all (param n p₁ y₁ : V) : V - ex (param n p₁ y₁ : V) : V - allChanges (param n : V) : V - exChanges (param n : V) : V - rel_defined : 𝚺₁.DefinedFunction (fun v ↦ rel (v 0) (v 1) (v 2) (v 3) (v 4)) φ.rel - nrel_defined : 𝚺₁.DefinedFunction (fun v ↦ nrel (v 0) (v 1) (v 2) (v 3) (v 4)) φ.nrel - verum_defined : 𝚺₁.DefinedFunction (fun v ↦ verum (v 0) (v 1)) φ.verum - falsum_defined : 𝚺₁.DefinedFunction (fun v ↦ falsum (v 0) (v 1)) φ.falsum - and_defined : 𝚺₁.DefinedFunction (fun v ↦ and (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) φ.and - or_defined : 𝚺₁.DefinedFunction (fun v ↦ or (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) φ.or - all_defined : 𝚺₁.DefinedFunction (fun v ↦ all (v 0) (v 1) (v 2) (v 3)) φ.all - ex_defined : 𝚺₁.DefinedFunction (fun v ↦ ex (v 0) (v 1) (v 2) (v 3)) φ.ex - allChanges_defined : 𝚺₁-Function₂ allChanges via φ.allChanges - exChanges_defined : 𝚺₁-Function₂ exChanges via φ.exChanges + rel (param k R v : V) : V + nrel (param k R v : V) : V + verum (param : V) : V + falsum (param : V) : V + and (param p₁ p₂ y₁ y₂ : V) : V + or (param p₁ p₂ y₁ y₂ : V) : V + all (param p₁ y₁ : V) : V + ex (param p₁ y₁ : V) : V + allChanges (param : V) : V + exChanges (param : V) : V + rel_defined : 𝚺₁-Function₄ rel via φ.rel + nrel_defined : 𝚺₁-Function₄ nrel via φ.nrel + verum_defined : 𝚺₁-Function₁ verum via φ.verum + falsum_defined : 𝚺₁-Function₁ falsum via φ.falsum + and_defined : 𝚺₁-Function₅ and via φ.and + or_defined : 𝚺₁-Function₅ or via φ.or + all_defined : 𝚺₁-Function₃ all via φ.all + ex_defined : 𝚺₁-Function₃ ex via φ.ex + allChanges_defined : 𝚺₁-Function₁ allChanges via φ.allChanges + exChanges_defined : 𝚺₁-Function₁ exChanges via φ.exChanges variable {V} @@ -693,73 +662,73 @@ variable {β : Blueprint pL} (c : Construction V L β) def Phi (C : Set V) (pr : V) : Prop := ∃ param p y, pr = ⟪param, p, y⟫ ∧ - L.UFormula p ∧ ( - (∃ n k r v, p = ^rel n k r v ∧ y = c.rel param n k r v) ∨ - (∃ n k r v, p = ^nrel n k r v ∧ y = c.nrel param n k r v) ∨ - (∃ n, p = ^⊤[n] ∧ y = c.verum param n) ∨ - (∃ n, p = ^⊥[n] ∧ y = c.falsum param n) ∨ - (∃ n p₁ p₂ y₁ y₂, ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋏[n] p₂ ∧ y = c.and param n p₁ p₂ y₁ y₂) ∨ - (∃ n p₁ p₂ y₁ y₂, ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋎[n] p₂ ∧ y = c.or param n p₁ p₂ y₁ y₂) ∨ - (∃ n p₁ y₁, ⟪c.allChanges param n, p₁, y₁⟫ ∈ C ∧ p = ^∀[n] p₁ ∧ y = c.all param n p₁ y₁) ∨ - (∃ n p₁ y₁, ⟪c.exChanges param n, p₁, y₁⟫ ∈ C ∧ p = ^∃[n] p₁ ∧ y = c.ex param n p₁ y₁) ) + L.IsUFormula p ∧ ( + (∃ k r v, p = ^rel k r v ∧ y = c.rel param k r v) ∨ + (∃ k r v, p = ^nrel k r v ∧ y = c.nrel param k r v) ∨ + (p = ^⊤ ∧ y = c.verum param) ∨ + (p = ^⊥ ∧ y = c.falsum param) ∨ + (∃ p₁ p₂ y₁ y₂, ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋏ p₂ ∧ y = c.and param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ p₂ y₁ y₂, ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋎ p₂ ∧ y = c.or param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ y₁, ⟪c.allChanges param, p₁, y₁⟫ ∈ C ∧ p = ^∀ p₁ ∧ y = c.all param p₁ y₁) ∨ + (∃ p₁ y₁, ⟪c.exChanges param, p₁, y₁⟫ ∈ C ∧ p = ^∃ p₁ ∧ y = c.ex param p₁ y₁) ) private lemma phi_iff (C pr : V) : c.Phi {x | x ∈ C} pr ↔ - ∃ param ≤ pr, ∃ p ≤ pr, ∃ y ≤ pr, pr = ⟪param, p, y⟫ ∧ L.UFormula p ∧ - ((∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, p = ^rel n k R v ∧ y = c.rel param n k R v) ∨ - (∃ n < p, ∃ k < p, ∃ R < p, ∃ v < p, p = ^nrel n k R v ∧ y = c.nrel param n k R v) ∨ - (∃ n < p, p = ^⊤[n] ∧ y = c.verum param n) ∨ - (∃ n < p, p = ^⊥[n] ∧ y = c.falsum param n) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋏[n] p₂ ∧ y = c.and param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, - ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋎[n] p₂ ∧ y = c.or param n p₁ p₂ y₁ y₂) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - ⟪c.allChanges param n, p₁, y₁⟫ ∈ C ∧ p = ^∀[n] p₁ ∧ y = c.all param n p₁ y₁) ∨ - (∃ n < p, ∃ p₁ < p, ∃ y₁ < C, - ⟪c.exChanges param n, p₁, y₁⟫ ∈ C ∧ p = ^∃[n] p₁ ∧ y = c.ex param n p₁ y₁)) := by + ∃ param ≤ pr, ∃ p ≤ pr, ∃ y ≤ pr, pr = ⟪param, p, y⟫ ∧ L.IsUFormula p ∧ + ((∃ k < p, ∃ R < p, ∃ v < p, p = ^rel k R v ∧ y = c.rel param k R v) ∨ + (∃ k < p, ∃ R < p, ∃ v < p, p = ^nrel k R v ∧ y = c.nrel param k R v) ∨ + (p = ^⊤ ∧ y = c.verum param) ∨ + (p = ^⊥ ∧ y = c.falsum param) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋏ p₂ ∧ y = c.and param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ p₂ < p, ∃ y₁ < C, ∃ y₂ < C, + ⟪param, p₁, y₁⟫ ∈ C ∧ ⟪param, p₂, y₂⟫ ∈ C ∧ p = p₁ ^⋎ p₂ ∧ y = c.or param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ < p, ∃ y₁ < C, + ⟪c.allChanges param, p₁, y₁⟫ ∈ C ∧ p = ^∀ p₁ ∧ y = c.all param p₁ y₁) ∨ + (∃ p₁ < p, ∃ y₁ < C, + ⟪c.exChanges param, p₁, y₁⟫ ∈ C ∧ p = ^∃ p₁ ∧ y = c.ex param p₁ y₁)) := by constructor · rintro ⟨param, p, y, rfl, hp, H⟩ refine ⟨param, by simp, p, le_trans (le_pair_left p y) (le_pair_right _ _), y, le_trans (le_pair_right p y) (le_pair_right _ _), rfl, hp, ?_⟩ - rcases H with (⟨n, k, r, v, rfl, rfl⟩ | ⟨n, k, r, v, rfl, rfl⟩ | H) - · left; exact ⟨n, by simp, k, by simp, r, by simp, v, by simp, rfl, rfl⟩ - · right; left; exact ⟨n, by simp, k, by simp, r, by simp, v, by simp, rfl, rfl⟩ + rcases H with (⟨k, r, v, rfl, rfl⟩ | ⟨k, r, v, rfl, rfl⟩ | H) + · left; exact ⟨k, by simp, r, by simp, v, by simp, rfl, rfl⟩ + · right; left; exact ⟨k, by simp, r, by simp, v, by simp, rfl, rfl⟩ right; right - rcases H with (⟨n, rfl, rfl⟩ | ⟨n, rfl, rfl⟩ | H) - · left; exact ⟨n, by simp, rfl, rfl⟩ - · right; left; exact ⟨n, by simp, rfl, rfl⟩ + rcases H with (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ | H) + · left; exact ⟨rfl, rfl⟩ + · right; left; exact ⟨rfl, rfl⟩ right; right - rcases H with (⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | H) - · left; exact ⟨n, by simp, p₁, by simp, p₂, by simp, + rcases H with (⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | H) + · left; exact ⟨p₁, by simp, p₂, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), y₂, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₂), h₁, h₂, rfl, rfl⟩ - · right; left; exact ⟨n, by simp, p₁, by simp, p₂, by simp, + · right; left; exact ⟨p₁, by simp, p₂, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), y₂, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₂), h₁, h₂, rfl, rfl⟩ right; right - rcases H with (⟨n, p₁, y₁, h₁, rfl, rfl⟩ | ⟨n, p₁, y₁, h₁, rfl, rfl⟩) - · left; exact ⟨n, by simp, p₁, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), h₁, rfl, rfl⟩ - · right; exact ⟨n, by simp, p₁, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), h₁, rfl, rfl⟩ + rcases H with (⟨p₁, y₁, h₁, rfl, rfl⟩ | ⟨p₁, y₁, h₁, rfl, rfl⟩) + · left; exact ⟨p₁, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), h₁, rfl, rfl⟩ + · right; exact ⟨p₁, by simp, y₁, lt_of_le_of_lt (by simp) (lt_of_mem_rng h₁), h₁, rfl, rfl⟩ · rintro ⟨param, _, p, _, y, _, rfl, hp, H⟩ refine ⟨param, p, y, rfl, hp, ?_⟩ - rcases H with (⟨n, _, k, _, r, _, v, _, rfl, rfl⟩ | ⟨n, _, k, _, r, _, v, _, rfl, rfl⟩ | H) - · left; exact ⟨n, k, r, v, rfl, rfl⟩ - · right; left; exact ⟨n, k, r, v, rfl, rfl⟩ + rcases H with (⟨k, _, r, _, v, _, rfl, rfl⟩ | ⟨k, _, r, _, v, _, rfl, rfl⟩ | H) + · left; exact ⟨k, r, v, rfl, rfl⟩ + · right; left; exact ⟨k, r, v, rfl, rfl⟩ right; right - rcases H with (⟨n, _, rfl, rfl⟩ | ⟨n, _, rfl, rfl⟩ | H) - · left; exact ⟨n, rfl, rfl⟩ - · right; left; exact ⟨n, rfl, rfl⟩ + rcases H with (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ | H) + · left; exact ⟨rfl, rfl⟩ + · right; left; exact ⟨rfl, rfl⟩ right; right - rcases H with (⟨n, _, p₁, _, p₂, _, y₁, _, y₂, _, h₁, h₂, rfl, rfl⟩ | - ⟨n, _, p₁, _, p₂, _, y₁, _, y₂, _, h₁, h₂, rfl, rfl⟩ | H) - · left; exact ⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ - · right; left; exact ⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ + rcases H with (⟨p₁, _, p₂, _, y₁, _, y₂, _, h₁, h₂, rfl, rfl⟩ | + ⟨p₁, _, p₂, _, y₁, _, y₂, _, h₁, h₂, rfl, rfl⟩ | H) + · left; exact ⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ + · right; left; exact ⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ right; right - rcases H with (⟨n, _, p₁, _, y₁, _, h₁, rfl, rfl⟩ | ⟨n, _, p₁, _, y₁, _, h₁, rfl, rfl⟩) - · left; exact ⟨n, p₁, y₁, h₁, rfl, rfl⟩ - · right; exact ⟨n, p₁, y₁, h₁, rfl, rfl⟩ + rcases H with (⟨p₁, _, y₁, _, h₁, rfl, rfl⟩ | ⟨p₁, _, y₁, _, h₁, rfl, rfl⟩) + · left; exact ⟨p₁, y₁, h₁, rfl, rfl⟩ + · right; exact ⟨p₁, y₁, h₁, rfl, rfl⟩ def construction : Fixpoint.Construction V (β.blueprint) where Φ := fun _ ↦ c.Phi @@ -767,7 +736,7 @@ def construction : Fixpoint.Construction V (β.blueprint) where ⟨by intro v /- simp? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, - eval_uformulaDef L, (uformula_defined L).proper.iff', + L.isUFormula_defined.df.iff, L.isUFormula_defined.proper.iff', c.rel_defined.iff, c.rel_defined.graph_delta.proper.iff', c.nrel_defined.iff, c.nrel_defined.graph_delta.proper.iff', c.verum_defined.iff, c.verum_defined.graph_delta.proper.iff', @@ -780,31 +749,32 @@ def construction : Fixpoint.Construction V (β.blueprint) where c.exChanges_defined.iff, c.exChanges_defined.graph_delta.proper.iff'] -/ simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Blueprint.blueprint, Fin.isValue, - HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.sigma_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, - Semiformula.eval_bexLTSucc', Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, - Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, - LogicalConnective.HomClass.map_and, Semiformula.eval_substs, Matrix.comp_vecCons', - Matrix.cons_val_three, Fin.succ_one_eq_two, Matrix.cons_val_zero, Matrix.cons_val_fin_one, - Matrix.constant_eq_singleton, eval_pair₃Def, eval_uformulaDef L, - LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, Matrix.cons_val_four, - Matrix.cons_val_succ, Matrix.cons_app_five, eval_qqRelDef, Matrix.cons_app_six, - c.rel_defined.iff, LogicalConnective.Prop.and_eq, c.nrel_defined.iff, eval_qqVerumDef, - c.verum_defined.iff, eval_qqFalsumDef, c.falsum_defined.iff, Matrix.cons_app_seven, - Matrix.cons_app_eight, Semiformula.eval_operator₄, TermRec.Construction.cons_app_9, - eval_memRel₃, eval_qqAndDef, c.and_defined.iff, eval_qqOrDef, c.or_defined.iff, - Semiformula.eval_ex, c.allChanges_defined.iff, exists_eq_left, eval_qqAllDef, - c.all_defined.iff, c.exChanges_defined.iff, eval_qqExDef, c.ex_defined.iff, - LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, - (uformula_defined L).proper.iff', c.rel_defined.graph_delta.proper.iff', - HierarchySymbol.Semiformula.graphDelta_val, c.nrel_defined.graph_delta.proper.iff', - c.verum_defined.graph_delta.proper.iff', c.falsum_defined.graph_delta.proper.iff', - c.and_defined.graph_delta.proper.iff', c.or_defined.graph_delta.proper.iff', - Semiformula.eval_all, LogicalConnective.HomClass.map_imply, LogicalConnective.Prop.arrow_eq, - forall_eq, c.all_defined.graph_delta.proper.iff', c.ex_defined.graph_delta.proper.iff'], + HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.sigma_mkDelta, + HierarchySymbol.Semiformula.val_mkSigma, Semiformula.eval_bexLTSucc', Semiterm.val_bvar, + Matrix.cons_val_one, Matrix.vecHead, Matrix.cons_val_two, Matrix.vecTail, + Function.comp_apply, Fin.succ_zero_eq_one, LogicalConnective.HomClass.map_and, + Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_three, Fin.succ_one_eq_two, + Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, eval_pair₃Def, + L.isUFormula_defined.df.iff, LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, + Matrix.cons_val_four, Matrix.cons_val_succ, eval_qqRelDef, Matrix.cons_app_five, + c.rel_defined.iff, LogicalConnective.Prop.and_eq, eval_qqNRelDef, c.nrel_defined.iff, + eval_qqVerumDef, c.verum_defined.iff, eval_qqFalsumDef, c.falsum_defined.iff, + Matrix.cons_app_six, Matrix.cons_app_seven, Semiformula.eval_operator₄, + Matrix.cons_app_eight, eval_memRel₃, eval_qqAndDef, c.and_defined.iff, eval_qqOrDef, + c.or_defined.iff, Semiformula.eval_ex, c.allChanges_defined.iff, exists_eq_left, + eval_qqAllDef, c.all_defined.iff, c.exChanges_defined.iff, eval_qqExDef, c.ex_defined.iff, + LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, + HierarchySymbol.Semiformula.val_mkPi, L.isUFormula_defined.proper.iff', + c.rel_defined.graph_delta.proper.iff', HierarchySymbol.Semiformula.graphDelta_val, + c.nrel_defined.graph_delta.proper.iff', c.verum_defined.graph_delta.proper.iff', + c.falsum_defined.graph_delta.proper.iff', c.and_defined.graph_delta.proper.iff', + c.or_defined.graph_delta.proper.iff', Semiformula.eval_all, + LogicalConnective.HomClass.map_imply, LogicalConnective.Prop.arrow_eq, forall_eq, + c.all_defined.graph_delta.proper.iff', c.ex_defined.graph_delta.proper.iff'], by intro v /- - simpa [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, - eval_uformulaDef L, + simpa? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, + L.isUFormula_defined.df.iff, c.rel_defined.iff, c.nrel_defined.iff, c.verum_defined.iff, @@ -817,21 +787,22 @@ def construction : Fixpoint.Construction V (β.blueprint) where c.exChanges_defined.iff] using c.phi_iff _ _ -/ simpa only [Fin.isValue, Nat.succ_eq_add_one, Nat.reduceAdd, Blueprint.blueprint, - HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.val_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, - Semiformula.eval_bexLTSucc', Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, - Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, - LogicalConnective.HomClass.map_and, Semiformula.eval_substs, Matrix.comp_vecCons', - Matrix.cons_val_three, Fin.succ_one_eq_two, Matrix.cons_val_zero, Matrix.cons_val_fin_one, - Matrix.constant_eq_singleton, eval_pair₃Def, eval_uformulaDef L, - LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, Matrix.cons_val_four, - Matrix.cons_val_succ, Matrix.cons_app_five, eval_qqRelDef, Matrix.cons_app_six, - c.rel_defined.iff, LogicalConnective.Prop.and_eq, eval_qqNRelDef, c.nrel_defined.iff, - eval_qqVerumDef, c.verum_defined.iff, eval_qqFalsumDef, c.falsum_defined.iff, - Matrix.cons_app_seven, Matrix.cons_app_eight, Semiformula.eval_operator₄, - TermRec.Construction.cons_app_9, eval_memRel₃, eval_qqAndDef, c.and_defined.iff, - eval_qqOrDef, c.or_defined.iff, Semiformula.eval_ex, c.allChanges_defined.iff, - exists_eq_left, eval_qqAllDef, c.all_defined.iff, c.exChanges_defined.iff, eval_qqExDef, - c.ex_defined.iff, LogicalConnective.Prop.or_eq] using c.phi_iff _ _⟩ + HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.val_mkDelta, + HierarchySymbol.Semiformula.val_mkSigma, Semiformula.eval_bexLTSucc', Semiterm.val_bvar, + Matrix.cons_val_one, Matrix.vecHead, Matrix.cons_val_two, Matrix.vecTail, + Function.comp_apply, Fin.succ_zero_eq_one, LogicalConnective.HomClass.map_and, + Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_three, Fin.succ_one_eq_two, + Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, + eval_pair₃Def, L.isUFormula_defined.df.iff, LogicalConnective.HomClass.map_or, + Semiformula.eval_bexLT, Matrix.cons_val_four, Matrix.cons_val_succ, eval_qqRelDef, + Matrix.cons_app_five, c.rel_defined.iff, LogicalConnective.Prop.and_eq, eval_qqNRelDef, + c.nrel_defined.iff, eval_qqVerumDef, c.verum_defined.iff, eval_qqFalsumDef, + c.falsum_defined.iff, Matrix.cons_app_six, Matrix.cons_app_seven, + Semiformula.eval_operator₄, Matrix.cons_app_eight, eval_memRel₃, eval_qqAndDef, + c.and_defined.iff, eval_qqOrDef, c.or_defined.iff, Semiformula.eval_ex, + c.allChanges_defined.iff, exists_eq_left, eval_qqAllDef, c.all_defined.iff, + c.exChanges_defined.iff, eval_qqExDef, c.ex_defined.iff, + LogicalConnective.Prop.or_eq] using c.phi_iff _ _⟩ monotone := by unfold Phi rintro C C' hC _ _ ⟨param, p, y, rfl, hp, H⟩ @@ -842,34 +813,34 @@ def construction : Fixpoint.Construction V (β.blueprint) where · right; right; left; exact h · right; right; right; left; exact h right; right; right; right - rcases H with (⟨n, p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩ | ⟨n, p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩ | H) - · left; exact ⟨n, p₁, p₂, r₁, r₂, hC h₁, hC h₂, rfl, rfl⟩ - · right; left; exact ⟨n, p₁, p₂, r₁, r₂, hC h₁, hC h₂, rfl, rfl⟩ + rcases H with (⟨p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩ | ⟨p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩ | H) + · left; exact ⟨p₁, p₂, r₁, r₂, hC h₁, hC h₂, rfl, rfl⟩ + · right; left; exact ⟨p₁, p₂, r₁, r₂, hC h₁, hC h₂, rfl, rfl⟩ right; right - rcases H with (⟨n, p₁, r₁, h₁, rfl, rfl⟩ | ⟨n, p₁, r₁, h₁, rfl, rfl⟩) - · left; exact ⟨n, p₁, r₁, hC h₁, rfl, rfl⟩ - · right; exact ⟨n, p₁, r₁, hC h₁, rfl, rfl⟩ + rcases H with (⟨p₁, r₁, h₁, rfl, rfl⟩ | ⟨p₁, r₁, h₁, rfl, rfl⟩) + · left; exact ⟨p₁, r₁, hC h₁, rfl, rfl⟩ + · right; exact ⟨p₁, r₁, hC h₁, rfl, rfl⟩ instance : c.construction.Finite where finite {C _ pr h} := by rcases h with ⟨param, p, y, rfl, hp, (h | h | h | h | - ⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨n, p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨n, p₁, y₁, h₁, rfl, rfl⟩ | ⟨n, p₁, y₁, h₁, rfl, rfl⟩)⟩ + ⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨p₁, p₂, y₁, y₂, h₁, h₂, rfl, rfl⟩ | ⟨p₁, y₁, h₁, rfl, rfl⟩ | ⟨p₁, y₁, h₁, rfl, rfl⟩)⟩ · exact ⟨0, param, _, _, rfl, hp, Or.inl h⟩ · exact ⟨0, param, _, _, rfl, hp, Or.inr <| Or.inl h⟩ · exact ⟨0, param, _, _, rfl, hp, Or.inr <| Or.inr <| Or.inl h⟩ · exact ⟨0, param, _, _, rfl, hp, Or.inr <| Or.inr <| Or.inr <| Or.inl h⟩ · exact ⟨Max.max ⟪param, p₁, y₁⟫ ⟪param, p₂, y₂⟫ + 1, param, _, _, rfl, hp, by right; right; right; right; left - exact ⟨n, p₁, p₂, y₁, y₂, by simp [h₁, lt_succ_iff_le], by simp [h₂, lt_succ_iff_le], rfl, rfl⟩⟩ + exact ⟨p₁, p₂, y₁, y₂, by simp [h₁, lt_succ_iff_le], by simp [h₂, lt_succ_iff_le], rfl, rfl⟩⟩ · exact ⟨Max.max ⟪param, p₁, y₁⟫ ⟪param, p₂, y₂⟫ + 1, param, _, _, rfl, hp, by right; right; right; right; right; left - exact ⟨n, p₁, p₂, y₁, y₂, by simp [h₁, lt_succ_iff_le], by simp [h₂, lt_succ_iff_le], rfl, rfl⟩⟩ - · exact ⟨⟪c.allChanges param n, p₁, y₁⟫ + 1, param, _, _, rfl, hp, by + exact ⟨p₁, p₂, y₁, y₂, by simp [h₁, lt_succ_iff_le], by simp [h₂, lt_succ_iff_le], rfl, rfl⟩⟩ + · exact ⟨⟪c.allChanges param, p₁, y₁⟫ + 1, param, _, _, rfl, hp, by right; right; right; right; right; right; left - exact ⟨n, p₁, y₁, by simp [h₁], rfl, rfl⟩⟩ - · exact ⟨⟪c.exChanges param n, p₁, y₁⟫ + 1, param, _, _, rfl, hp, by + exact ⟨p₁, y₁, by simp [h₁], rfl, rfl⟩⟩ + · exact ⟨⟪c.exChanges param, p₁, y₁⟫ + 1, param, _, _, rfl, hp, by right; right; right; right; right; right; right - exact ⟨n, p₁, y₁, by simp [h₁], rfl, rfl⟩⟩ + exact ⟨p₁, y₁, by simp [h₁], rfl, rfl⟩⟩ def Graph (param : V) (x y : V) : Prop := c.construction.Fixpoint ![] ⟪param, x, y⟫ @@ -879,15 +850,15 @@ variable {c} lemma Graph.case_iff {p y : V} : c.Graph param p y ↔ - L.UFormula p ∧ ( - (∃ n k R v, p = ^rel n k R v ∧ y = c.rel param n k R v) ∨ - (∃ n k R v, p = ^nrel n k R v ∧ y = c.nrel param n k R v) ∨ - (∃ n, p = ^⊤[n] ∧ y = c.verum param n) ∨ - (∃ n, p = ^⊥[n] ∧ y = c.falsum param n) ∨ - (∃ n p₁ p₂ y₁ y₂, c.Graph param p₁ y₁ ∧ c.Graph param p₂ y₂ ∧ p = p₁ ^⋏[n] p₂ ∧ y = c.and param n p₁ p₂ y₁ y₂) ∨ - (∃ n p₁ p₂ y₁ y₂, c.Graph param p₁ y₁ ∧ c.Graph param p₂ y₂ ∧ p = p₁ ^⋎[n] p₂ ∧ y = c.or param n p₁ p₂ y₁ y₂) ∨ - (∃ n p₁ y₁, c.Graph (c.allChanges param n) p₁ y₁ ∧ p = ^∀[n] p₁ ∧ y = c.all param n p₁ y₁) ∨ - (∃ n p₁ y₁, c.Graph (c.exChanges param n) p₁ y₁ ∧ p = ^∃[n] p₁ ∧ y = c.ex param n p₁ y₁) ) := + L.IsUFormula p ∧ ( + (∃ k R v, p = ^rel k R v ∧ y = c.rel param k R v) ∨ + (∃ k R v, p = ^nrel k R v ∧ y = c.nrel param k R v) ∨ + (p = ^⊤ ∧ y = c.verum param) ∨ + (p = ^⊥ ∧ y = c.falsum param) ∨ + (∃ p₁ p₂ y₁ y₂, c.Graph param p₁ y₁ ∧ c.Graph param p₂ y₂ ∧ p = p₁ ^⋏ p₂ ∧ y = c.and param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ p₂ y₁ y₂, c.Graph param p₁ y₁ ∧ c.Graph param p₂ y₂ ∧ p = p₁ ^⋎ p₂ ∧ y = c.or param p₁ p₂ y₁ y₂) ∨ + (∃ p₁ y₁, c.Graph (c.allChanges param) p₁ y₁ ∧ p = ^∀ p₁ ∧ y = c.all param p₁ y₁) ∨ + (∃ p₁ y₁, c.Graph (c.exChanges param) p₁ y₁ ∧ p = ^∃ p₁ ∧ y = c.ex param p₁ y₁) ) := Iff.trans c.construction.case (by constructor · rintro ⟨param, p, y, e, H⟩; @@ -908,14 +879,14 @@ instance graph_definable : 𝚺-[0 + 1]-Relation₃ c.Graph := (c.graph_defined) variable {β} lemma graph_dom_uformula {p r} : - c.Graph param p r → L.UFormula p := fun h ↦ Graph.case_iff.mp h |>.1 + c.Graph param p r → L.IsUFormula p := fun h ↦ Graph.case_iff.mp h |>.1 -lemma graph_rel_iff {n k r v y} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : - c.Graph param (^rel n k r v) y ↔ y = c.rel param n k r v := by +lemma graph_rel_iff {k r v y} (hkr : L.Rel k r) (hv : L.IsUTermVec k v) : + c.Graph param (^rel k r v) y ↔ y = c.rel param k r v := by constructor · intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨n, k, r, v, H, rfl⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨k, r, v, H, rfl⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqRel] at H; rcases H with ⟨rfl, rfl, rfl, rfl⟩; rfl · simp [qqRel, qqNRel] at H · simp [qqRel, qqVerum] at H @@ -924,14 +895,14 @@ lemma graph_rel_iff {n k r v y} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : · simp [qqRel, qqOr] at H · simp [qqRel, qqAll] at H · simp [qqRel, qqEx] at H - · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inl ⟨n, k, r, v, rfl, rfl⟩⟩ + · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inl ⟨k, r, v, rfl, rfl⟩⟩ -lemma graph_nrel_iff {n k r v y} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : - c.Graph param (^nrel n k r v) y ↔ y = c.nrel param n k r v := by +lemma graph_nrel_iff {k r v y} (hkr : L.Rel k r) (hv : L.IsUTermVec k v) : + c.Graph param (^nrel k r v) y ↔ y = c.nrel param k r v := by constructor · intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, rfl⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, rfl⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqNRel, qqRel] at H · simp [qqNRel] at H; rcases H with ⟨rfl, rfl, rfl, rfl⟩; rfl · simp [qqNRel, qqVerum] at H @@ -940,14 +911,14 @@ lemma graph_nrel_iff {n k r v y} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : · simp [qqNRel, qqOr] at H · simp [qqNRel, qqAll] at H · simp [qqNRel, qqEx] at H - · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inr <| Or.inl ⟨n, k, r, v, rfl, rfl⟩⟩ + · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inr <| Or.inl ⟨k, r, v, rfl, rfl⟩⟩ -lemma graph_verum_iff {n y} : - c.Graph param ^⊤[n] y ↔ y = c.verum param n := by +lemma graph_verum_iff {y} : + c.Graph param ^⊤ y ↔ y = c.verum param := by constructor · intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, rfl⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, rfl⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqVerum, qqRel] at H · simp [qqVerum, qqNRel] at H · simp [qqVerum, qqVerum] at H; rcases H; rfl @@ -956,14 +927,14 @@ lemma graph_verum_iff {n y} : · simp [qqVerum, qqOr] at H · simp [qqVerum, qqAll] at H · simp [qqVerum, qqEx] at H - · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inl ⟨n, rfl, rfl⟩⟩ + · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inl ⟨rfl, rfl⟩⟩ -lemma graph_falsum_iff {n y} : - c.Graph param ^⊥[n] y ↔ y = c.falsum param n := by +lemma graph_falsum_iff {y} : + c.Graph param ^⊥ y ↔ y = c.falsum param := by constructor · intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, rfl⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, rfl⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqFalsum, qqRel] at H · simp [qqFalsum, qqNRel] at H · simp [qqFalsum, qqVerum] at H @@ -972,35 +943,32 @@ lemma graph_falsum_iff {n y} : · simp [qqFalsum, qqOr] at H · simp [qqFalsum, qqAll] at H · simp [qqFalsum, qqEx] at H - · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, rfl, rfl⟩⟩ + · rintro rfl; exact (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨rfl, rfl⟩⟩ -lemma graph_rel {n k r v} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : - c.Graph param (^rel n k r v) (c.rel param n k r v) := - (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inl ⟨n, k, r, v, rfl, rfl⟩⟩ +lemma graph_rel {k r v} (hkr : L.Rel k r) (hv : L.IsUTermVec k v) : + c.Graph param (^rel k r v) (c.rel param k r v) := + (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inl ⟨k, r, v, rfl, rfl⟩⟩ -lemma graph_nrel {n k r v} (hkr : L.Rel k r) (hv : L.SemitermVec k n v) : - c.Graph param (^nrel n k r v) (c.nrel param n k r v) := - (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inr <| Or.inl ⟨n, k, r, v, rfl, rfl⟩⟩ +lemma graph_nrel {k r v} (hkr : L.Rel k r) (hv : L.IsUTermVec k v) : + c.Graph param (^nrel k r v) (c.nrel param k r v) := + (Graph.case_iff).mpr ⟨by simp [hkr, hv], Or.inr <| Or.inl ⟨k, r, v, rfl, rfl⟩⟩ -lemma graph_verum (n : V) : - c.Graph param (^⊤[n]) (c.verum param n) := - (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inl ⟨n, rfl, rfl⟩⟩ +lemma graph_verum : + c.Graph param ^⊤ (c.verum param) := (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inl ⟨rfl, rfl⟩⟩ -lemma graph_falsum (n : V) : - c.Graph param (^⊥[n]) (c.falsum param n) := - (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, rfl, rfl⟩⟩ +lemma graph_falsum : + c.Graph param ^⊥ (c.falsum param) := (Graph.case_iff).mpr ⟨by simp, Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨rfl, rfl⟩⟩ -lemma graph_and {n p₁ p₂ r₁ r₂ : V} (hp₁ : L.Semiformula n p₁) (hp₂ : L.Semiformula n p₂) +lemma graph_and {p₁ p₂ r₁ r₂ : V} (hp₁ : L.IsUFormula p₁) (hp₂ : L.IsUFormula p₂) (h₁ : c.Graph param p₁ r₁) (h₂ : c.Graph param p₂ r₂) : - c.Graph param (p₁ ^⋏[n] p₂) (c.and param n p₁ p₂ r₁ r₂) := - (Graph.case_iff).mpr ⟨by simp [hp₁, hp₂], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, - p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩⟩ + c.Graph param (p₁ ^⋏ p₂) (c.and param p₁ p₂ r₁ r₂) := + (Graph.case_iff).mpr ⟨by simp [hp₁, hp₂], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩⟩ -lemma graph_and_inv {n p₁ p₂ r : V} : - c.Graph param (p₁ ^⋏[n] p₂) r → ∃ r₁ r₂, c.Graph param p₁ r₁ ∧ c.Graph param p₂ r₂ ∧ r = c.and param n p₁ p₂ r₁ r₂ := by +lemma graph_and_inv {p₁ p₂ r : V} : + c.Graph param (p₁ ^⋏ p₂) r → ∃ r₁ r₂, c.Graph param p₁ r₁ ∧ c.Graph param p₂ r₂ ∧ r = c.and param p₁ p₂ r₁ r₂ := by intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, rfl⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, rfl⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqAnd, qqRel] at H · simp [qqAnd, qqNRel] at H · simp [qqAnd, qqVerum] at H @@ -1011,17 +979,16 @@ lemma graph_and_inv {n p₁ p₂ r : V} : · simp [qqAnd, qqAll] at H · simp [qqAnd, qqEx] at H -lemma graph_or {n p₁ p₂ r₁ r₂ : V} (hp₁ : L.Semiformula n p₁) (hp₂ : L.Semiformula n p₂) +lemma graph_or {p₁ p₂ r₁ r₂ : V} (hp₁ : L.IsUFormula p₁) (hp₂ : L.IsUFormula p₂) (h₁ : c.Graph param p₁ r₁) (h₂ : c.Graph param p₂ r₂) : - c.Graph param (p₁ ^⋎[n] p₂) (c.or param n p₁ p₂ r₁ r₂) := - (Graph.case_iff).mpr ⟨by simp [hp₁, hp₂], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, - p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩⟩ + c.Graph param (p₁ ^⋎ p₂) (c.or param p₁ p₂ r₁ r₂) := + (Graph.case_iff).mpr ⟨by simp [hp₁, hp₂], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, p₂, r₁, r₂, h₁, h₂, rfl, rfl⟩⟩ -lemma graph_or_inv {n p₁ p₂ r : V} : - c.Graph param (p₁ ^⋎[n] p₂) r → ∃ r₁ r₂, c.Graph param p₁ r₁ ∧ c.Graph param p₂ r₂ ∧ r = c.or param n p₁ p₂ r₁ r₂ := by +lemma graph_or_inv {p₁ p₂ r : V} : + c.Graph param (p₁ ^⋎ p₂) r → ∃ r₁ r₂, c.Graph param p₁ r₁ ∧ c.Graph param p₂ r₂ ∧ r = c.or param p₁ p₂ r₁ r₂ := by intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, rfl⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, rfl⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqOr, qqRel] at H · simp [qqOr, qqNRel] at H · simp [qqOr, qqVerum] at H @@ -1032,16 +999,15 @@ lemma graph_or_inv {n p₁ p₂ r : V} : · simp [qqOr, qqAll] at H · simp [qqOr, qqEx] at H -lemma graph_all {n p₁ r₁ : V} (hp₁ : L.Semiformula (n + 1) p₁) (h₁ : c.Graph (c.allChanges param n) p₁ r₁) : - c.Graph param (^∀[n] p₁) (c.all param n p₁ r₁) := - (Graph.case_iff).mpr ⟨by simp [hp₁], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨n, - p₁, r₁, h₁, rfl, rfl⟩⟩ +lemma graph_all {p₁ r₁ : V} (hp₁ : L.IsUFormula p₁) (h₁ : c.Graph (c.allChanges param) p₁ r₁) : + c.Graph param (^∀ p₁) (c.all param p₁ r₁) := + (Graph.case_iff).mpr ⟨by simp [hp₁], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, r₁, h₁, rfl, rfl⟩⟩ -lemma graph_all_inv {n p₁ r : V} : - c.Graph param (^∀[n] p₁) r → ∃ r₁, c.Graph (c.allChanges param n) p₁ r₁ ∧ r = c.all param n p₁ r₁ := by +lemma graph_all_inv {p₁ r : V} : + c.Graph param (^∀ p₁) r → ∃ r₁, c.Graph (c.allChanges param) p₁ r₁ ∧ r = c.all param p₁ r₁ := by intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, rfl⟩ | ⟨_, _, _, _, H, _⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, rfl⟩ | ⟨_, _, _, H, _⟩)⟩ · simp [qqAll, qqRel] at H · simp [qqAll, qqNRel] at H · simp [qqAll, qqVerum] at H @@ -1052,16 +1018,15 @@ lemma graph_all_inv {n p₁ r : V} : exact ⟨_, by assumption, rfl⟩ · simp [qqAll, qqEx] at H -lemma graph_ex {n p₁ r₁ : V} (hp₁ : L.Semiformula (n + 1) p₁) (h₁ : c.Graph (c.exChanges param n) p₁ r₁) : - c.Graph param (^∃[n] p₁) (c.ex param n p₁ r₁) := - (Graph.case_iff).mpr ⟨by simp [hp₁], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr ⟨n, - p₁, r₁, h₁, rfl, rfl⟩⟩ +lemma graph_ex {p₁ r₁ : V} (hp₁ : L.IsUFormula p₁) (h₁ : c.Graph (c.exChanges param) p₁ r₁) : + c.Graph param (^∃ p₁) (c.ex param p₁ r₁) := + (Graph.case_iff).mpr ⟨by simp [hp₁], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr ⟨p₁, r₁, h₁, rfl, rfl⟩⟩ -lemma graph_ex_inv {n p₁ r : V} : - c.Graph param (^∃[n] p₁) r → ∃ r₁, c.Graph (c.exChanges param n) p₁ r₁ ∧ r = c.ex param n p₁ r₁ := by +lemma graph_ex_inv {p₁ r : V} : + c.Graph param (^∃ p₁) r → ∃ r₁, c.Graph (c.exChanges param) p₁ r₁ ∧ r = c.ex param p₁ r₁ := by intro h - rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, H, _⟩ | ⟨_, H, _⟩ | - ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, H, _⟩ | ⟨_, _, _, _, H, rfl⟩)⟩ + rcases Graph.case_iff.mp h with ⟨_, (⟨_, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨H, _⟩ | ⟨H, _⟩ | + ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, _, _, _, H, _⟩ | ⟨_, _, _, H, _⟩ | ⟨_, _, _, H, rfl⟩)⟩ · simp [qqEx, qqRel] at H · simp [qqEx, qqNRel] at H · simp [qqEx, qqVerum] at H @@ -1074,81 +1039,81 @@ lemma graph_ex_inv {n p₁ r : V} : variable (param) -lemma graph_exists {p : V} : L.UFormula p → ∃ y, c.Graph param p y := by - haveI : 𝚺₁-Function₂ c.allChanges := c.allChanges_defined.to_definable - haveI : 𝚺₁-Function₂ c.exChanges := c.exChanges_defined.to_definable - let f : V → V → V := fun p param ↦ max param (max (c.allChanges param (fstIdx p)) (c.exChanges param (fstIdx p))) +lemma graph_exists {p : V} : L.IsUFormula p → ∃ y, c.Graph param p y := by + haveI : 𝚺₁-Function₁ c.allChanges := c.allChanges_defined.to_definable + haveI : 𝚺₁-Function₁ c.exChanges := c.exChanges_defined.to_definable + let f : V → V → V := fun _ param ↦ Max.max param (Max.max (c.allChanges param) (c.exChanges param)) have hf : 𝚺₁-Function₂ f := by simp [f]; definability - apply sigma₁_order_ball_induction hf ?_ ?_ p param + apply order_ball_induction_sigma1 hf ?_ ?_ p param · definability intro p param ih hp rcases hp.case with - (⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, k, r, v, hkr, hv, rfl⟩ | - ⟨n, rfl⟩ | ⟨n, rfl⟩ | - ⟨n, p₁, p₂, hp₁, hp₂, rfl⟩ | ⟨n, p₁, p₂, hp₁, hp₂, rfl⟩ | - ⟨n, p₁, hp₁, rfl⟩ | ⟨n, p₁, hp₁, rfl⟩) - · exact ⟨c.rel param n k r v, c.graph_rel hkr hv⟩ - · exact ⟨c.nrel param n k r v, c.graph_nrel hkr hv⟩ - · exact ⟨c.verum param n, c.graph_verum n⟩ - · exact ⟨c.falsum param n, c.graph_falsum n⟩ - · rcases ih p₁ (by simp) param (by simp [f]) hp₁.1 with ⟨y₁, h₁⟩ - rcases ih p₂ (by simp) param (by simp [f]) hp₂.1 with ⟨y₂, h₂⟩ - exact ⟨c.and param n p₁ p₂ y₁ y₂, c.graph_and hp₁ hp₂ h₁ h₂⟩ - · rcases ih p₁ (by simp) param (by simp [f]) hp₁.1 with ⟨y₁, h₁⟩ - rcases ih p₂ (by simp) param (by simp [f]) hp₂.1 with ⟨y₂, h₂⟩ - exact ⟨c.or param n p₁ p₂ y₁ y₂, c.graph_or hp₁ hp₂ h₁ h₂⟩ - · rcases ih p₁ (by simp) (c.allChanges param n) (by simp [f]) hp₁.1 with ⟨y₁, h₁⟩ - exact ⟨c.all param n p₁ y₁, c.graph_all hp₁ h₁⟩ - · rcases ih p₁ (by simp) (c.exChanges param n) (by simp [f]) hp₁.1 with ⟨y₁, h₁⟩ - exact ⟨c.ex param n p₁ y₁, c.graph_ex hp₁ h₁⟩ - -lemma graph_unique {p : V} : L.UFormula p → ∀ {param r r'}, c.Graph param p r → c.Graph param p r' → r = r' := by - apply Language.UFormula.induction 𝚷 (P := fun p ↦ ∀ {param r r'}, c.Graph param p r → c.Graph param p r' → r = r') + (⟨k, r, v, hkr, hv, rfl⟩ | ⟨k, r, v, hkr, hv, rfl⟩ | + rfl | rfl | + ⟨p₁, p₂, hp₁, hp₂, rfl⟩ | ⟨p₁, p₂, hp₁, hp₂, rfl⟩ | + ⟨p₁, hp₁, rfl⟩ | ⟨p₁, hp₁, rfl⟩) + · exact ⟨c.rel param k r v, c.graph_rel hkr hv⟩ + · exact ⟨c.nrel param k r v, c.graph_nrel hkr hv⟩ + · exact ⟨c.verum param, c.graph_verum⟩ + · exact ⟨c.falsum param, c.graph_falsum⟩ + · rcases ih p₁ (by simp) param (by simp [f]) hp₁ with ⟨y₁, h₁⟩ + rcases ih p₂ (by simp) param (by simp [f]) hp₂ with ⟨y₂, h₂⟩ + exact ⟨c.and param p₁ p₂ y₁ y₂, c.graph_and hp₁ hp₂ h₁ h₂⟩ + · rcases ih p₁ (by simp) param (by simp [f]) hp₁ with ⟨y₁, h₁⟩ + rcases ih p₂ (by simp) param (by simp [f]) hp₂ with ⟨y₂, h₂⟩ + exact ⟨c.or param p₁ p₂ y₁ y₂, c.graph_or hp₁ hp₂ h₁ h₂⟩ + · rcases ih p₁ (by simp) (c.allChanges param) (by simp [f]) hp₁ with ⟨y₁, h₁⟩ + exact ⟨c.all param p₁ y₁, c.graph_all hp₁ h₁⟩ + · rcases ih p₁ (by simp) (c.exChanges param) (by simp [f]) hp₁ with ⟨y₁, h₁⟩ + exact ⟨c.ex param p₁ y₁, c.graph_ex hp₁ h₁⟩ + +lemma graph_unique {p : V} : L.IsUFormula p → ∀ {param r r'}, c.Graph param p r → c.Graph param p r' → r = r' := by + apply Language.IsUFormula.induction_pi1 (P := fun p ↦ ∀ {param r r'}, c.Graph param p r → c.Graph param p r' → r = r') (by definability) case hrel => - intro n k R v hkR hv + intro k R v hkR hv simp [c.graph_rel_iff hkR hv] case hnrel => - intro n k R v hkR hv + intro k R v hkR hv simp [c.graph_nrel_iff hkR hv] case hverum => - intro n; simp [c.graph_verum_iff] + simp [c.graph_verum_iff] case hfalsum => - intro n; simp [c.graph_falsum_iff] + simp [c.graph_falsum_iff] case hand => - intro n p₁ p₂ _ _ ih₁ ih₂ param r r' hr hr' + intro p₁ p₂ _ _ ih₁ ih₂ param r r' hr hr' rcases c.graph_and_inv hr with ⟨r₁, r₂, h₁, h₂, rfl⟩ rcases c.graph_and_inv hr' with ⟨r₁', r₂', h₁', h₂', rfl⟩ rcases ih₁ h₁ h₁'; rcases ih₂ h₂ h₂'; rfl case hor => - intro n p₁ p₂ _ _ ih₁ ih₂ param r r' hr hr' + intro p₁ p₂ _ _ ih₁ ih₂ param r r' hr hr' rcases c.graph_or_inv hr with ⟨r₁, r₂, h₁, h₂, rfl⟩ rcases c.graph_or_inv hr' with ⟨r₁', r₂', h₁', h₂', rfl⟩ rcases ih₁ h₁ h₁'; rcases ih₂ h₂ h₂'; rfl case hall => - intro n p _ ih param r r' hr hr' + intro p _ ih param r r' hr hr' rcases c.graph_all_inv hr with ⟨r₁, h₁, rfl⟩ rcases c.graph_all_inv hr' with ⟨r₁', h₁', rfl⟩ rcases ih h₁ h₁'; rfl case hex => - intro n p _ ih param r r' hr hr' + intro p _ ih param r r' hr hr' rcases c.graph_ex_inv hr with ⟨r₁, h₁, rfl⟩ rcases c.graph_ex_inv hr' with ⟨r₁', h₁', rfl⟩ rcases ih h₁ h₁'; rfl -lemma exists_unique {p : V} (hp : L.UFormula p) : ∃! r, c.Graph param p r := by +lemma exists_unique {p : V} (hp : L.IsUFormula p) : ∃! r, c.Graph param p r := by rcases c.graph_exists param hp with ⟨r, hr⟩ exact ExistsUnique.intro r hr (fun r' hr' ↦ c.graph_unique hp hr' hr) -lemma exists_unique_all (p : V) : ∃! r, (L.UFormula p → c.Graph param p r) ∧ (¬L.UFormula p → r = 0) := by - by_cases hp : L.UFormula p <;> simp [hp, exists_unique] +lemma exists_unique_all (p : V) : ∃! r, (L.IsUFormula p → c.Graph param p r) ∧ (¬L.IsUFormula p → r = 0) := by + by_cases hp : L.IsUFormula p <;> simp [hp, exists_unique] def result (p : V) : V := Classical.choose! (c.exists_unique_all param p) -lemma result_prop {p : V} (hp : L.UFormula p) : c.Graph param p (c.result param p) := +lemma result_prop {p : V} (hp : L.IsUFormula p) : c.Graph param p (c.result param p) := Classical.choose!_spec (c.exists_unique_all param p) |>.1 hp -lemma result_prop_not {p : V} (hp : ¬L.UFormula p) : c.result param p = 0 := +lemma result_prop_not {p : V} (hp : ¬L.IsUFormula p) : c.result param p = 0 := Classical.choose!_spec (c.exists_unique_all param p) |>.2 hp variable {param} @@ -1156,75 +1121,70 @@ variable {param} lemma result_eq_of_graph {p r} (h : c.Graph param p r) : c.result param p = r := Eq.symm <| Classical.choose_uniq (c.exists_unique_all param p) (by simp [c.graph_dom_uformula h, h]) -@[simp] lemma result_rel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - c.result param (^rel n k R v) = c.rel param n k R v := +@[simp] lemma result_rel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + c.result param (^rel k R v) = c.rel param k R v := c.result_eq_of_graph (c.graph_rel hR hv) -@[simp] lemma result_nrel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - c.result param (^nrel n k R v) = c.nrel param n k R v := +@[simp] lemma result_nrel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + c.result param (^nrel k R v) = c.nrel param k R v := c.result_eq_of_graph (c.graph_nrel hR hv) -@[simp] lemma result_verum {n} : c.result param ^⊤[n] = c.verum param n := c.result_eq_of_graph (c.graph_verum n) +@[simp] lemma result_verum : c.result param ^⊤ = c.verum param := c.result_eq_of_graph c.graph_verum -@[simp] lemma result_falsum {n} : c.result param ^⊥[n] = c.falsum param n := c.result_eq_of_graph (c.graph_falsum n) +@[simp] lemma result_falsum : c.result param ^⊥ = c.falsum param := c.result_eq_of_graph c.graph_falsum -@[simp] lemma result_and {n p q} - (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - c.result param (p ^⋏[n] q) = c.and param n p q (c.result param p) (c.result param q) := - c.result_eq_of_graph (c.graph_and hp hq (c.result_prop param hp.1) (c.result_prop param hq.1)) +@[simp] lemma result_and {p q} + (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + c.result param (p ^⋏ q) = c.and param p q (c.result param p) (c.result param q) := + c.result_eq_of_graph (c.graph_and hp hq (c.result_prop param hp) (c.result_prop param hq)) -@[simp] lemma result_or {n p q} - (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - c.result param (p ^⋎[n] q) = c.or param n p q (c.result param p) (c.result param q) := - c.result_eq_of_graph (c.graph_or hp hq (c.result_prop param hp.1) (c.result_prop param hq.1)) +@[simp] lemma result_or {p q} + (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + c.result param (p ^⋎ q) = c.or param p q (c.result param p) (c.result param q) := + c.result_eq_of_graph (c.graph_or hp hq (c.result_prop param hp) (c.result_prop param hq)) -@[simp] lemma result_all {n p} (hp : L.Semiformula (n + 1) p) : - c.result param (^∀[n] p) = c.all param n p (c.result (c.allChanges param n) p) := - c.result_eq_of_graph (c.graph_all hp (c.result_prop (c.allChanges param n) hp.1)) +@[simp] lemma result_all {p} (hp : L.IsUFormula p) : + c.result param (^∀ p) = c.all param p (c.result (c.allChanges param) p) := + c.result_eq_of_graph (c.graph_all hp (c.result_prop (c.allChanges param) hp)) -@[simp] lemma result_ex {n p} (hp : L.Semiformula (n + 1) p) : - c.result param (^∃[n] p) = c.ex param n p (c.result (c.exChanges param n) p) := - c.result_eq_of_graph (c.graph_ex hp (c.result_prop _ hp.1)) +@[simp] lemma result_ex {p} (hp : L.IsUFormula p) : + c.result param (^∃ p) = c.ex param p (c.result (c.exChanges param) p) := + c.result_eq_of_graph (c.graph_ex hp (c.result_prop _ hp)) section lemma result_defined : 𝚺₁-Function₂ c.result via β.result := by intro v - simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, eval_uformulaDef L, (uformula_defined L).proper.iff', c.eval_graphDef] + simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, L.isUFormula_defined.df.iff, L.isUFormula_defined.proper.iff', c.eval_graphDef] exact Classical.choose!_eq_iff (c.exists_unique_all (v 1) (v 2)) -@[definability] instance result_definable : 𝚺₁-Function₂ c.result := c.result_defined.to_definable _ +instance result_definable : 𝚺-[0 + 1]-Function₂ c.result := c.result_defined.to_definable end lemma uformula_result_induction {P : V → V → V → Prop} (hP : 𝚺₁-Relation₃ P) - (hRel : ∀ param n k R v, L.Rel k R → L.SemitermVec k n v → P param (^rel n k R v) (c.rel param n k R v)) - (hNRel : ∀ param n k R v, L.Rel k R → L.SemitermVec k n v → P param (^nrel n k R v) (c.nrel param n k R v)) - (hverum : ∀ param n, P param (^⊤[n]) (c.verum param n)) - (hfalsum : ∀ param n, P param (^⊥[n]) (c.falsum param n)) - (hand : ∀ param n p q, L.Semiformula n p → L.Semiformula n q → - P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋏[n] q) (c.and param n p q (c.result param p) (c.result param q))) - (hor : ∀ param n p q, L.Semiformula n p → L.Semiformula n q → - P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋎[n] q) (c.or param n p q (c.result param p) (c.result param q))) - (hall : ∀ param n p, L.Semiformula (n + 1) p → - P (c.allChanges param n) p (c.result (c.allChanges param n) p) → - P param (^∀[n] p) (c.all param n p (c.result (c.allChanges param n) p))) - (hex : ∀ param n p, L.Semiformula (n + 1) p → - P (c.exChanges param n) p (c.result (c.exChanges param n) p) → - P param (^∃[n] p) (c.ex param n p (c.result (c.exChanges param n) p))) : - ∀ {param p : V}, L.UFormula p → P param p (c.result param p) := by + (hRel : ∀ param k R v, L.Rel k R → L.IsUTermVec k v → P param (^rel k R v) (c.rel param k R v)) + (hNRel : ∀ param k R v, L.Rel k R → L.IsUTermVec k v → P param (^nrel k R v) (c.nrel param k R v)) + (hverum : ∀ param, P param ^⊤ (c.verum param)) + (hfalsum : ∀ param, P param ^⊥ (c.falsum param)) + (hand : ∀ param p q, L.IsUFormula p → L.IsUFormula q → + P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋏ q) (c.and param p q (c.result param p) (c.result param q))) + (hor : ∀ param p q, L.IsUFormula p → L.IsUFormula q → + P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋎ q) (c.or param p q (c.result param p) (c.result param q))) + (hall : ∀ param p, L.IsUFormula p → + P (c.allChanges param) p (c.result (c.allChanges param) p) → + P param (^∀ p) (c.all param p (c.result (c.allChanges param) p))) + (hex : ∀ param p, L.IsUFormula p → + P (c.exChanges param) p (c.result (c.exChanges param) p) → + P param (^∃ p) (c.ex param p (c.result (c.exChanges param) p))) : + ∀ {param p : V}, L.IsUFormula p → P param p (c.result param p) := by haveI : 𝚺₁-Function₂ c.result := c.result_definable + haveI : 𝚺₁-Function₁ c.allChanges := c.allChanges_defined.to_definable + haveI : 𝚺₁-Function₁ c.exChanges := c.exChanges_defined.to_definable + let f : V → V → V := fun _ param ↦ Max.max param (Max.max (c.allChanges param) (c.exChanges param)) + have hf : 𝚺₁-Function₂ f := by simp [f]; definability intro param p - haveI : 𝚺₁-Function₂ c.allChanges := c.allChanges_defined.to_definable - haveI : 𝚺₁-Function₂ c.exChanges := c.exChanges_defined.to_definable - let f : V → V → V := fun p param ↦ max param (max (c.allChanges param (fstIdx p)) (c.exChanges param (fstIdx p))) - have hf : 𝚺₁-Function₂ f := - HierarchySymbol.BoldfaceFunction.comp₂ (f := Max.max) - (HierarchySymbol.BoldfaceFunction.var _) - (HierarchySymbol.BoldfaceFunction.comp₂ - (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _))) - (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)))) - apply sigma₁_order_ball_induction hf ?_ ?_ p param + apply order_ball_induction_sigma1 hf ?_ ?_ p param · apply HierarchySymbol.Boldface.imp (HierarchySymbol.Boldface.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) (HierarchySymbol.Boldface.comp₃ @@ -1233,43 +1193,41 @@ lemma uformula_result_induction {P : V → V → V → Prop} (hP : 𝚺₁-Relat (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _))) intro p param ih hp rcases hp.case with - (⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, k, r, v, hkr, hv, rfl⟩ | - ⟨n, rfl⟩ | ⟨n, rfl⟩ | - ⟨n, p₁, p₂, hp₁, hp₂, rfl⟩ | ⟨n, p₁, p₂, hp₁, hp₂, rfl⟩ | - ⟨n, p₁, hp₁, rfl⟩ | ⟨n, p₁, hp₁, rfl⟩) - · simpa [hkr, hv] using hRel param n k r v hkr hv - · simpa [hkr, hv] using hNRel param n k r v hkr hv - · simpa using hverum param n - · simpa using hfalsum param n + (⟨k, r, v, hkr, hv, rfl⟩ | ⟨k, r, v, hkr, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, hp₁, hp₂, rfl⟩ | ⟨p₁, p₂, hp₁, hp₂, rfl⟩ | ⟨p₁, hp₁, rfl⟩ | ⟨p₁, hp₁, rfl⟩) + · simpa [hkr, hv] using hRel param k r v hkr hv + · simpa [hkr, hv] using hNRel param k r v hkr hv + · simpa using hverum param + · simpa using hfalsum param · simpa [c.result_and hp₁ hp₂] using - hand param n p₁ p₂ hp₁ hp₂ (ih p₁ (by simp) param (by simp [f]) hp₁.1) (ih p₂ (by simp) param (by simp [f]) hp₂.1) + hand param p₁ p₂ hp₁ hp₂ (ih p₁ (by simp) param (by simp [f]) hp₁) (ih p₂ (by simp) param (by simp [f]) hp₂) · simpa [c.result_or hp₁ hp₂] using - hor param n p₁ p₂ hp₁ hp₂ (ih p₁ (by simp) param (by simp [f]) hp₁.1) (ih p₂ (by simp) param (by simp [f]) hp₂.1) + hor param p₁ p₂ hp₁ hp₂ (ih p₁ (by simp) param (by simp [f]) hp₁) (ih p₂ (by simp) param (by simp [f]) hp₂) · simpa [c.result_all hp₁] using - hall param n p₁ hp₁ (ih p₁ (by simp) (c.allChanges param n) (by simp [f]) hp₁.1) + hall param p₁ hp₁ (ih p₁ (by simp) (c.allChanges param) (by simp [f]) hp₁) · simpa [c.result_ex hp₁] using - hex param n p₁ hp₁ (ih p₁ (by simp) (c.exChanges param n) (by simp [f]) hp₁.1) + hex param p₁ hp₁ (ih p₁ (by simp) (c.exChanges param) (by simp [f]) hp₁) +/- lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺₁-Relation₄ P) - (hRel : ∀ param n k R v, L.Rel k R → L.SemitermVec k n v → P param n (^rel n k R v) (c.rel param n k R v)) - (hNRel : ∀ param n k R v, L.Rel k R → L.SemitermVec k n v → P param n (^nrel n k R v) (c.nrel param n k R v)) - (hverum : ∀ param n, P param n (^⊤[n]) (c.verum param n)) - (hfalsum : ∀ param n, P param n (^⊥[n]) (c.falsum param n)) - (hand : ∀ param n p q, L.Semiformula n p → L.Semiformula n q → - P param n p (c.result param p) → P param n q (c.result param q) → P param n (p ^⋏[n] q) (c.and param n p q (c.result param p) (c.result param q))) - (hor : ∀ param n p q, L.Semiformula n p → L.Semiformula n q → - P param n p (c.result param p) → P param n q (c.result param q) → P param n (p ^⋎[n] q) (c.or param n p q (c.result param p) (c.result param q))) - (hall : ∀ param n p, L.Semiformula (n + 1) p → - P (c.allChanges param n) (n + 1) p (c.result (c.allChanges param n) p) → - P param n (^∀[n] p) (c.all param n p (c.result (c.allChanges param n) p))) - (hex : ∀ param n p, L.Semiformula (n + 1) p → - P (c.exChanges param n) (n + 1) p (c.result (c.exChanges param n) p) → - P param n (^∃[n] p) (c.ex param n p (c.result (c.exChanges param n) p))) : - ∀ {param n p : V}, L.Semiformula n p → P param n p (c.result param p) := by - suffices ∀ {param p : V}, L.UFormula p → ∀ n ≤ p, n = fstIdx p → P param n p (c.result param p) - by intro param n p hp; exact @this param p hp.1 n (by simp [hp.2]) hp.2 + (hRel : ∀ param k R v, L.Rel k R → L.SemitermVec k n v → P param (^rel n k R v) (c.rel param k R v)) + (hNRel : ∀ param k R v, L.Rel k R → L.SemitermVec k n v → P param (^nrel n k R v) (c.nrel param k R v)) + (hverum : ∀ param, P param (^⊤[n]) (c.verum param)) + (hfalsum : ∀ param, P param (^⊥[n]) (c.falsum param)) + (hand : ∀ param p q, L.Semiformula n p → L.Semiformula n q → + P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋏ q) (c.and param p q (c.result param p) (c.result param q))) + (hor : ∀ param p q, L.Semiformula n p → L.Semiformula n q → + P param p (c.result param p) → P param q (c.result param q) → P param (p ^⋎ q) (c.or param p q (c.result param p) (c.result param q))) + (hall : ∀ param p, L.Semiformula (n + 1) p → + P (c.allChanges param) (n + 1) p (c.result (c.allChanges param) p) → + P param (^∀ p) (c.all param p (c.result (c.allChanges param) p))) + (hex : ∀ param p, L.Semiformula (n + 1) p → + P (c.exChanges param) (n + 1) p (c.result (c.exChanges param) p) → + P param (^∃ p) (c.ex param p (c.result (c.exChanges param) p))) : + ∀ {param p : V}, L.Semiformula n p → P param p (c.result param p) := by + suffices ∀ {param p : V}, L.IsUFormula p → ∀ n ≤ p, n = fstIdx p → P param p (c.result param p) + by intro param p hp; exact @this param p hp.1 n (by simp [hp.2]) hp.2 intro param p hp - apply c.uformula_result_induction (P := fun param p y ↦ ∀ n ≤ p, n = fstIdx p → P param n p y) + apply c.uformula_result_induction (P := fun param p y ↦ ∀ n ≤ p, n = fstIdx p → P param p y) ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp · apply HierarchySymbol.Boldface.ball_le (HierarchySymbol.BoldfaceFunction.var _) simp_all only [zero_add, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, Fin.succ_one_eq_two, @@ -1288,90 +1246,426 @@ lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺 simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] - · rintro param n k R v hkR hv _ _ rfl; simpa using hRel param n k R v hkR hv - · rintro param n k R v hkR hv _ _ rfl; simpa using hNRel param n k R v hkR hv - · rintro param n _ _ rfl; simpa using hverum param n - · rintro param n _ _ rfl; simpa using hfalsum param n - · rintro param n p q hp hq ihp ihq _ _ rfl - have ihp : P param n p (c.result param p) := ihp n (by simp [hp.2]) hp.2 - have ihq : P param n q (c.result param q) := ihq n (by simp [hq.2]) hq.2 - simpa using hand param n p q hp hq ihp ihq - · rintro param n p q hp hq ihp ihq _ _ rfl - have ihp : P param n p (c.result param p) := ihp n (by simp [hp.2]) hp.2 - have ihq : P param n q (c.result param q) := ihq n (by simp [hq.2]) hq.2 - simpa using hor param n p q hp hq ihp ihq - · rintro param n p hp ihp _ _ rfl - have ihp : P (c.allChanges param n) (n + 1) p (c.result (c.allChanges param n) p) := ihp (n + 1) (by simp [hp.2]) hp.2 - simpa using hall param n p hp ihp - · rintro param n p hp ihp _ _ rfl - have ihp : P (c.exChanges param n) (n + 1) p (c.result (c.exChanges param n) p) := ihp (n + 1) (by simp [hp.2]) hp.2 - simpa using hex param n p hp ihp + · rintro param k R v hkR hv _ _ rfl; simpa using hRel param k R v hkR hv + · rintro param k R v hkR hv _ _ rfl; simpa using hNRel param k R v hkR hv + · rintro param _ _ rfl; simpa using hverum param + · rintro param _ _ rfl; simpa using hfalsum param + · rintro param p q hp hq ihp ihq _ _ rfl + have ihp : P param p (c.result param p) := ihp n (by simp [hp.2]) hp.2 + have ihq : P param q (c.result param q) := ihq n (by simp [hq.2]) hq.2 + simpa using hand param p q hp hq ihp ihq + · rintro param p q hp hq ihp ihq _ _ rfl + have ihp : P param p (c.result param p) := ihp n (by simp [hp.2]) hp.2 + have ihq : P param q (c.result param q) := ihq n (by simp [hq.2]) hq.2 + simpa using hor param p q hp hq ihp ihq + · rintro param p hp ihp _ _ rfl + have ihp : P (c.allChanges param) (n + 1) p (c.result (c.allChanges param) p) := ihp (n + 1) (by simp [hp.2]) hp.2 + simpa using hall param p hp ihp + · rintro param p hp ihp _ _ rfl + have ihp : P (c.exChanges param) (n + 1) p (c.result (c.exChanges param) p) := ihp (n + 1) (by simp [hp.2]) hp.2 + simpa using hex param p hp ihp +-/ end Construction end Language.UformulaRec1 -/- -namespace Language.UformulaRec - -structure Blueprint (pL : LDef) (arity : ℕ) where - rel : 𝚺₁.Semisentence (arity + 5) - nrel : 𝚺₁.Semisentence (arity + 5) - verum : 𝚺₁.Semisentence (arity + 2) - falsum : 𝚺₁.Semisentence (arity + 2) - and : 𝚺₁.Semisentence (arity + 6) - or : 𝚺₁.Semisentence (arity + 6) - all : 𝚺₁.Semisentence (arity + 4) - ex : 𝚺₁.Semisentence (arity + 4) - allChanges : Fin arity → 𝚺₁.Semisentence (arity + 2) - exChanges : Fin arity → 𝚺₁.Semisentence (arity + 2) - -structure Construction (L : Arith.Language V) {arity} (φ : Blueprint pL arity) where - rel (param : Fin arity → V) (n k R v : V) : V - nrel (param : Fin arity → V) (n k R v : V) : V - verum (param : Fin arity → V) (n : V) : V - falsum (param : Fin arity → V) (n : V) : V - and (param : Fin arity → V) (n p₁ p₂ y₁ y₂ : V) : V - or (param : Fin arity → V) (n p₁ p₂ y₁ y₂ : V) : V - all (param : Fin arity → V) (n p₁ y₁ : V) : V - ex (param : Fin arity → V) (n p₁ y₁ : V) : V - allChanges (i : Fin arity) (param : Fin arity → V) (n : V) : V - exChanges (i : Fin arity) (param : Fin arity → V) (n : V) : V - rel_defined : DefinedFunction (fun v ↦ rel (v ·.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3)) φ.rel - nrel_defined : DefinedFunction (fun v ↦ nrel (v ·.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3)) φ.nrel - verum_defined : DefinedFunction (fun v ↦ verum (v ·.succ) (v 0)) φ.verum - falsum_defined : DefinedFunction (fun v ↦ falsum (v ·.succ) (v 0)) φ.falsum - and_defined : DefinedFunction (fun v ↦ and (v ·.succ.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3) (v 4)) φ.and - or_defined : DefinedFunction (fun v ↦ or (v ·.succ.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3) (v 4)) φ.or - all_defined : DefinedFunction (fun v ↦ all (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) φ.all - ex_defined : DefinedFunction (fun v ↦ ex (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) φ.ex - allChanges_defined (i : Fin arity) : DefinedFunction (fun v ↦ allChanges i (v ·.succ) (v 0)) (φ.allChanges i) - exChanges_defined (i : Fin arity) : DefinedFunction (fun v ↦ exChanges i (v ·.succ) (v 0)) (φ.exChanges i) - -variable {arity} (β : Blueprint pL arity) +section bv -namespace Blueprint +namespace BV -def decomp {n : ℕ} (s : 𝚺₁.Semisentence n) : 𝚺₁.Semisentence 1 := - .mkSigma (∃^[n] (Matrix.conj fun i : Fin n ↦ - (unNpairDef i).val/[#(i.natAdd 1), #⟨n, by simp⟩]) ⋏ (Rew.substs fun i : Fin n ↦ #(i.natAdd 1)).hom s) (by simp) - -def toRec1 : UformulaRec1.Blueprint pL where - rel := .mkSigma “y param n k R v | !qqNRelDef y n k R v” (by simp) - nrel := .mkSigma “y param n k R v | !qqRelDef y n k R v” (by simp) - verum := .mkSigma “y param n | !qqFalsumDef y n” (by simp) - falsum := .mkSigma “y param n | !qqVerumDef y n” (by simp) - and := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqOrDef y n y₁ y₂” (by simp) - or := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqAndDef y n y₁ y₂” (by simp) - all := .mkSigma “y param n p₁ y₁ | !qqExDef y n y₁” (by simp) - ex := .mkSigma “y param n p₁ y₁ | !qqAllDef y n y₁” (by simp) - allChanges := .mkSigma “param' param n | param' = 0” (by simp) - exChanges := .mkSigma “param' param n | param' = 0” (by simp) +def blueprint (pL : LDef) : Language.UformulaRec1.Blueprint pL where + rel := .mkSigma “y param k R v | ∃ M, !pL.termBVVecDef M k v ∧ !listMaxDef y M” (by simp) + nrel := .mkSigma “y param k R v | ∃ M, !pL.termBVVecDef M k v ∧ !listMaxDef y M” (by simp) + verum := .mkSigma “y param | y = 0” (by simp) + falsum := .mkSigma “y param | y = 0” (by simp) + and := .mkSigma “y param p₁ p₂ y₁ y₂ | !max y y₁ y₂” (by simp) + or := .mkSigma “y param p₁ p₂ y₁ y₂ | !max y y₁ y₂” (by simp) + all := .mkSigma “y param p₁ y₁ | !subDef y y₁ 1” (by simp) + ex := .mkSigma “y param p₁ y₁ | !subDef y y₁ 1” (by simp) + allChanges := .mkSigma “param' param | param' = 0” (by simp) + exChanges := .mkSigma “param' param | param' = 0” (by simp) -end Blueprint +variable (L) -end Language.UformulaRec --/ +def construction : Language.UformulaRec1.Construction V L (blueprint pL) where + rel {_} := fun k _ v ↦ listMax (L.termBVVec k v) + nrel {_} := fun k _ v ↦ listMax (L.termBVVec k v) + verum {_} := 0 + falsum {_} := 0 + and {_} := fun _ _ y₁ y₂ ↦ Max.max y₁ y₂ + or {_} := fun _ _ y₁ y₂ ↦ Max.max y₁ y₂ + all {_} := fun _ y₁ ↦ y₁ - 1 + ex {_} := fun _ y₁ ↦ y₁ - 1 + allChanges := fun _ ↦ 0 + exChanges := fun _ ↦ 0 + rel_defined := by intro v; simp [blueprint, L.termBVVec_defined.df.iff]; rfl + nrel_defined := by intro v; simp [blueprint, L.termBVVec_defined.df.iff]; rfl + verum_defined := by intro v; simp [blueprint] + falsum_defined := by intro v; simp [blueprint] + and_defined := by intro v; simp [blueprint]; rfl + or_defined := by intro v; simp [blueprint]; rfl + all_defined := by intro v; simp [blueprint]; rfl + ex_defined := by intro v; simp [blueprint]; rfl + allChanges_defined := by intro v; simp [blueprint] + exChanges_defined := by intro v; simp [blueprint] + +end BV + +open BV + +variable (L) + +def Language.bv (p : V) : V := (construction L).result 0 p + +variable {L} + +section + +def _root_.LO.FirstOrder.Arith.LDef.bvDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result.rew (Rew.substs ![#0, ‘0’, #1]) + +variable (L) + +lemma Language.bv_defined : 𝚺₁-Function₁ L.bv via pL.bvDef := fun v ↦ by + simpa [LDef.bvDef] using (construction L).result_defined ![v 0, 0, v 1] + +instance Language.bv_definable : 𝚺₁-Function₁ L.bv := L.bv_defined.to_definable + +instance Language.bv_definable' (Γ) : Γ-[m + 1]-Function₁ L.bv := L.bv_definable.of_sigmaOne + +end + +@[simp] lemma bv_rel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.bv (^rel k R v) = listMax (L.termBVVec k v) := by simp [Language.bv, hR, hv, construction] + +@[simp] lemma bv_nrel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.bv (^nrel k R v) = listMax (L.termBVVec k v) := by simp [Language.bv, hR, hv, construction] + +@[simp] lemma bv_verum : L.bv ^⊤ = 0 := by simp [Language.bv, construction] + +@[simp] lemma bv_falsum : L.bv ^⊥ = 0 := by simp [Language.bv, construction] + +@[simp] lemma bv_and {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.bv (p ^⋏ q) = Max.max (L.bv p) (L.bv q) := by simp [Language.bv, hp, hq, construction] + +@[simp] lemma bv_or {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.bv (p ^⋎ q) = Max.max (L.bv p) (L.bv q) := by simp [Language.bv, hp, hq, construction] + +@[simp] lemma bv_all {p} (hp : L.IsUFormula p) : L.bv (^∀ p) = L.bv p - 1 := by simp [Language.bv, hp, construction] + +@[simp] lemma bv_ex {p} (hp : L.IsUFormula p) : L.bv (^∃ p) = L.bv p - 1 := by simp [Language.bv, hp, construction] + +lemma bv_eq_of_not_isUFormula {p} (h : ¬L.IsUFormula p) : L.bv p = 0 := (construction L).result_prop_not _ h + +end bv + +section isSemiformula + +variable (L) + +structure Language.IsSemiformula (n p : V) : Prop where + isUFormula : L.IsUFormula p + bv : L.bv p ≤ n + +abbrev Language.IsFormula (p : V) : Prop := L.IsSemiformula 0 p + +variable {L} + +lemma Language.isSemiformula_iff {n p : V} : + L.IsSemiformula n p ↔ L.IsUFormula p ∧ L.bv p ≤ n := + ⟨fun h ↦ ⟨h.isUFormula, h.bv⟩, by rintro ⟨hp, h⟩; exact ⟨hp, h⟩⟩ + +section + +variable (L) + +def _root_.LO.FirstOrder.Arith.LDef.isSemiformulaDef (pL : LDef) : 𝚫₁.Semisentence 2 := .mkDelta + (.mkSigma “n p | !pL.isUFormulaDef.sigma p ∧ ∃ b, !pL.bvDef b p ∧ b ≤ n” (by simp)) + (.mkPi “n p | !pL.isUFormulaDef.pi p ∧ ∀ b, !pL.bvDef b p → b ≤ n” (by simp)) + +lemma Language.isSemiformula_defined : 𝚫₁-Relation L.IsSemiformula via pL.isSemiformulaDef := by + constructor + · intro v; simp [FirstOrder.Arith.LDef.isSemiformulaDef, HierarchySymbol.Semiformula.val_sigma, + L.bv_defined.df.iff, L.isUFormula_defined.proper.iff'] + · intro v; simp [FirstOrder.Arith.LDef.isSemiformulaDef, HierarchySymbol.Semiformula.val_sigma, + L.bv_defined.df.iff, L.isUFormula_defined.df.iff, Language.isSemiformula_iff] + +instance Language.isSemiformula_definable : 𝚫₁-Relation L.IsSemiformula := L.isSemiformula_defined.to_definable + +instance Language.isSemiformulaDef_definable' : Γ-[m + 1]-Relation L.IsSemiformula := L.isSemiformula_definable.of_deltaOne + +end + +@[simp] lemma Language.IsUFormula.isSemiformula {p : V} (h : L.IsUFormula p) : L.IsSemiformula (L.bv p) p where + isUFormula := h + bv := by rfl + +@[simp] lemma Language.IsSemiformula.rel {n k r v : V} : + L.IsSemiformula n (^rel k r v) ↔ L.Rel k r ∧ L.IsSemitermVec k n v := by + constructor + · intro h + have hrv : L.Rel k r ∧ L.IsUTermVec k v := by simpa using h.isUFormula + exact ⟨hrv.1, hrv.2, fun i hi ↦ by + have : listMax (L.termBVVec k v) ≤ n := by simpa [hrv] using h.bv + exact le_trans (le_trans (by simp_all) (nth_le_listMax (i := i) (by simp_all))) this⟩ + · rintro ⟨hr, hv⟩ + exact ⟨by simp [hr, hv.isUTerm], by + rw [bv_rel hr hv.isUTerm] + apply listMaxss_le + intro i hi + have := hv.bv (i := i) (by simpa [hv.isUTerm] using hi) + rwa [nth_termBVVec hv.isUTerm (by simpa [hv.isUTerm] using hi)]⟩ + +@[simp] lemma Language.IsSemiformula.nrel {n k r v : V} : + L.IsSemiformula n (^nrel k r v) ↔ L.Rel k r ∧ L.IsSemitermVec k n v := by + constructor + · intro h + have hrv : L.Rel k r ∧ L.IsUTermVec k v := by simpa using h.isUFormula + exact ⟨hrv.1, hrv.2, fun i hi ↦ by + have : listMax (L.termBVVec k v) ≤ n := by simpa [hrv] using h.bv + exact le_trans (le_trans (by simp_all) (nth_le_listMax (i := i) (by simp_all))) this⟩ + · rintro ⟨hr, hv⟩ + exact ⟨by simp [hr, hv.isUTerm], by + rw [bv_nrel hr hv.isUTerm] + apply listMaxss_le + intro i hi + have := hv.bv (i := i) (by simpa [hv.isUTerm] using hi) + rwa [nth_termBVVec hv.isUTerm (by simpa [hv.isUTerm] using hi)]⟩ + +@[simp] lemma Language.IsSemiformula.verum {n} : L.IsSemiformula n ^⊤ := ⟨by simp, by simp⟩ + +@[simp] lemma Language.IsSemiformula.falsum {n} : L.IsSemiformula n ^⊥ := ⟨by simp, by simp⟩ + +@[simp] lemma Language.IsSemiformula.and {n p q : V} : + L.IsSemiformula n (p ^⋏ q) ↔ L.IsSemiformula n p ∧ L.IsSemiformula n q := by + constructor + · intro h + have hpq : L.IsUFormula p ∧ L.IsUFormula q := by simpa using h.isUFormula + have hbv : L.bv p ≤ n ∧ L.bv q ≤ n := by simpa [hpq] using h.bv + exact ⟨⟨hpq.1, hbv.1⟩, ⟨hpq.2, hbv.2⟩⟩ + · rintro ⟨hp, hq⟩ + exact ⟨by simp [hp.isUFormula, hq.isUFormula], by simp [hp.isUFormula, hq.isUFormula, hp.bv, hq.bv]⟩ + +@[simp] lemma Language.IsSemiformula.or {n p q : V} : + L.IsSemiformula n (p ^⋎ q) ↔ L.IsSemiformula n p ∧ L.IsSemiformula n q := by + constructor + · intro h + have hpq : L.IsUFormula p ∧ L.IsUFormula q := by simpa using h.isUFormula + have hbv : L.bv p ≤ n ∧ L.bv q ≤ n := by simpa [hpq] using h.bv + exact ⟨⟨hpq.1, hbv.1⟩, ⟨hpq.2, hbv.2⟩⟩ + · rintro ⟨hp, hq⟩ + exact ⟨by simp [hp.isUFormula, hq.isUFormula], by simp [hp.isUFormula, hq.isUFormula, hp.bv, hq.bv]⟩ + +@[simp] lemma Language.IsSemiformula.all {n p : V} : + L.IsSemiformula n (^∀ p) ↔ L.IsSemiformula (n + 1) p := by + constructor + · intro h + exact ⟨by simpa using h.isUFormula, by + simpa [show L.IsUFormula p by simpa using h.isUFormula] using h.bv⟩ + · intro h + exact ⟨by simp [h.isUFormula], by simp [h.isUFormula, h.bv]⟩ + +@[simp] lemma Language.IsSemiformula.ex {n p : V} : + L.IsSemiformula n (^∃ p) ↔ L.IsSemiformula (n + 1) p := by + constructor + · intro h + exact ⟨by simpa using h.isUFormula, by + simpa [show L.IsUFormula p by simpa using h.isUFormula] using h.bv⟩ + · intro h + exact ⟨by simp [h.isUFormula], by simp [h.isUFormula, h.bv]⟩ + +lemma Language.IsSemiformula.case_iff {n p : V} : + L.IsSemiformula n p ↔ + (∃ k R v, L.Rel k R ∧ L.IsSemitermVec k n v ∧ p = ^rel k R v) ∨ + (∃ k R v, L.Rel k R ∧ L.IsSemitermVec k n v ∧ p = ^nrel k R v) ∨ + (p = ^⊤) ∨ + (p = ^⊥) ∨ + (∃ p₁ p₂, L.IsSemiformula n p₁ ∧ L.IsSemiformula n p₂ ∧ p = p₁ ^⋏ p₂) ∨ + (∃ p₁ p₂, L.IsSemiformula n p₁ ∧ L.IsSemiformula n p₂ ∧ p = p₁ ^⋎ p₂) ∨ + (∃ p₁, L.IsSemiformula (n + 1) p₁ ∧ p = ^∀ p₁) ∨ + (∃ p₁, L.IsSemiformula (n + 1) p₁ ∧ p = ^∃ p₁) := by + constructor + · intro h + rcases h.isUFormula.case with + (⟨k, r, v, _, _, rfl⟩ | ⟨k, r, v, _, _, rfl⟩ | rfl | rfl | ⟨p₁, p₂, _, _, rfl⟩ | ⟨p₁, p₂, _, _, rfl⟩ | ⟨p₁, _, rfl⟩ | ⟨p₁, _, rfl⟩) + · have : L.Rel k r ∧ L.IsSemitermVec k n v := by simpa using h + exact Or.inl ⟨k, r, v, by simp [this]⟩ + · have : L.Rel k r ∧ L.IsSemitermVec k n v := by simpa using h + exact Or.inr <| Or.inl ⟨k, r, v, by simp [this]⟩ + · exact Or.inr <| Or.inr <| Or.inl rfl + · exact Or.inr <| Or.inr <| Or.inr <| Or.inl rfl + · have : L.IsSemiformula n p₁ ∧ L.IsSemiformula n p₂ := by simpa using h + exact Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, p₂, by simp [this]⟩ + · have : L.IsSemiformula n p₁ ∧ L.IsSemiformula n p₂ := by simpa using h + exact Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, p₂, by simp [this]⟩ + · have : L.IsSemiformula (n + 1) p₁ := by simpa using h + exact Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨p₁, by simp [this]⟩ + · have : L.IsSemiformula (n + 1) p₁ := by simpa using h + exact Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr ⟨p₁, by simp [this]⟩ + · rintro (⟨k, R, v, hR, hv, rfl⟩ | ⟨k, R, v, hR, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, h₁, rfl⟩ | ⟨p₁, h₁, rfl⟩) <;> simp [*] + +lemma Language.IsSemiformula.case {P : V → V → Prop} {n p} (hp : L.IsSemiformula n p) + (hrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^rel k r v)) + (hnrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^nrel k r v)) + (hverum : ∀ n, P n ^⊤) + (hfalsum : ∀ n, P n ^⊥) + (hand : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.IsSemiformula (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.IsSemiformula (n + 1) p → P n (^∃ p)) : P n p := by + rcases Language.IsSemiformula.case_iff.mp hp with + (⟨k, R, v, hR, hv, rfl⟩ | ⟨k, R, v, hR, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, h₁, rfl⟩ | ⟨p₁, h₁, rfl⟩) + · exact hrel _ _ _ _ hR hv + · exact hnrel _ _ _ _ hR hv + · exact hverum n + · exact hfalsum n + · exact hand _ _ _ h₁ h₂ + · exact hor _ _ _ h₁ h₂ + · exact hall _ _ h₁ + · exact hex _ _ h₁ + +lemma Language.IsSemiformula.induction_sigma1 {P : V → V → Prop} (hP : 𝚺₁-Relation P) + (hrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^rel k r v)) + (hnrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^nrel k r v)) + (hverum : ∀ n, P n ^⊤) + (hfalsum : ∀ n, P n ^⊥) + (hand : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) {n p} : + L.IsSemiformula n p → P n p := by + have : 𝚺₁-Function₂ (fun _ (n : V) ↦ n + 1) := by definability + apply order_ball_induction_sigma1 this ?_ ?_ p n + · apply HierarchySymbol.Boldface.imp + · apply HierarchySymbol.Boldface.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.Boldface.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _) + intro p n ih hp + rcases Language.IsSemiformula.case_iff.mp hp with + (⟨k, R, v, hR, hv, rfl⟩ | ⟨k, R, v, hR, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, h₁, rfl⟩ | ⟨p₁, h₁, rfl⟩) + · apply hrel _ _ _ _ hR hv + · apply hnrel _ _ _ _ hR hv + · apply hverum + · apply hfalsum + · apply hand _ _ _ h₁ h₂ (ih p₁ (by simp) n (by simp) h₁) (ih p₂ (by simp) n (by simp) h₂) + · apply hor _ _ _ h₁ h₂ (ih p₁ (by simp) n (by simp) h₁) (ih p₂ (by simp) n (by simp) h₂) + · apply hall _ _ h₁ (ih p₁ (by simp) (n + 1) (by simp) h₁) + · apply hex _ _ h₁ (ih p₁ (by simp) (n + 1) (by simp) h₁) + +lemma Language.IsSemiformula.induction_pi1 {P : V → V → Prop} (hP : 𝚷₁-Relation P) + (hrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^rel k r v)) + (hnrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^nrel k r v)) + (hverum : ∀ n, P n ^⊤) + (hfalsum : ∀ n, P n ^⊥) + (hand : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) {n p} : + L.IsSemiformula n p → P n p := by + suffices L.IsUFormula p → ∀ n, L.IsSemiformula n p → P n p by intro h; exact this h.isUFormula n h + apply Language.IsUFormula.induction_pi1 (P := fun p ↦ ∀ n, L.IsSemiformula n p → P n p) + · definability + · intro k R v hR _ n h + have : L.Rel k R ∧ L.IsSemitermVec k n v := by simpa using h + exact hrel _ _ _ _ hR this.2 + · intro k R v hR _ n h + have : L.Rel k R ∧ L.IsSemitermVec k n v := by simpa using h + exact hnrel _ _ _ _ hR this.2 + · intro n _; apply hverum + · intro n _; apply hfalsum + · intro p q _ _ ihp ihq n h + have : L.IsSemiformula n p ∧ L.IsSemiformula n q := by simpa using h + apply hand _ _ _ this.1 this.2 (ihp n this.1) (ihq n this.2) + · intro p q _ _ ihp ihq n h + have : L.IsSemiformula n p ∧ L.IsSemiformula n q := by simpa using h + apply hor _ _ _ this.1 this.2 (ihp n this.1) (ihq n this.2) + · intro p _ ihp n h + have : L.IsSemiformula (n + 1) p := by simpa using h + apply hall _ _ this (ihp _ this) + · intro p _ ihp n h + have : L.IsSemiformula (n + 1) p := by simpa using h + apply hex _ _ this (ihp _ this) + +lemma Language.IsSemiformula.induction1 (Γ) {P : V → V → Prop} (hP : Γ-[1]-Relation P) + (hrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^rel k r v)) + (hnrel : ∀ n k r v, L.Rel k r → L.IsSemitermVec k n v → P n (^nrel k r v)) + (hverum : ∀ n, P n ^⊤) + (hfalsum : ∀ n, P n ^⊥) + (hand : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋏ q)) + (hor : ∀ n p q, L.IsSemiformula n p → L.IsSemiformula n q → P n p → P n q → P n (p ^⋎ q)) + (hall : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∀ p)) + (hex : ∀ n p, L.IsSemiformula (n + 1) p → P (n + 1) p → P n (^∃ p)) {n p} : + L.IsSemiformula n p → P n p := + match Γ with + | 𝚺 => Language.IsSemiformula.induction_sigma1 hP hrel hnrel hverum hfalsum hand hor hall hex + | 𝚷 => Language.IsSemiformula.induction_pi1 hP hrel hnrel hverum hfalsum hand hor hall hex + | 𝚫 => Language.IsSemiformula.induction_sigma1 hP.of_delta hrel hnrel hverum hfalsum hand hor hall hex + + +lemma Language.IsSemiformula.pos {n p : V} (h : L.IsSemiformula n p) : 0 < p := h.isUFormula.pos + +@[simp] lemma Language.IsSemiformula.not_zero (m : V) : ¬L.IsSemiformula m (0 : V) := by intro h; simpa using h.pos + +end isSemiformula + +namespace Language.UformulaRec1.Construction + +variable {β : Blueprint pL} {c : Construction V L β} {param : V} + +lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺₁-Relation₄ P) + (hRel : ∀ n param k R v, L.Rel k R → L.IsSemitermVec k n v → P param n (^rel k R v) (c.rel param k R v)) + (hNRel : ∀ n param k R v, L.Rel k R → L.IsSemitermVec k n v → P param n (^nrel k R v) (c.nrel param k R v)) + (hverum : ∀ n param, P param n ^⊤ (c.verum param)) + (hfalsum : ∀ n param, P param n ^⊥ (c.falsum param)) + (hand : ∀ n param p q, L.IsSemiformula n p → L.IsSemiformula n q → + P param n p (c.result param p) → P param n q (c.result param q) → P param n (p ^⋏ q) (c.and param p q (c.result param p) (c.result param q))) + (hor : ∀ n param p q, L.IsSemiformula n p → L.IsSemiformula n q → + P param n p (c.result param p) → P param n q (c.result param q) → P param n (p ^⋎ q) (c.or param p q (c.result param p) (c.result param q))) + (hall : ∀ n param p, L.IsSemiformula (n + 1) p → + P (c.allChanges param) (n + 1) p (c.result (c.allChanges param) p) → + P param n (^∀ p) (c.all param p (c.result (c.allChanges param) p))) + (hex : ∀ n param p, L.IsSemiformula (n + 1) p → + P (c.exChanges param) (n + 1) p (c.result (c.exChanges param) p) → + P param n (^∃ p) (c.ex param p (c.result (c.exChanges param) p))) : + ∀ {param n p : V}, L.IsSemiformula n p → P param n p (c.result param p) := by + haveI : 𝚺₁-Function₂ c.result := c.result_definable + haveI : 𝚺₁-Function₁ c.allChanges := c.allChanges_defined.to_definable + haveI : 𝚺₁-Function₁ c.exChanges := c.exChanges_defined.to_definable + let f : V → V → V → V := fun _ param _ ↦ Max.max param (Max.max (c.allChanges param) (c.exChanges param)) + have hf : 𝚺₁-Function₃ f := by simp [f]; definability + let g : V → V → V → V := fun _ _ n ↦ n + 1 + have hg : 𝚺₁-Function₃ g := by simp [g]; definability + intro param n p + apply order_ball_induction₂_sigma1 hf hg ?_ ?_ p param n + · apply HierarchySymbol.Boldface.imp + · apply HierarchySymbol.Boldface.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _) + · apply HierarchySymbol.Boldface.comp₄ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _) + intro p param n ih hp + rcases Language.IsSemiformula.case_iff.mp hp with + (⟨k, R, v, hR, hv, rfl⟩ | ⟨k, R, v, hR, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, h₁, rfl⟩ | ⟨p₁, h₁, rfl⟩) + · simpa [hR, hv.isUTerm] using hRel n param k R v hR hv + · simpa [hR, hv.isUTerm] using hNRel n param k R v hR hv + · simpa using hverum n param + · simpa using hfalsum n param + · simpa [h₁.isUFormula, h₂.isUFormula] using + hand n param p₁ p₂ h₁ h₂ + (ih p₁ (by simp) param (by simp [f]) n (by simp [g]) h₁) + (ih p₂ (by simp) param (by simp [f]) n (by simp [g]) h₂) + · simpa [h₁.isUFormula, h₂.isUFormula] using + hor n param p₁ p₂ h₁ h₂ + (ih p₁ (by simp) param (by simp [f]) n (by simp [g]) h₁) + (ih p₂ (by simp) param (by simp [f]) n (by simp [g]) h₂) + · simpa [h₁.isUFormula] using + hall n param p₁ h₁ + (ih p₁ (by simp) (c.allChanges param) (by simp [f]) (n + 1) (by simp [g]) h₁) + · simpa [h₁.isUFormula] using + hex n param p₁ h₁ + (ih p₁ (by simp) (c.exChanges param) (by simp [f]) (n + 1) (by simp [g]) h₁) + +end Language.UformulaRec1.Construction end LO.Arith diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index 8aa0f99..cda58c2 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -16,30 +16,30 @@ section negation namespace Negation def blueprint (pL : LDef) : Language.UformulaRec1.Blueprint pL where - rel := .mkSigma “y param n k R v | !qqNRelDef y n k R v” (by simp) - nrel := .mkSigma “y param n k R v | !qqRelDef y n k R v” (by simp) - verum := .mkSigma “y param n | !qqFalsumDef y n” (by simp) - falsum := .mkSigma “y param n | !qqVerumDef y n” (by simp) - and := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqOrDef y n y₁ y₂” (by simp) - or := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqAndDef y n y₁ y₂” (by simp) - all := .mkSigma “y param n p₁ y₁ | !qqExDef y n y₁” (by simp) - ex := .mkSigma “y param n p₁ y₁ | !qqAllDef y n y₁” (by simp) - allChanges := .mkSigma “param' param n | param' = 0” (by simp) - exChanges := .mkSigma “param' param n | param' = 0” (by simp) + rel := .mkSigma “y param k R v | !qqNRelDef y k R v” (by simp) + nrel := .mkSigma “y param k R v | !qqRelDef y k R v” (by simp) + verum := .mkSigma “y param | !qqFalsumDef y” (by simp) + falsum := .mkSigma “y param | !qqVerumDef y” (by simp) + and := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqOrDef y y₁ y₂” (by simp) + or := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqAndDef y y₁ y₂” (by simp) + all := .mkSigma “y param p₁ y₁ | !qqExDef y y₁” (by simp) + ex := .mkSigma “y param p₁ y₁ | !qqAllDef y y₁” (by simp) + allChanges := .mkSigma “param' param | param' = 0” (by simp) + exChanges := .mkSigma “param' param | param' = 0” (by simp) variable (L) def construction : Language.UformulaRec1.Construction V L (blueprint pL) where - rel {_} := fun n k R v ↦ ^nrel n k R v - nrel {_} := fun n k R v ↦ ^rel n k R v - verum {_} := fun n ↦ ^⊥[n] - falsum {_} := fun n ↦ ^⊤[n] - and {_} := fun n _ _ y₁ y₂ ↦ y₁ ^⋎[n] y₂ - or {_} := fun n _ _ y₁ y₂ ↦ y₁ ^⋏[n] y₂ - all {_} := fun n _ y₁ ↦ ^∃[n] y₁ - ex {_} := fun n _ y₁ ↦ ^∀[n] y₁ - allChanges := fun _ _ ↦ 0 - exChanges := fun _ _ ↦ 0 + rel {_} := fun k R v ↦ ^nrel k R v + nrel {_} := fun k R v ↦ ^rel k R v + verum {_} := ^⊥ + falsum {_} := ^⊤ + and {_} := fun _ _ y₁ y₂ ↦ y₁ ^⋎ y₂ + or {_} := fun _ _ y₁ y₂ ↦ y₁ ^⋏ y₂ + all {_} := fun _ y₁ ↦ ^∃ y₁ + ex {_} := fun _ y₁ ↦ ^∀ y₁ + allChanges := fun _ ↦ 0 + exChanges := fun _ ↦ 0 rel_defined := by intro v; simp [blueprint]; rfl nrel_defined := by intro v; simp [blueprint]; rfl verum_defined := by intro v; simp [blueprint] @@ -67,124 +67,134 @@ def _root_.LO.FirstOrder.Arith.LDef.negDef (pL : LDef) : 𝚺₁.Semisentence 2 variable (L) -lemma neg_defined : 𝚺₁-Function₁ L.neg via pL.negDef := fun v ↦ by +lemma Language.neg_defined : 𝚺₁-Function₁ L.neg via pL.negDef := fun v ↦ by simpa [LDef.negDef] using (construction L).result_defined ![v 0, 0, v 1] -@[simp] lemma neg_defined_iff (v : Fin 2 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.negDef ↔ v 0 = L.neg (v 1) := (neg_defined L).df.iff v +instance Language.neg_definable : 𝚺₁-Function₁ L.neg := L.neg_defined.to_definable -instance neg_definable : 𝚺₁-Function₁ L.neg := - (neg_defined L).to_definable - -instance neg_definable' (Γ) : Γ-[m + 1]-Function₁ L.neg := - .of_sigmaOne (neg_definable L) _ _ +instance Language.neg_definable' (Γ) : Γ-[m + 1]-Function₁ L.neg := + .of_sigmaOne (neg_definable L) end -@[simp] lemma neg_rel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.neg (^rel n k R v) = ^nrel n k R v := by simp [Language.neg, hR, hv, construction] +@[simp] lemma neg_rel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.neg (^rel k R v) = ^nrel k R v := by simp [Language.neg, hR, hv, construction] -@[simp] lemma neg_nrel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.neg (^nrel n k R v) = ^rel n k R v := by simp [Language.neg, hR, hv, construction] +@[simp] lemma neg_nrel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.neg (^nrel k R v) = ^rel k R v := by simp [Language.neg, hR, hv, construction] -@[simp] lemma neg_verum (n) : - L.neg ^⊤[n] = ^⊥[n] := by simp [Language.neg, construction] +@[simp] lemma neg_verum : + L.neg ^⊤ = ^⊥ := by simp [Language.neg, construction] -@[simp] lemma neg_falsum (n) : - L.neg ^⊥[n] = ^⊤[n] := by simp [Language.neg, construction] +@[simp] lemma neg_falsum : + L.neg ^⊥ = ^⊤ := by simp [Language.neg, construction] -@[simp] lemma neg_and {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.neg (p ^⋏[n] q) = L.neg p ^⋎[n] L.neg q := by simp [Language.neg, hp, hq, construction] +@[simp] lemma neg_and {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.neg (p ^⋏ q) = L.neg p ^⋎ L.neg q := by simp [Language.neg, hp, hq, construction] -@[simp] lemma neg_or {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.neg (p ^⋎[n] q) = L.neg p ^⋏[n] L.neg q := by simp [Language.neg, hp, hq, construction] +@[simp] lemma neg_or {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.neg (p ^⋎ q) = L.neg p ^⋏ L.neg q := by simp [Language.neg, hp, hq, construction] -@[simp] lemma neg_all {n p} (hp : L.Semiformula (n + 1) p) : - L.neg (^∀[n] p) = ^∃[n] (L.neg p) := by simp [Language.neg, hp, construction] +@[simp] lemma neg_all {p} (hp : L.IsUFormula p) : + L.neg (^∀ p) = ^∃ (L.neg p) := by simp [Language.neg, hp, construction] -@[simp] lemma neg_ex {n p} (hp : L.Semiformula (n + 1) p) : - L.neg (^∃[n] p) = ^∀[n] (L.neg p) := by simp [Language.neg, hp, construction] +@[simp] lemma neg_ex {p} (hp : L.IsUFormula p) : + L.neg (^∃ p) = ^∀ (L.neg p) := by simp [Language.neg, hp, construction] -lemma neg_not_uformula {x} (h : ¬L.UFormula x) : +lemma neg_not_uformula {x} (h : ¬L.IsUFormula x) : L.neg x = 0 := (construction L).result_prop_not _ h -lemma fstIdx_neg {p} (h : L.UFormula p) : fstIdx (L.neg p) = fstIdx p := by - rcases h.case with (⟨_, _, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, _, rfl⟩ | ⟨_, rfl⟩ | ⟨_, rfl⟩ | - ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, rfl⟩ | ⟨_, _, _, rfl⟩) <;> - simp [*] +lemma Language.IsUFormula.neg {p : V} : L.IsUFormula p → L.IsUFormula (L.neg p) := by + apply Language.IsUFormula.induction_sigma1 + · definability + · intro k r v hr hv; simp [hr, hv] + · intro k r v hr hv; simp [hr, hv] + · simp + · simp + · intro p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] + · intro p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] + · intro p hp ihp; simp [hp, ihp] + · intro p hp ihp; simp [hp, ihp] -lemma Language.Semiformula.neg {p : V} : L.Semiformula n p → L.Semiformula n (L.neg p) := by - apply Language.Semiformula.induction_sigma₁ +@[simp] lemma Language.IsUFormula.bv_neg {p : V} : L.IsUFormula p → L.bv (L.neg p) = L.bv p := by + apply Language.IsUFormula.induction_sigma1 · definability - · intro n k r v hr hv; simp [hr, hv] - · intro n k r v hr hv; simp [hr, hv] + · intro k R v hR hv; simp [*] + · intro k R v hR hv; simp [*] · simp · simp - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p hp ihp; simp [hp, ihp] - · intro n p hp ihp; simp [hp, ihp] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p hp ihp; simp [hp, hp.neg, ihp] + · intro p hp ihp; simp [hp, hp.neg, ihp] -@[simp] lemma Language.Semiformula.neg_iff {p : V} : L.Semiformula n (L.neg p) ↔ L.Semiformula n p := - ⟨fun h ↦ by - rcases h with ⟨h, rfl⟩ - have : L.UFormula p := by - by_contra hp - simp [neg_not_uformula hp] at h - exact ⟨this, by simp [fstIdx_neg this]⟩, - Language.Semiformula.neg⟩ - -@[simp] lemma neg_neg {p : V} : L.Semiformula n p → L.neg (L.neg p) = p := by - apply Language.Semiformula.induction_sigma₁ +@[simp] lemma Language.IsUFormula.neg_neg {p : V} : L.IsUFormula p → L.neg (L.neg p) = p := by + apply Language.IsUFormula.induction_sigma1 · definability - · intro n k r v hr hv; simp [hr, hv] - · intro n k r v hr hv; simp [hr, hv] - · intro n; simp - · intro n; simp - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p hp ihp; simp [hp, ihp] - · intro n p hp ihp; simp [hp, ihp] - -@[simp] lemma neg_inj_iff (hp : L.Semiformula n p) (hq : L.Semiformula n q) : L.neg p = L.neg q ↔ p = q := by + · intro k r v hr hv; simp [hr, hv] + · intro k r v hr hv; simp [hr, hv] + · simp + · simp + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p hp ihp; simp [hp, hp.neg, ihp] + · intro p hp ihp; simp [hp, hp.neg, ihp] + +@[simp] lemma Language.IsUFormula.neg_iff {p : V} : L.IsUFormula (L.neg p) ↔ L.IsUFormula p := by + constructor + · intro h; by_contra hp + have Hp : L.IsUFormula p := by by_contra hp; simp [neg_not_uformula hp] at h + contradiction + · exact Language.IsUFormula.neg + +@[simp] lemma Language.IsSemiformula.neg_iff {p : V} : L.IsSemiformula n (L.neg p) ↔ L.IsSemiformula n p := by constructor - · intro h; simpa [neg_neg hp, neg_neg hq] using congrArg L.neg h + · intro h; by_contra hp + have Hp : L.IsUFormula p := by by_contra hp; simp [neg_not_uformula hp] at h + have : L.IsSemiformula n p := ⟨Hp, by simpa [Hp.bv_neg] using h.bv⟩ + contradiction + · intro h; exact ⟨by simp [h.isUFormula], by simpa [h.isUFormula] using h.bv⟩ + +alias ⟨Language.IsSemiformula.elim_neg, Language.IsSemiformula.neg⟩ := Language.IsSemiformula.neg_iff + +@[simp] lemma neg_inj_iff (hp : L.IsUFormula p) (hq : L.IsUFormula q) : L.neg p = L.neg q ↔ p = q := by + constructor + · intro h; simpa [hp.neg_neg, hq.neg_neg] using congrArg L.neg h · rintro rfl; rfl end negation variable (L) -def Language.imp (n p q : V) : V := L.neg p ^⋎[n] q +def Language.imp (p q : V) : V := L.neg p ^⋎ q -notation:60 p:61 " ^→[" L "; " n "] " q:60 => Language.imp L n p q +notation:60 p:61 " ^→[" L "] " q:60 => Language.imp L p q variable {L} section imp -@[simp] lemma Language.Semiformula.imp {n p q : V} : - L.Semiformula n (L.imp n p q) ↔ L.Semiformula n p ∧ L.Semiformula n q := by +@[simp] lemma Language.IsUFormula.imp {p q : V} : + L.IsUFormula (L.imp p q) ↔ L.IsUFormula p ∧ L.IsUFormula q := by + simp [Language.imp] + +@[simp] lemma Language.IsSemiformula.imp {n p q : V} : + L.IsSemiformula n (L.imp p q) ↔ L.IsSemiformula n p ∧ L.IsSemiformula n q := by simp [Language.imp] section -def _root_.LO.FirstOrder.Arith.LDef.impDef (pL : LDef) : 𝚺₁.Semisentence 4 := .mkSigma - “r n p q | ∃ np, !pL.negDef np p ∧ !qqOrDef r n np q” (by simp) +def _root_.LO.FirstOrder.Arith.LDef.impDef (pL : LDef) : 𝚺₁.Semisentence 3 := .mkSigma + “r p q | ∃ np, !pL.negDef np p ∧ !qqOrDef r np q” (by simp) variable (L) -lemma imp_defined : 𝚺₁-Function₃ L.imp via pL.impDef := fun v ↦ by - simp [LDef.impDef, (neg_defined L).df.iff]; rfl - -@[simp] lemma eval_impDef (v : Fin 4 → V) : - Semiformula.Evalbm V v pL.impDef.val ↔ v 0 = L.imp (v 1) (v 2) (v 3) := (imp_defined L).df.iff v +lemma Language.imp_defined : 𝚺₁-Function₂ L.imp via pL.impDef := fun v ↦ by + simp [LDef.impDef, L.neg_defined.df.iff]; rfl -instance imp_definable : 𝚺₁-Function₃ L.imp := - (imp_defined L).to_definable +instance Language.imp_definable : 𝚺₁-Function₂ L.imp := L.imp_defined.to_definable -instance imp_definable' (Γ) : Γ-[m + 1]-Function₃ L.imp := - .of_sigmaOne (imp_definable L) _ _ +instance imp_definable' : Γ-[m + 1]-Function₂ L.imp := L.imp_definable.of_sigmaOne end @@ -195,32 +205,32 @@ section shift namespace Shift def blueprint (pL : LDef) : Language.UformulaRec1.Blueprint pL where - rel := .mkSigma “y param n k R v | ∃ v', !pL.termShiftVecDef v' k n v ∧ !qqRelDef y n k R v'” (by simp) - nrel := .mkSigma “y param n k R v | ∃ v', !pL.termShiftVecDef v' k n v ∧ !qqNRelDef y n k R v'” (by simp) - verum := .mkSigma “y param n | !qqVerumDef y n” (by simp) - falsum := .mkSigma “y param n | !qqFalsumDef y n” (by simp) - and := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqAndDef y n y₁ y₂” (by simp) - or := .mkSigma “y param n p₁ p₂ y₁ y₂ | !qqOrDef y n y₁ y₂” (by simp) - all := .mkSigma “y param n p₁ y₁ | !qqAllDef y n y₁” (by simp) - ex := .mkSigma “y param n p₁ y₁ | !qqExDef y n y₁” (by simp) - allChanges := .mkSigma “param' param n | param' = 0” (by simp) - exChanges := .mkSigma “param' param n | param' = 0” (by simp) + rel := .mkSigma “y param k R v | ∃ v', !pL.termShiftVecDef v' k v ∧ !qqRelDef y k R v'” (by simp) + nrel := .mkSigma “y param k R v | ∃ v', !pL.termShiftVecDef v' k v ∧ !qqNRelDef y k R v'” (by simp) + verum := .mkSigma “y param | !qqVerumDef y” (by simp) + falsum := .mkSigma “y param | !qqFalsumDef y” (by simp) + and := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqAndDef y y₁ y₂” (by simp) + or := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqOrDef y y₁ y₂” (by simp) + all := .mkSigma “y param p₁ y₁ | !qqAllDef y y₁” (by simp) + ex := .mkSigma “y param p₁ y₁ | !qqExDef y y₁” (by simp) + allChanges := .mkSigma “param' param | param' = 0” (by simp) + exChanges := .mkSigma “param' param | param' = 0” (by simp) variable (L) def construction : Language.UformulaRec1.Construction V L (blueprint pL) where - rel {_} := fun n k R v ↦ ^rel n k R (L.termShiftVec k n v) - nrel {_} := fun n k R v ↦ ^nrel n k R (L.termShiftVec k n v) - verum {_} := fun n ↦ ^⊤[n] - falsum {_} := fun n ↦ ^⊥[n] - and {_} := fun n _ _ y₁ y₂ ↦ y₁ ^⋏[n] y₂ - or {_} := fun n _ _ y₁ y₂ ↦ y₁ ^⋎[n] y₂ - all {_} := fun n _ y₁ ↦ ^∀[n] y₁ - ex {_} := fun n _ y₁ ↦ ^∃[n] y₁ - allChanges := fun _ _ ↦ 0 - exChanges := fun _ _ ↦ 0 - rel_defined := by intro v; simp [blueprint, (termShiftVec_defined L).df.iff]; rfl - nrel_defined := by intro v; simp [blueprint, (termShiftVec_defined L).df.iff]; rfl + rel {_} := fun k R v ↦ ^rel k R (L.termShiftVec k v) + nrel {_} := fun k R v ↦ ^nrel k R (L.termShiftVec k v) + verum {_} := ^⊤ + falsum {_} := ^⊥ + and {_} := fun _ _ y₁ y₂ ↦ y₁ ^⋏ y₂ + or {_} := fun _ _ y₁ y₂ ↦ y₁ ^⋎ y₂ + all {_} := fun _ y₁ ↦ ^∀ y₁ + ex {_} := fun _ y₁ ↦ ^∃ y₁ + allChanges := fun _ ↦ 0 + exChanges := fun _ ↦ 0 + rel_defined := by intro v; simp [blueprint, L.termShiftVec_defined.df.iff]; rfl + nrel_defined := by intro v; simp [blueprint, L.termShiftVec_defined.df.iff]; rfl verum_defined := by intro v; simp [blueprint] falsum_defined := by intro v; simp [blueprint] and_defined := by intro v; simp [blueprint]; rfl @@ -246,84 +256,94 @@ def _root_.LO.FirstOrder.Arith.LDef.shiftDef (pL : LDef) : 𝚺₁.Semisentence variable (L) -lemma shift_defined : 𝚺₁-Function₁ L.shift via pL.shiftDef := fun v ↦ by +lemma Language.shift_defined : 𝚺₁-Function₁ L.shift via pL.shiftDef := fun v ↦ by simpa [LDef.shiftDef] using (construction L).result_defined ![v 0, 0, v 1] -@[simp] lemma eval_shiftDef (v : Fin 2 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.shiftDef ↔ v 0 = L.shift (v 1) := (shift_defined L).df.iff v +instance Language.shift_definable : 𝚺₁-Function₁ L.shift := L.shift_defined.to_definable -instance shift_definable : 𝚺₁-Function₁ L.shift := - (shift_defined L).to_definable - -instance shift_definable' (Γ) : Γ-[m + 1]-Function₁ L.shift := - .of_sigmaOne (shift_definable L) _ _ +instance language.shift_definable' : Γ-[m + 1]-Function₁ L.shift := L.shift_definable.of_sigmaOne end -@[simp] lemma shift_rel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.shift (^rel n k R v) = ^rel n k R (L.termShiftVec k n v) := by simp [Language.shift, hR, hv, construction] +@[simp] lemma shift_rel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.shift (^relk R v) = ^relk R (L.termShiftVec k v) := by simp [Language.shift, hR, hv, construction] -@[simp] lemma shift_nrel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.shift (^nrel n k R v) = ^nrel n k R (L.termShiftVec k n v) := by simp [Language.shift, hR, hv, construction] +@[simp] lemma shift_nrel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.shift (^nrelk R v) = ^nrelk R (L.termShiftVec k v) := by simp [Language.shift, hR, hv, construction] -@[simp] lemma shift_verum (n) : - L.shift ^⊤[n] = ^⊤[n] := by simp [Language.shift, construction] +@[simp] lemma shift_verum : L.shift ^⊤ = ^⊤ := by simp [Language.shift, construction] -@[simp] lemma shift_falsum (n) : - L.shift ^⊥[n] = ^⊥[n] := by simp [Language.shift, construction] +@[simp] lemma shift_falsum : L.shift ^⊥ = ^⊥ := by simp [Language.shift, construction] -@[simp] lemma shift_and {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.shift (p ^⋏[n] q) = L.shift p ^⋏[n] L.shift q := by simp [Language.shift, hp, hq, construction] +@[simp] lemma shift_and {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.shift (p ^⋏ q) = L.shift p ^⋏ L.shift q := by simp [Language.shift, hp, hq, construction] -@[simp] lemma shift_or {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.shift (p ^⋎[n] q) = L.shift p ^⋎[n] L.shift q := by simp [Language.shift, hp, hq, construction] +@[simp] lemma shift_or {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.shift (p ^⋎ q) = L.shift p ^⋎ L.shift q := by simp [Language.shift, hp, hq, construction] -@[simp] lemma shift_all {n p} (hp : L.Semiformula (n + 1) p) : - L.shift (^∀[n] p) = ^∀[n] (L.shift p) := by simp [Language.shift, hp, construction] +@[simp] lemma shift_all {p} (hp : L.IsUFormula p) : + L.shift (^∀ p) = ^∀ (L.shift p) := by simp [Language.shift, hp, construction] -@[simp] lemma shift_ex {n p} (hp : L.Semiformula (n + 1) p) : - L.shift (^∃[n] p) = ^∃[n] (L.shift p) := by simp [Language.shift, hp, construction] +@[simp] lemma shift_ex {p} (hp : L.IsUFormula p) : + L.shift (^∃ p) = ^∃ (L.shift p) := by simp [Language.shift, hp, construction] -lemma shift_not_uformula {x} (h : ¬L.UFormula x) : +lemma shift_not_uformula {x} (h : ¬L.IsUFormula x) : L.shift x = 0 := (construction L).result_prop_not _ h -lemma Language.Semiformula.shift {p : V} : L.Semiformula n p → L.Semiformula n (L.shift p) := by - apply Language.Semiformula.induction_sigma₁ +lemma Language.IsUFormula.shift {p : V} : L.IsUFormula p → L.IsUFormula (L.shift p) := by + apply Language.IsUFormula.induction_sigma1 · definability - · intro n k r v hr hv; simp [hr, hv] - · intro n k r v hr hv; simp [hr, hv] + · intro k r v hr hv; simp [hr, hv] + · intro k r v hr hv; simp [hr, hv] · simp · simp - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p q hp hq ihp ihq; simp [hp, hq, ihp, ihq] - · intro n p hp ihp; simp [hp, ihp] - · intro n p hp ihp; simp [hp, ihp] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq] + · intro p hp ihp; simp [hp, hp.neg, ihp] + · intro p hp ihp; simp [hp, hp.neg, ihp] -lemma fstIdx_shift {p} (h : L.UFormula p) : fstIdx (L.shift p) = fstIdx p := by - rcases h.case with (⟨_, _, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, _, rfl⟩ | ⟨_, rfl⟩ | ⟨_, rfl⟩ | - ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, _, _, rfl⟩ | ⟨_, _, _, rfl⟩ | ⟨_, _, _, rfl⟩) <;> - simp [*] +lemma Language.IsUFormula.bv_shift {p : V} : L.IsUFormula p → L.bv (L.shift p) = L.bv p := by + apply Language.IsUFormula.induction_sigma1 + · definability + · intro k r v hr hv; simp [hr, hv] + · intro k r v hr hv; simp [hr, hv] + · simp + · simp + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq, hp.shift, hq.shift] + · intro p q hp hq ihp ihq; simp [hp, hq, hp.neg, hq.neg, ihp, ihq, hp.shift, hq.shift] + · intro p hp ihp; simp [hp, hp.neg, ihp, hp.shift] + · intro p hp ihp; simp [hp, hp.neg, ihp, hp.shift] -@[simp] lemma Language.Semiformula.shift_iff {p : V} : L.Semiformula n (L.shift p) ↔ L.Semiformula n p := +lemma Language.IsSemiformula.shift {p : V} : L.IsSemiformula n p → L.IsSemiformula n (L.shift p) := by + apply Language.IsSemiformula.induction_sigma1 + · definability + · intro n k r v hr hv; simp [hr, hv, hv.isUTerm] + · intro n k r v hr hv; simp [hr, hv, hv.isUTerm] + · simp + · simp + · intro n p q hp hq ihp ihq; simp [hp, hq, hp.isUFormula, hq.isUFormula, ihp, ihq] + · intro n p q hp hq ihp ihq; simp [hp, hq, hp.isUFormula, hq.isUFormula, ihp, ihq] + · intro n p hp ihp; simp [hp, hp.isUFormula, ihp] + · intro n p hp ihp; simp [hp, hp.isUFormula, ihp] + + +@[simp] lemma Language.IsSemiformula.shift_iff {p : V} : L.IsSemiformula n (L.shift p) ↔ L.IsSemiformula n p := ⟨fun h ↦ by - rcases h with ⟨h, rfl⟩ - have : L.UFormula p := by - by_contra hp - simp [shift_not_uformula hp] at h - exact ⟨this, by simp [fstIdx_shift this]⟩, - Language.Semiformula.shift⟩ - -lemma shift_neg {p : V} (hp : L.Semiformula n p) : L.shift (L.neg p) = L.neg (L.shift p) := by - apply Language.Semiformula.induction_sigma₁ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + have : L.IsUFormula p := by by_contra hp; simp [shift_not_uformula hp] at h + exact ⟨this, by simpa [this.bv_shift] using h.bv⟩, + Language.IsSemiformula.shift⟩ + +lemma shift_neg {p : V} (hp : L.IsSemiformula n p) : L.shift (L.neg p) = L.neg (L.shift p) := by + apply Language.IsSemiformula.induction_sigma1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp · definability - · intro n k R v hR hv; simp [*] - · intro n k R v hR hv; simp [*] - · intro n; simp - · intro n; simp - · intro n p q hp hq ihp ihq; simp [*] - · intro n p q hp hq ihp ihq; simp [*] - · intro n p hp ih; simp [*] - · intro n p hp ih; simp [*] + · intro n k R v hR hv; simp [hR, hv.isUTerm, hv.termShiftVec.isUTerm] + · intro n k R v hR hv; simp [hR, hv.isUTerm, hv.termShiftVec.isUTerm] + · simp + · simp + · intro n p q hp hq ihp ihq; simp [hp.isUFormula, hq.isUFormula, hp.shift.isUFormula, hq.shift.isUFormula, ihp, ihq] + · intro n p q hp hq ihp ihq; simp [hp.isUFormula, hq.isUFormula, hp.shift.isUFormula, hq.shift.isUFormula, ihp, ihq] + · intro n p hp ih; simp [hp.isUFormula, hp.shift.isUFormula, ih] + · intro n p hp ih; simp [hp.isUFormula, hp.shift.isUFormula, ih] end shift @@ -333,55 +353,55 @@ section variable (L) -def _root_.LO.FirstOrder.Arith.LDef.qVecDef (pL : LDef) : 𝚺₁.Semisentence 4 := .mkSigma - “w' k n w | ∃ sw, !pL.termBShiftVecDef sw k n w ∧ ∃ t, !qqBvarDef t 0 ∧ !consDef w' t sw” (by simp) +def _root_.LO.FirstOrder.Arith.LDef.qVecDef (pL : LDef) : 𝚺₁.Semisentence 2 := .mkSigma + “w' w | ∃ k, !lenDef k w ∧ ∃ sw, !pL.termBShiftVecDef sw k w ∧ ∃ t, !qqBvarDef t 0 ∧ !consDef w' t sw” (by simp) -lemma qVec_defined : 𝚺₁-Function₃ L.qVec via pL.qVecDef := by - intro v; simp [LDef.qVecDef, eval_termBShiftVecDef L]; rfl +lemma Language.qVec_defined : 𝚺₁-Function₁ L.qVec via pL.qVecDef := by + intro v; simp [LDef.qVecDef, L.termBShiftVec_defined.df.iff]; rfl -instance qVec_definable : 𝚺₁-Function₃ L.qVec := (qVec_defined L).to_definable +instance Language.qVec_definable : 𝚺₁-Function₁ L.qVec := L.qVec_defined.to_definable -@[simp, definability] instance qVec_definable' (Γ m) : Γ-[m + 1]-Function₃ L.qVec := .of_sigmaOne (qVec_definable L) _ _ +instance Language.qVec_definable' : Γ-[m + 1]-Function₁ L.qVec := L.qVec_definable.of_sigmaOne end namespace Substs def blueprint (pL : LDef) : Language.UformulaRec1.Blueprint pL where - rel := .mkSigma “y param n k R v | ∃ m, !pi₁Def m param ∧ ∃ w, !pi₂Def w param ∧ ∃ v', !pL.termSubstVecDef v' k n m w v ∧ !qqRelDef y m k R v'” (by simp) - nrel := .mkSigma “y param n k R v | ∃ m, !pi₁Def m param ∧ ∃ w, !pi₂Def w param ∧ ∃ v', !pL.termSubstVecDef v' k n m w v ∧ !qqNRelDef y m k R v'” (by simp) - verum := .mkSigma “y param n | ∃ m, !pi₁Def m param ∧ !qqVerumDef y m” (by simp) - falsum := .mkSigma “y param n | ∃ m, !pi₁Def m param ∧ !qqFalsumDef y m” (by simp) - and := .mkSigma “y param n p₁ p₂ y₁ y₂ | ∃ m, !pi₁Def m param ∧ !qqAndDef y m y₁ y₂” (by simp) - or := .mkSigma “y param n p₁ p₂ y₁ y₂ | ∃ m, !pi₁Def m param ∧ !qqOrDef y m y₁ y₂” (by simp) - all := .mkSigma “y param n p₁ y₁ | ∃ m, !pi₁Def m param ∧ !qqAllDef y m y₁” (by simp) - ex := .mkSigma “y param n p₁ y₁ | ∃ m, !pi₁Def m param ∧ !qqExDef y m y₁” (by simp) - allChanges := .mkSigma “param' param n | ∃ m, !pi₁Def m param ∧ ∃ w, !pi₂Def w param ∧ ∃ qseq, !pL.qVecDef qseq n m w ∧ !pairDef param' (m + 1) qseq” (by simp) - exChanges := .mkSigma “param' param n | ∃ m, !pi₁Def m param ∧ ∃ w, !pi₂Def w param ∧ ∃ qseq, !pL.qVecDef qseq n m w ∧ !pairDef param' (m + 1) qseq” (by simp) + rel := .mkSigma “y param k R v | ∃ v', !pL.termSubstVecDef v' k param v ∧ !qqRelDef y k R v'” (by simp) + nrel := .mkSigma “y param k R v | ∃ v', !pL.termSubstVecDef v' k param v ∧ !qqNRelDef y k R v'” (by simp) + verum := .mkSigma “y param | !qqVerumDef y” (by simp) + falsum := .mkSigma “y param | !qqFalsumDef y” (by simp) + and := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqAndDef y y₁ y₂” (by simp) + or := .mkSigma “y param p₁ p₂ y₁ y₂ | !qqOrDef y y₁ y₂” (by simp) + all := .mkSigma “y param p₁ y₁ | !qqAllDef y y₁” (by simp) + ex := .mkSigma “y param p₁ y₁ | !qqExDef y y₁” (by simp) + allChanges := .mkSigma “param' param | !pL.qVecDef param' param” (by simp) + exChanges := .mkSigma “param' param | !pL.qVecDef param' param” (by simp) variable (L) def construction : Language.UformulaRec1.Construction V L (blueprint pL) where - rel (param) := fun n k R v ↦ ^rel (π₁ param) k R (L.termSubstVec k n (π₁ param) (π₂ param) v) - nrel (param) := fun n k R v ↦ ^nrel (π₁ param) k R (L.termSubstVec k n (π₁ param) (π₂ param) v) - verum (param) := fun _ ↦ ^⊤[π₁ param] - falsum (param) := fun _ ↦ ^⊥[π₁ param] - and (param) := fun _ _ _ y₁ y₂ ↦ y₁ ^⋏[π₁ param] y₂ - or (param) := fun _ _ _ y₁ y₂ ↦ y₁ ^⋎[π₁ param] y₂ - all (param) := fun _ _ y₁ ↦ ^∀[π₁ param] y₁ - ex (param) := fun _ _ y₁ ↦ ^∃[π₁ param] y₁ - allChanges (param n) := ⟪π₁ param + 1, L.qVec n (π₁ param) (π₂ param)⟫ - exChanges (param n) := ⟪π₁ param + 1, L.qVec n (π₁ param) (π₂ param)⟫ - rel_defined := by intro v; simp [blueprint, (termSubstVec_defined L).df.iff]; rfl - nrel_defined := by intro v; simp [blueprint, (termSubstVec_defined L).df.iff]; rfl + rel (param) := fun k R v ↦ ^rel k R (L.termSubstVec k param v) + nrel (param) := fun k R v ↦ ^nrel k R (L.termSubstVec k param v) + verum _ := ^⊤ + falsum _ := ^⊥ + and _ := fun _ _ y₁ y₂ ↦ y₁ ^⋏ y₂ + or _ := fun _ _ y₁ y₂ ↦ y₁ ^⋎ y₂ + all _ := fun _ y₁ ↦ ^∀ y₁ + ex _ := fun _ y₁ ↦ ^∃ y₁ + allChanges (param) := L.qVec param + exChanges (param) := L.qVec param + rel_defined := by intro v; simp [blueprint, L.termSubstVec_defined.df.iff]; rfl + nrel_defined := by intro v; simp [blueprint, L.termSubstVec_defined.df.iff]; rfl verum_defined := by intro v; simp [blueprint] falsum_defined := by intro v; simp [blueprint] and_defined := by intro v; simp [blueprint]; rfl or_defined := by intro v; simp [blueprint]; rfl all_defined := by intro v; simp [blueprint]; rfl ex_defined := by intro v; simp [blueprint]; rfl - allChanges_defined := by intro v; simp [blueprint, (qVec_defined L).df.iff] - exChanges_defined := by intro v; simp [blueprint, (qVec_defined L).df.iff] + allChanges_defined := by intro v; simp [blueprint, L.qVec_defined.df.iff] + exChanges_defined := by intro v; simp [blueprint, L.qVec_defined.df.iff] end Substs @@ -389,310 +409,334 @@ open Substs variable (L) -def Language.substs (m w p : V) : V := (construction L).result ⟪m, w⟫ p +def Language.substs (w p : V) : V := (construction L).result w p variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.substsDef (pL : LDef) : 𝚺₁.Semisentence 4 := .mkSigma - “q m w p | ∃ mw, !pairDef mw m w ∧ !(blueprint pL).result q mw p” (by simp) +def _root_.LO.FirstOrder.Arith.LDef.substsDef (pL : LDef) : 𝚺₁.Semisentence 3 := (blueprint pL).result variable (L) -lemma substs_defined : 𝚺₁-Function₃ L.substs via pL.substsDef := fun v ↦ by - simp [LDef.substsDef, (construction L).result_defined.df.iff]; rfl - -@[simp] lemma eval_substsDef (v : Fin 4 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.substsDef ↔ v 0 = L.substs (v 1) (v 2) (v 3) := (substs_defined L).df.iff v +lemma Language.substs_defined : 𝚺₁-Function₂ L.substs via pL.substsDef := (construction L).result_defined -instance substs_definable : 𝚺₁-Function₃ L.substs := - (substs_defined L).to_definable +instance Language.substs_definable : 𝚺₁-Function₂ L.substs := L.substs_defined.to_definable -instance substs_definable' (Γ) : Γ-[m + 1]-Function₃ L.substs := - .of_sigmaOne (substs_definable L) _ _ +instance Language.substs_definable' : Γ-[m + 1]-Function₂ L.substs := L.substs_definable.of_sigmaOne end variable {m w : V} -@[simp] lemma substs_rel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.substs m w (^rel n k R v) = ^rel m k R (L.termSubstVec k n m w v) := by simp [Language.substs, hR, hv, construction] - -@[simp] lemma substs_nrel {n k R v} (hR : L.Rel k R) (hv : L.SemitermVec k n v) : - L.substs m w (^nrel n k R v) = ^nrel m k R (L.termSubstVec k n m w v) := by simp [Language.substs, hR, hv, construction] - -@[simp] lemma substs_verum (n) : - L.substs m w ^⊤[n] = ^⊤[m] := by simp [Language.substs, construction] - -@[simp] lemma substs_falsum (n) : - L.substs m w ^⊥[n] = ^⊥[m] := by simp [Language.substs, construction] - -@[simp] lemma substs_and {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.substs m w (p ^⋏[n] q) = L.substs m w p ^⋏[m] L.substs m w q := by simp [Language.substs, hp, hq, construction] - -@[simp] lemma substs_or {n p q} (hp : L.Semiformula n p) (hq : L.Semiformula n q) : - L.substs m w (p ^⋎[n] q) = L.substs m w p ^⋎[m] L.substs m w q := by simp [Language.substs, hp, hq, construction] - -@[simp] lemma substs_all {n p} (hp : L.Semiformula (n + 1) p) : - L.substs m w (^∀[n] p) = ^∀[m] (L.substs (m + 1) (L.qVec n m w) p) := by simp [Language.substs, hp, construction] - -@[simp] lemma substs_ex {n p} (hp : L.Semiformula (n + 1) p) : - L.substs m w (^∃[n] p) = ^∃[m] (L.substs (m + 1) (L.qVec n m w) p) := by simp [Language.substs, hp, construction] - -lemma uformula_subst_induction {P : V → V → V → V → Prop} (hP : 𝚺₁-Relation₄ P) - (hRel : ∀ n m w k R v, L.Rel k R → L.SemitermVec k n v → P m w (^rel n k R v) (^rel m k R (L.termSubstVec k n m w v))) - (hNRel : ∀ n m w k R v, L.Rel k R → L.SemitermVec k n v → P m w (^nrel n k R v) (^nrel m k R (L.termSubstVec k n m w v))) - (hverum : ∀ n m w, P m w (^⊤[n]) (^⊤[m])) - (hfalsum : ∀ n m w, P m w (^⊥[n]) (^⊥[m])) - (hand : ∀ n m w p q, L.Semiformula n p → L.Semiformula n q → - P m w p (L.substs m w p) → P m w q (L.substs m w q) → P m w (p ^⋏[n] q) (L.substs m w p ^⋏[m] L.substs m w q)) - (hor : ∀ n m w p q, L.Semiformula n p → L.Semiformula n q → - P m w p (L.substs m w p) → P m w q (L.substs m w q) → P m w (p ^⋎[n] q) (L.substs m w p ^⋎[m] L.substs m w q)) - (hall : ∀ n m w p, L.Semiformula (n + 1) p → - P (m + 1) (L.qVec n m w) p (L.substs (m + 1) (L.qVec n m w) p) → - P m w (^∀[n] p) (^∀[m] (L.substs (m + 1) (L.qVec n m w) p))) - (hex : ∀ n m w p, L.Semiformula (n + 1) p → - P (m + 1) (L.qVec n m w) p (L.substs (m + 1) (L.qVec n m w) p) → - P m w (^∃[n] p) (^∃[m] (L.substs (m + 1) (L.qVec n m w) p))) : - ∀ {p m w}, L.UFormula p → P m w p (L.substs m w p) := by - suffices ∀ param p, L.UFormula p → P (π₁ param) (π₂ param) p ((construction L).result param p) by - intro p m w hp; simpa using this ⟪m, w⟫ p hp - apply (construction L).uformula_result_induction (P := fun param p y ↦ P (π₁ param) (π₂ param) p y) - · apply HierarchySymbol.Boldface.comp₄ - (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) - (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) - (HierarchySymbol.BoldfaceFunction.var _) - (HierarchySymbol.BoldfaceFunction.var _) - · intro param n k R v hkR hv; simpa using hRel n (π₁ param) (π₂ param) k R v hkR hv - · intro param n k R v hkR hv; simpa using hNRel n (π₁ param) (π₂ param) k R v hkR hv - · intro param n; simpa using hverum n (π₁ param) (π₂ param) - · intro param n; simpa using hfalsum n (π₁ param) (π₂ param) - · intro param n p q hp hq ihp ihq +@[simp] lemma substs_rel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.substs w (^relk R v) = ^rel k R (L.termSubstVec k w v) := by simp [Language.substs, hR, hv, construction] + +@[simp] lemma substs_nrel {k R v} (hR : L.Rel k R) (hv : L.IsUTermVec k v) : + L.substs w (^nrelk R v) = ^nrel k R (L.termSubstVec k w v) := by simp [Language.substs, hR, hv, construction] + +@[simp] lemma substs_verum (w) : L.substs w ^⊤ = ^⊤ := by simp [Language.substs, construction] + +@[simp] lemma substs_falsum (w) : L.substs w ^⊥ = ^⊥ := by simp [Language.substs, construction] + +@[simp] lemma substs_and {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.substs w (p ^⋏ q) = L.substs w p ^⋏ L.substs w q := by simp [Language.substs, hp, hq, construction] + +@[simp] lemma substs_or {p q} (hp : L.IsUFormula p) (hq : L.IsUFormula q) : + L.substs w (p ^⋎ q) = L.substs w p ^⋎ L.substs w q := by simp [Language.substs, hp, hq, construction] + +@[simp] lemma substs_all {p} (hp : L.IsUFormula p) : + L.substs w (^∀ p) = ^∀ (L.substs (L.qVec w) p) := by simp [Language.substs, hp, construction] + +@[simp] lemma substs_ex {p} (hp : L.IsUFormula p) : + L.substs w (^∃ p) = ^∃ (L.substs (L.qVec w) p) := by simp [Language.substs, hp, construction] + +lemma isUFormula_subst_induction_sigma1 {P : V → V → V → Prop} (hP : 𝚺₁-Relation₃ P) + (hRel : ∀ w k R v, L.Rel k R → L.IsUTermVec k v → P w (^relk R v) (^rel k R (L.termSubstVec k w v))) + (hNRel : ∀ w k R v, L.Rel k R → L.IsUTermVec k v → P w (^nrelk R v) (^nrel k R (L.termSubstVec k w v))) + (hverum : ∀ w, P w ^⊤ ^⊤) + (hfalsum : ∀ w, P w ^⊥ ^⊥) + (hand : ∀ w p q, L.IsUFormula p → L.IsUFormula q → + P w p (L.substs w p) → P w q (L.substs w q) → P w (p ^⋏ q) (L.substs w p ^⋏ L.substs w q)) + (hor : ∀ w p q, L.IsUFormula p → L.IsUFormula q → + P w p (L.substs w p) → P w q (L.substs w q) → P w (p ^⋎ q) (L.substs w p ^⋎ L.substs w q)) + (hall : ∀ w p, L.IsUFormula p → P (L.qVec w) p (L.substs (L.qVec w) p) → P w (^∀ p) (^∀ (L.substs (L.qVec w) p))) + (hex : ∀ w p, L.IsUFormula p → P (L.qVec w) p (L.substs (L.qVec w) p) → P w (^∃ p) (^∃ (L.substs (L.qVec w) p))) : + ∀ {w p}, L.IsUFormula p → P w p (L.substs w p) := by + suffices ∀ param p, L.IsUFormula p → P param p ((construction L).result param p) by + intro w p hp; simpa using this w p hp + apply (construction L).uformula_result_induction (P := fun param p y ↦ P param p y) + · definability + · intro param k R v hkR hv; simpa using hRel param k R v hkR hv + · intro param k R v hkR hv; simpa using hNRel param k R v hkR hv + · intro param; simpa using hverum param + · intro param; simpa using hfalsum param + · intro param p q hp hq ihp ihq simpa [Language.substs] using - hand n (π₁ param) (π₂ param) p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) - · intro param n p q hp hq ihp ihq + hand param p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) + · intro param p q hp hq ihp ihq simpa [Language.substs] using - hor n (π₁ param) (π₂ param) p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) - · intro param n p hp ihp - simpa using hall n (π₁ param) (π₂ param) p hp (by simpa [construction] using ihp) - · intro param n p hp ihp - simpa using hex n (π₁ param) (π₂ param) p hp (by simpa [construction] using ihp) - -lemma semiformula_subst_induction {P : V → V → V → V → V → Prop} (hP : 𝚺₁-Relation₅ P) - (hRel : ∀ n m w k R v, L.Rel k R → L.SemitermVec k n v → P n m w (^rel n k R v) (^rel m k R (L.termSubstVec k n m w v))) - (hNRel : ∀ n m w k R v, L.Rel k R → L.SemitermVec k n v → P n m w (^nrel n k R v) (^nrel m k R (L.termSubstVec k n m w v))) - (hverum : ∀ n m w, P n m w (^⊤[n]) (^⊤[m])) - (hfalsum : ∀ n m w, P n m w (^⊥[n]) (^⊥[m])) - (hand : ∀ n m w p q, L.Semiformula n p → L.Semiformula n q → - P n m w p (L.substs m w p) → P n m w q (L.substs m w q) → P n m w (p ^⋏[n] q) (L.substs m w p ^⋏[m] L.substs m w q)) - (hor : ∀ n m w p q, L.Semiformula n p → L.Semiformula n q → - P n m w p (L.substs m w p) → P n m w q (L.substs m w q) → P n m w (p ^⋎[n] q) (L.substs m w p ^⋎[m] L.substs m w q)) - (hall : ∀ n m w p, L.Semiformula (n + 1) p → - P (n + 1) (m + 1) (L.qVec n m w) p (L.substs (m + 1) (L.qVec n m w) p) → - P n m w (^∀[n] p) (^∀[m] (L.substs (m + 1) (L.qVec n m w) p))) - (hex : ∀ n m w p, L.Semiformula (n + 1) p → - P (n + 1) (m + 1) (L.qVec n m w) p (L.substs (m + 1) (L.qVec n m w) p) → - P n m w (^∃[n] p) (^∃[m] (L.substs (m + 1) (L.qVec n m w) p))) : - ∀ {n p m w}, L.Semiformula n p → P n m w p (L.substs m w p) := by - suffices ∀ param n p, L.Semiformula n p → P n (π₁ param) (π₂ param) p ((construction L).result param p) by - intro n p m w hp; simpa using this ⟪m, w⟫ n p hp - apply (construction L).semiformula_result_induction (P := fun param n p y ↦ P n (π₁ param) (π₂ param) p y) - · apply HierarchySymbol.Boldface.comp₅ - (HierarchySymbol.BoldfaceFunction.var _) - (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) - (HierarchySymbol.BoldfaceFunction.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) - (HierarchySymbol.BoldfaceFunction.var _) - (HierarchySymbol.BoldfaceFunction.var _) - · intro param n k R v hkR hv; simpa using hRel n (π₁ param) (π₂ param) k R v hkR hv - · intro param n k R v hkR hv; simpa using hNRel n (π₁ param) (π₂ param) k R v hkR hv - · intro param n; simpa using hverum n (π₁ param) (π₂ param) - · intro param n; simpa using hfalsum n (π₁ param) (π₂ param) - · intro param n p q hp hq ihp ihq + hor param p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) + · intro param p hp ihp + simpa using hall param p hp (by simpa [construction] using ihp) + · intro param p hp ihp + simpa using hex param p hp (by simpa [construction] using ihp) + +lemma semiformula_subst_induction {P : V → V → V → V → Prop} (hP : 𝚺₁-Relation₄ P) + (hRel : ∀ n w k R v, L.Rel k R → L.IsSemitermVec k n v → P n w (^relk R v) (^rel k R (L.termSubstVec k w v))) + (hNRel : ∀ n w k R v, L.Rel k R → L.IsSemitermVec k n v → P n w (^nrelk R v) (^nrel k R (L.termSubstVec k w v))) + (hverum : ∀ n w, P n w ^⊤ ^⊤) + (hfalsum : ∀ n w, P n w ^⊥ ^⊥) + (hand : ∀ n w p q, L.IsSemiformula n p → L.IsSemiformula n q → + P n w p (L.substs w p) → P n w q (L.substs w q) → P n w (p ^⋏ q) (L.substs w p ^⋏ L.substs w q)) + (hor : ∀ n w p q, L.IsSemiformula n p → L.IsSemiformula n q → + P n w p (L.substs w p) → P n w q (L.substs w q) → P n w (p ^⋎ q) (L.substs w p ^⋎ L.substs w q)) + (hall : ∀ n w p, L.IsSemiformula (n + 1) p → + P (n + 1) (L.qVec w) p (L.substs (L.qVec w) p) → P n w (^∀ p) (^∀ (L.substs (L.qVec w) p))) + (hex : ∀ n w p, L.IsSemiformula (n + 1) p → + P (n + 1) (L.qVec w) p (L.substs (L.qVec w) p) → P n w (^∃ p) (^∃ (L.substs (L.qVec w) p))) : + ∀ {n p w}, L.IsSemiformula n p → P n w p (L.substs w p) := by + suffices ∀ param n p, L.IsSemiformula n p → P n param p ((construction L).result param p) by + intro n p w hp; simpa using this w n p hp + apply (construction L).semiformula_result_induction (P := fun param n p y ↦ P n param p y) + · definability + · intro n param k R v hkR hv; simpa using hRel n param k R v hkR hv + · intro n param k R v hkR hv; simpa using hNRel n param k R v hkR hv + · intro n param; simpa using hverum n param + · intro n param; simpa using hfalsum n param + · intro n param p q hp hq ihp ihq simpa [Language.substs] using - hand n (π₁ param) (π₂ param) p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) - · intro param n p q hp hq ihp ihq + hand n param p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) + · intro n param p q hp hq ihp ihq simpa [Language.substs] using - hor n (π₁ param) (π₂ param) p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) - · intro param n p hp ihp - simpa using hall n (π₁ param) (π₂ param) p hp (by simpa [construction] using ihp) - · intro param n p hp ihp - simpa using hex n (π₁ param) (π₂ param) p hp (by simpa [construction] using ihp) - -@[simp] lemma Language.Semiformula.substs {n p m w : V} : - L.Semiformula n p → L.SemitermVec n m w → L.Semiformula m (L.substs m w p) := by - apply semiformula_subst_induction (P := fun n m w _ y ↦ L.SemitermVec n m w → L.Semiformula m y) + hor n param p q hp hq (by simpa [Language.substs] using ihp) (by simpa [Language.substs] using ihq) + · intro n param p hp ihp + simpa using hall n param p hp (by simpa [construction] using ihp) + · intro n param p hp ihp + simpa using hex n param p hp (by simpa [construction] using ihp) + +@[simp] lemma Language.IsSemiformula.substs {n p m w : V} : + L.IsSemiformula n p → L.IsSemitermVec n m w → L.IsSemiformula m (L.substs w p) := by + let fw : V → V → V → V → V := fun _ w _ _ ↦ Max.max w (L.qVec w) + have hfw : 𝚺₁-Function₄ fw := by simp [fw]; definability + let fn : V → V → V → V → V := fun _ _ n _ ↦ n + 1 + have hfn : 𝚺₁-Function₄ fn := by simp [fn]; definability + let fm : V → V → V → V → V := fun _ _ _ m ↦ m + 1 + have hfm : 𝚺₁-Function₄ fm := by simp [fm]; definability + apply order_ball_induction₃_sigma1 hfw hfn hfm ?_ ?_ p w n m · definability - case hRel => intro n m w k R v hR hv hw; simp [hR, hv, hw] - case hNRel => intro n m w k R v hR hv hw; simp [hR, hv, hw] - case hverum => simp - case hfalsum => simp - case hand => - intro n m w p q _ _ ihp ihq hw - simp [ihp hw, ihq hw] - case hor => - intro n m w p q _ _ ihp ihq hw - simp [ihp hw, ihq hw] - case hall => - intro n m w p _ ih hw - simpa using ih hw.qVec - case hex => - intro n m w p _ ih hw - simpa using ih hw.qVec - -lemma substs_not_uformula {m w x} (h : ¬L.UFormula x) : - L.substs m w x = 0 := (construction L).result_prop_not _ h - -lemma substs_neg {p} (hp : L.Semiformula n p) : - L.SemitermVec n m w → L.substs m w (L.neg p) = L.neg (L.substs m w p) := by + intro p w n m ih hp hw + rcases Language.IsSemiformula.case_iff.mp hp with + (⟨k, R, v, hR, hv, rfl⟩ | ⟨k, R, v, hR, hv, rfl⟩ | rfl | rfl | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, p₂, h₁, h₂, rfl⟩ | ⟨p₁, h₁, rfl⟩ | ⟨p₁, h₁, rfl⟩) + · simp [hR, hv.isUTerm, hw.termSubstVec hv] + · simp [hR, hv.isUTerm, hw.termSubstVec hv] + · simp + · simp + · have ih₁ : L.IsSemiformula m (L.substs w p₁) := ih p₁ (by simp) w (by simp [fw]) n (by simp [fn]) m (by simp [fm]) h₁ hw + have ih₂ : L.IsSemiformula m (L.substs w p₂) := ih p₂ (by simp) w (by simp [fw]) n (by simp [fn]) m (by simp [fm]) h₂ hw + simp [h₁.isUFormula, h₂.isUFormula, ih₁, ih₂] + · have ih₁ : L.IsSemiformula m (L.substs w p₁) := ih p₁ (by simp) w (by simp [fw]) n (by simp [fn]) m (by simp [fm]) h₁ hw + have ih₂ : L.IsSemiformula m (L.substs w p₂) := ih p₂ (by simp) w (by simp [fw]) n (by simp [fn]) m (by simp [fm]) h₂ hw + simp [h₁.isUFormula, h₂.isUFormula, ih₁, ih₂] + · simpa [h₁.isUFormula] using ih p₁ (by simp) (L.qVec w) (by simp [fw]) (n + 1) (by simp [fn]) (m + 1) (by simp [fm]) h₁ hw.qVec + · simpa [h₁.isUFormula] using ih p₁ (by simp) (L.qVec w) (by simp [fw]) (n + 1) (by simp [fn]) (m + 1) (by simp [fm]) h₁ hw.qVec + +lemma substs_not_uformula {w x} (h : ¬L.IsUFormula x) : + L.substs w x = 0 := (construction L).result_prop_not _ h + +lemma substs_neg {p} (hp : L.IsSemiformula n p) : + L.IsSemitermVec n m w → L.substs w (L.neg p) = L.neg (L.substs w p) := by revert m w - apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.IsSemiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp · definability - · intros; simp [*] - · intros; simp [*] + · intros n k R v hR hv m w hw + rw [neg_rel hR hv.isUTerm, substs_nrel hR hv.isUTerm, substs_rel hR hv.isUTerm, neg_rel hR (hw.termSubstVec hv).isUTerm] + · intros n k R v hR hv m w hw + rw [neg_nrel hR hv.isUTerm, substs_rel hR hv.isUTerm, substs_nrel hR hv.isUTerm, neg_nrel hR (hw.termSubstVec hv).isUTerm] · intros; simp [*] · intros; simp [*] · intro n p q hp hq ihp ihq m w hw - simp [hp, hq, hw, hp.substs, hq.substs, ihp hw, ihq hw] + rw [neg_and hp.isUFormula hq.isUFormula, + substs_or hp.neg.isUFormula hq.neg.isUFormula, + substs_and hp.isUFormula hq.isUFormula, + neg_and (hp.substs hw).isUFormula (hq.substs hw).isUFormula, + ihp hw, ihq hw] · intro n p q hp hq ihp ihq m w hw - simp [hp, hq, hw, hp.substs, hq.substs, ihp hw, ihq hw] + rw [neg_or hp.isUFormula hq.isUFormula, + substs_and hp.neg.isUFormula hq.neg.isUFormula, + substs_or hp.isUFormula hq.isUFormula, + neg_or (hp.substs hw).isUFormula (hq.substs hw).isUFormula, + ihp hw, ihq hw] · intro n p hp ih m w hw - simp [hp, hw, hp.substs hw.qVec, ih hw.qVec] + rw [neg_all hp.isUFormula, substs_ex hp.neg.isUFormula, + substs_all hp.isUFormula, neg_all (hp.substs hw.qVec).isUFormula, ih hw.qVec] · intro n p hp ih m w hw - simp [hp, hw, hp.substs hw.qVec, ih hw.qVec] + rw [neg_ex hp.isUFormula, substs_all hp.neg.isUFormula, + substs_ex hp.isUFormula, neg_ex (hp.substs hw.qVec).isUFormula, ih hw.qVec] -lemma shift_substs {p} (hp : L.Semiformula n p) : - L.SemitermVec n m w → L.shift (L.substs m w p) = L.substs m (L.termShiftVec n m w) (L.shift p) := by +lemma shift_substs {p} (hp : L.IsSemiformula n p) : + L.IsSemitermVec n m w → L.shift (L.substs w p) = L.substs (L.termShiftVec n w) (L.shift p) := by revert m w - apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.IsSemiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp · definability · intro n k R v hR hv m w hw - simp only [substs_rel, Language.SemitermVec.termSubstVec, shift_rel, - Language.SemitermVec.termShiftVec, qqRel_inj, true_and, hR, hv, hw] - apply nth_ext' k (by simp [*]) (by simp [*]) + rw [substs_rel hR hv.isUTerm, + shift_rel hR (hw.termSubstVec hv).isUTerm, + shift_rel hR hv.isUTerm, + substs_rel hR hv.termShiftVec.isUTerm] + simp only [qqRel_inj, true_and] + apply nth_ext' k + (by rw [len_termShiftVec (hw.termSubstVec hv).isUTerm]) + (by rw [len_termSubstVec hv.termShiftVec.isUTerm]) intro i hi - rw [nth_termShiftVec (hw.termSubstVec hv) hi, - nth_termSubstVec hv hi, - nth_termSubstVec hv.termShiftVec hi, - nth_termShiftVec hv hi, - termShift_termSubsts (hv.2 i hi) hw] + rw [nth_termShiftVec (hw.termSubstVec hv).isUTerm hi, + nth_termSubstVec hv.isUTerm hi, + nth_termSubstVec hv.termShiftVec.isUTerm hi, + nth_termShiftVec hv.isUTerm hi, + termShift_termSubsts (hv.nth hi) hw] · intro n k R v hR hv m w hw - simp only [substs_nrel, Language.SemitermVec.termSubstVec, shift_nrel, - Language.SemitermVec.termShiftVec, qqNRel_inj, true_and, hR, hv, hw] - apply nth_ext' k (by simp [*]) (by simp [*]) + rw [substs_nrel hR hv.isUTerm, + shift_nrel hR (hw.termSubstVec hv).isUTerm, + shift_nrel hR hv.isUTerm, + substs_nrel hR hv.termShiftVec.isUTerm] + simp only [qqNRel_inj, true_and] + apply nth_ext' k + (by rw [len_termShiftVec (hw.termSubstVec hv).isUTerm]) + (by rw [len_termSubstVec hv.termShiftVec.isUTerm]) intro i hi - rw [nth_termShiftVec (hw.termSubstVec hv) hi, - nth_termSubstVec hv hi, - nth_termSubstVec hv.termShiftVec hi, - nth_termShiftVec hv hi, - termShift_termSubsts (hv.2 i hi) hw] + rw [nth_termShiftVec (hw.termSubstVec hv).isUTerm hi, + nth_termSubstVec hv.isUTerm hi, + nth_termSubstVec hv.termShiftVec.isUTerm hi, + nth_termShiftVec hv.isUTerm hi, + termShift_termSubsts (hv.nth hi) hw] · intro n w hw; simp · intro n w hw; simp · intro n p q hp hq ihp ihq m w hw - simp [*] - rw [shift_and (hp.substs hw) (hq.substs hw), ihp hw, ihq hw] + rw [substs_and hp.isUFormula hq.isUFormula, + shift_and (hp.substs hw).isUFormula (hq.substs hw).isUFormula, + shift_and hp.isUFormula hq.isUFormula, + substs_and hp.shift.isUFormula hq.shift.isUFormula, + ihp hw, ihq hw] · intro n p q hp hq ihp ihq m w hw - simp [*] - rw [shift_or (hp.substs hw) (hq.substs hw), ihp hw, ihq hw] + rw [substs_or hp.isUFormula hq.isUFormula, + shift_or (hp.substs hw).isUFormula (hq.substs hw).isUFormula, + shift_or hp.isUFormula hq.isUFormula, + substs_or hp.shift.isUFormula hq.shift.isUFormula, + ihp hw, ihq hw] · intro n p hp ih m w hw - simp only [substs_all, shift_all, Language.Semiformula.shift_iff, hp] - rw [shift_all (hp.substs hw.qVec), ih hw.qVec, termShift_qVec hw] + rw [substs_all hp.isUFormula, + shift_all (hp.substs hw.qVec).isUFormula, + shift_all hp.isUFormula, + substs_all hp.shift.isUFormula, + ih hw.qVec, + termShift_qVec hw] · intro n p hp ih m w hw - simp only [substs_ex, shift_ex, Language.Semiformula.shift_iff, hp] - rw [shift_ex (hp.substs hw.qVec), ih hw.qVec, termShift_qVec hw] - -lemma substs_substs {p} (hp : L.Semiformula l p) : - L.SemitermVec n m w → L.SemitermVec l n v → L.substs m w (L.substs n v p) = L.substs m (L.termSubstVec l n m w v) p := by + rw [substs_ex hp.isUFormula, + shift_ex (hp.substs hw.qVec).isUFormula, + shift_ex hp.isUFormula, + substs_ex hp.shift.isUFormula, + ih hw.qVec, + termShift_qVec hw] + +lemma substs_substs {p} (hp : L.IsSemiformula l p) : + L.IsSemitermVec n m w → L.IsSemitermVec l n v → L.substs w (L.substs v p) = L.substs (L.termSubstVec l w v) p := by revert m w n v - apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp - · apply HierarchySymbol.Boldface.all - apply HierarchySymbol.Boldface.all - apply HierarchySymbol.Boldface.all - apply HierarchySymbol.Boldface.all - apply HierarchySymbol.Boldface.imp (by definability) - apply HierarchySymbol.Boldface.imp (by definability) - apply HierarchySymbol.Boldface.comp₂ (by simp; definability) - apply HierarchySymbol.BoldfaceFunction.comp₃ (by definability) ?_ (by definability) - apply HierarchySymbol.BoldfaceFunction₅.comp (termSubstVec_definable _) <;> definability + apply Language.IsSemiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp + · definability · intro l k R ts hR hts m w n v _ hv - simp only [substs_rel, Language.SemitermVec.termSubstVec, qqRel_inj, true_and, hR, hts, hv] - apply nth_ext' k (by simp [hv, hts]) (by simp [hts]) + rw [substs_rel hR hts.isUTerm, + substs_rel hR (hv.termSubstVec hts).isUTerm, + substs_rel hR hts.isUTerm] + simp only [qqRel_inj, true_and] + apply nth_ext' k (by rw [len_termSubstVec (hv.termSubstVec hts).isUTerm]) (by rw [len_termSubstVec hts.isUTerm]) intro i hi - rw [nth_termSubstVec (hv.termSubstVec hts) hi, - nth_termSubstVec hts hi, - nth_termSubstVec hts hi, - termSubst_termSubst hv (hts.2 i hi)] + rw [nth_termSubstVec (hv.termSubstVec hts).isUTerm hi, + nth_termSubstVec hts.isUTerm hi, + nth_termSubstVec hts.isUTerm hi, + termSubst_termSubst hv (hts.nth hi)] · intro l k R ts hR hts m w n v _ hv - simp only [substs_nrel, Language.SemitermVec.termSubstVec, qqNRel_inj, true_and, hR, hts, hv] - apply nth_ext' k (by simp [hv, hts]) (by simp [hts]) + rw [substs_nrel hR hts.isUTerm, + substs_nrel hR (hv.termSubstVec hts).isUTerm, + substs_nrel hR hts.isUTerm] + simp only [qqNRel_inj, true_and] + apply nth_ext' k (by rw [len_termSubstVec (hv.termSubstVec hts).isUTerm]) (by rw [len_termSubstVec hts.isUTerm]) intro i hi - rw [nth_termSubstVec (hv.termSubstVec hts) hi, - nth_termSubstVec hts hi, - nth_termSubstVec hts hi, - termSubst_termSubst hv (hts.2 i hi)] - · intro l m w n v _ _; simp [*] - · intro l m w n v _ _; simp [*] + rw [nth_termSubstVec (hv.termSubstVec hts).isUTerm hi, + nth_termSubstVec hts.isUTerm hi, + nth_termSubstVec hts.isUTerm hi, + termSubst_termSubst hv (hts.nth hi)] + · intros; simp + · intros; simp · intro l p q hp hq ihp ihq m w n v hw hv - simp only [substs_and, hp, hq] - rw [substs_and (hp.substs hv) (hq.substs hv), ihp hw hv, ihq hw hv] + rw [substs_and hp.isUFormula hq.isUFormula, + substs_and (hp.substs hv).isUFormula (hq.substs hv).isUFormula, + substs_and hp.isUFormula hq.isUFormula, + ihp hw hv, ihq hw hv] · intro l p q hp hq ihp ihq m w n v hw hv - simp only [substs_or, hp, hq] - rw [substs_or (hp.substs hv) (hq.substs hv), ihp hw hv, ihq hw hv] + rw [substs_or hp.isUFormula hq.isUFormula, + substs_or (hp.substs hv).isUFormula (hq.substs hv).isUFormula, + substs_or hp.isUFormula hq.isUFormula, + ihp hw hv, ihq hw hv] · intro l p hp ih m w n v hw hv - simp only [substs_all, hp] - rw [substs_all (hp.substs hv.qVec), ih hw.qVec hv.qVec, termSubstVec_qVec_qVec hv hw] + rw [substs_all hp.isUFormula, + substs_all (hp.substs hv.qVec).isUFormula, + substs_all hp.isUFormula, + ih hw.qVec hv.qVec, + termSubstVec_qVec_qVec hv hw] · intro l p hp ih m w n v hw hv - simp only [substs_ex, hp] - rw [substs_ex (hp.substs hv.qVec), ih hw.qVec hv.qVec, termSubstVec_qVec_qVec hv hw] - -lemma subst_eq_self {n w : V} (hp : L.Semiformula n p) (hw : L.SemitermVec n n w) (H : ∀ i < n, w.[i] = ^#i) : - L.substs n w p = p := by + rw [substs_ex hp.isUFormula, + substs_ex (hp.substs hv.qVec).isUFormula, + substs_ex hp.isUFormula, + ih hw.qVec hv.qVec, + termSubstVec_qVec_qVec hv hw] + +lemma subst_eq_self {n w : V} (hp : L.IsSemiformula n p) (hw : L.IsSemitermVec n n w) (H : ∀ i < n, w.[i] = ^#i) : + L.substs w p = p := by revert w - apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.IsSemiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp · definability · intro n k R v hR hv w _ H - simp only [substs_rel, qqRel_inj, true_and, hR, hv] - apply nth_ext' k (by simp [*]) (by simp [hv.1]) + simp only [substs_rel, qqRel_inj, true_and, hR, hv.isUTerm] + apply nth_ext' k (by simp [*, hv.isUTerm]) (by simp [hv.lh]) intro i hi - rw [nth_termSubstVec hv hi, termSubst_eq_self (hv.2 i hi) H] + rw [nth_termSubstVec hv.isUTerm hi, termSubst_eq_self (hv.nth hi) H] · intro n k R v hR hv w _ H - simp only [substs_nrel, qqNRel_inj, true_and, hR, hv] - apply nth_ext' k (by simp [*]) (by simp [hv.1]) + simp only [substs_nrel, qqNRel_inj, true_and, hR, hv.isUTerm] + apply nth_ext' k (by simp [*, hv.isUTerm]) (by simp [hv.lh]) intro i hi - rw [nth_termSubstVec hv hi, termSubst_eq_self (hv.2 i hi) H] + rw [nth_termSubstVec hv.isUTerm hi, termSubst_eq_self (hv.nth hi) H] · intro n w _ _; simp · intro n w _ _; simp · intro n p q hp hq ihp ihq w hw H - simp [*, ihp hw H, ihq hw H] + simp [*, hp.isUFormula, hq.isUFormula, ihp hw H, ihq hw H] · intro n p q hp hq ihp ihq w hw H - simp [*, ihp hw H, ihq hw H] + simp [*, hp.isUFormula, hq.isUFormula, ihp hw H, ihq hw H] · intro n p hp ih w hw H - have H : ∀ i < n + 1, (L.qVec n n w).[i] = ^#i := by + have H : ∀ i < n + 1, (L.qVec w).[i] = ^#i := by intro i hi rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp [Language.qVec] · have hi : i < n := by simpa using hi simp only [Language.qVec, nth_cons_succ] - rw [nth_termBShiftVec hw hi] - simp [H i hi, hi] - simp [*, ih hw.qVec H] + rw [nth_termBShiftVec (by simpa [hw.lh] using hw.isUTerm) (by simp [hw.lh, hi])] + simp [hw.lh, H i hi, hi] + simp [*, hp.isUFormula, ih hw.qVec H] · intro n p hp ih w hw H - have H : ∀ i < n + 1, (L.qVec n n w).[i] = ^#i := by + have H : ∀ i < n + 1, (L.qVec w).[i] = ^#i := by intro i hi rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp [Language.qVec] · have hi : i < n := by simpa using hi simp only [Language.qVec, nth_cons_succ] - rw [nth_termBShiftVec hw hi] + rw [nth_termBShiftVec (by simpa [hw.lh] using hw.isUTerm) (by simp [hw.lh, hi])] simp [H i hi, hi] - simp [*, ih hw.qVec H] + simp [*, hp.isUFormula, ih hw.qVec H] end substs - variable (L) -def Language.substs₁ (t u : V) : V := L.substs 0 ?[t] u +def Language.substs₁ (t u : V) : V := L.substs ?[t] u variable {L} @@ -701,19 +745,21 @@ section substs₁ section def _root_.LO.FirstOrder.Arith.LDef.substs₁Def (pL : LDef) : 𝚺₁.Semisentence 3 := .mkSigma - “ z t p | ∃ v, !consDef v t 0 ∧ !pL.substsDef z 0 v p” (by simp) + “ z t p | ∃ v, !consDef v t 0 ∧ !pL.substsDef z v p” (by simp) variable (L) -lemma substs₁_defined : 𝚺₁-Function₂ L.substs₁ via pL.substs₁Def := by - intro v; simp [LDef.substs₁Def, (substs_defined L).df.iff]; rfl +lemma Language.substs₁_defined : 𝚺₁-Function₂ L.substs₁ via pL.substs₁Def := by + intro v; simp [LDef.substs₁Def, L.substs_defined.df.iff]; rfl -@[simp] instance substs₁_definable : 𝚺₁-Function₂ L.substs₁ := (substs₁_defined L).to_definable +instance Language.substs₁_definable : 𝚺₁-Function₂ L.substs₁ := L.substs₁_defined.to_definable + +instance : Γ-[m + 1]-Function₂ L.substs₁ := L.substs₁_definable.of_sigmaOne end -lemma Language.Semiformula.substs₁ (ht : L.Term t) (hp : L.Semiformula 1 p) : L.Semiformula 0 (L.substs₁ t p) := - Language.Semiformula.substs hp (by simp [ht]) +lemma Language.IsSemiformula.substs₁ (ht : L.IsSemiterm n t) (hp : L.IsSemiformula 1 p) : L.IsSemiformula n (L.substs₁ t p) := + Language.IsSemiformula.substs hp (by simp [ht]) end substs₁ @@ -732,23 +778,26 @@ def _root_.LO.FirstOrder.Arith.LDef.freeDef (pL : LDef) : 𝚺₁.Semisentence 2 variable (L) -lemma free_defined : 𝚺₁-Function₁ L.free via pL.freeDef := by - intro v; simp [LDef.freeDef, (shift_defined L).df.iff, (substs₁_defined L).df.iff, Language.free] +lemma Language.free_defined : 𝚺₁-Function₁ L.free via pL.freeDef := by + intro v; simp [LDef.freeDef, L.shift_defined.df.iff, L.substs₁_defined.df.iff, Language.free] + +instance Language.free_definable : 𝚺₁-Function₁ L.free := L.free_defined.to_definable -instance free_definable : 𝚺-[0 + 1]-Function₁ L.free := (free_defined L).to_definable +instance Language.free_definable' : Γ-[m + 1]-Function₁ L.free := L.free_definable.of_sigmaOne end -@[simp] lemma Language.Semiformula.free (hp : L.Semiformula 1 p) : L.Formula (L.free p) := - Language.Semiformula.substs₁ (by simp) (by simp [hp]) +@[simp] lemma Language.IsSemiformula.free (hp : L.IsSemiformula 1 p) : L.IsFormula (L.free p) := + Language.IsSemiformula.substs₁ (by simp) hp.shift end free +/- section fvfree variable (L) -def Language.IsFVFree (n p : V) : Prop := L.Semiformula n p ∧ L.shift p = p +def Language.IsFVFree (n p : V) : Prop := L.IsSemiformula n p ∧ L.shift p = p section @@ -782,79 +831,72 @@ lemma Language.IsFVFree.ex {n p : V} (hp : L.IsFVFree (n + 1) p) : @[simp] lemma Language.IsFVFree.neg_iff : L.IsFVFree n (L.neg p) ↔ L.IsFVFree n p := by constructor · intro h - have hp : L.Semiformula n p := Language.Semiformula.neg_iff.mp h.1 + have hp : L.Semiformula n p := Language.IsSemiformula.neg_iff.mp h.1 have : L.shift (L.neg p) = L.neg p := h.2 simp [shift_neg hp, neg_inj_iff hp.shift hp] at this exact ⟨hp, this⟩ · intro h; exact ⟨by simp [h.1], by rw [shift_neg h.1, h.2]⟩ end fvfree +-/ namespace Formalized -def qqEQ (n x y : V) : V := ^rel n 2 (eqIndex : V) ?[x, y] - -def qqNEQ (n x y : V) : V := ^nrel n 2 (eqIndex : V) ?[x, y] - -def qqLT (n x y : V) : V := ^rel n 2 (ltIndex : V) ?[x, y] - -def qqNLT (n x y : V) : V := ^nrel n 2 (ltIndex : V) ?[x, y] - -notation:75 x:75 " ^=[" n "] " y:76 => qqEQ n x y +def qqEQ (x y : V) : V := ^rel 2 (eqIndex : V) ?[x, y] -notation:75 x:75 " ^= " y:76 => qqEQ 0 x y +def qqNEQ (x y : V) : V := ^nrel 2 (eqIndex : V) ?[x, y] -notation:75 x:75 " ^≠[" n "] " y:76 => qqNEQ n x y +def qqLT (x y : V) : V := ^rel 2 (ltIndex : V) ?[x, y] -notation:75 x:75 " ^≠ " y:76 => qqNEQ 0 x y +def qqNLT (x y : V) : V := ^nrel 2 (ltIndex : V) ?[x, y] -notation:78 x:78 " ^<[" n "] " y:79 => qqLT n x y +notation:75 x:75 " ^= " y:76 => qqEQ x y -notation:78 x:78 " ^< " y:79 => qqLT 0 x y +notation:75 x:75 " ^≠ " y:76 => qqNEQ x y -notation:78 x:78 " ^≮[" n "] " y:79 => qqNLT n x y +notation:78 x:78 " ^< " y:79 => qqLT x y -notation:78 x:78 " ^≮ " y:79 => qqNLT 0 x y +notation:78 x:78 " ^≮ " y:79 => qqNLT x y -def _root_.LO.FirstOrder.Arith.qqEQDef : 𝚺₁.Semisentence 4 := - .mkSigma “p n x y | ∃ v, !mkVec₂Def v x y ∧ !qqRelDef p n 2 (!(.Operator.numeral ℒₒᵣ eqIndex)) v” (by simp) +def _root_.LO.FirstOrder.Arith.qqEQDef : 𝚺₁.Semisentence 3 := + .mkSigma “p x y | ∃ v, !mkVec₂Def v x y ∧ !qqRelDef p 2 ↑eqIndex v” (by simp) -def _root_.LO.FirstOrder.Arith.qqNEQDef : 𝚺₁.Semisentence 4 := - .mkSigma “p n x y | ∃ v, !mkVec₂Def v x y ∧ !qqNRelDef p n 2 (!(.Operator.numeral ℒₒᵣ eqIndex)) v” (by simp) +def _root_.LO.FirstOrder.Arith.qqNEQDef : 𝚺₁.Semisentence 3 := + .mkSigma “p x y | ∃ v, !mkVec₂Def v x y ∧ !qqNRelDef p 2 ↑eqIndex v” (by simp) -def _root_.LO.FirstOrder.Arith.qqLTDef : 𝚺₁.Semisentence 4 := - .mkSigma “p n x y | ∃ v, !mkVec₂Def v x y ∧ !qqRelDef p n 2 (!(.Operator.numeral ℒₒᵣ ltIndex)) v” (by simp) +def _root_.LO.FirstOrder.Arith.qqLTDef : 𝚺₁.Semisentence 3 := + .mkSigma “p x y | ∃ v, !mkVec₂Def v x y ∧ !qqRelDef p 2 ↑ltIndex v” (by simp) -def _root_.LO.FirstOrder.Arith.qqNLTDef : 𝚺₁.Semisentence 4 := - .mkSigma “p n x y | ∃ v, !mkVec₂Def v x y ∧ !qqNRelDef p n 2 (!(.Operator.numeral ℒₒᵣ ltIndex)) v” (by simp) +def _root_.LO.FirstOrder.Arith.qqNLTDef : 𝚺₁.Semisentence 3 := + .mkSigma “p x y | ∃ v, !mkVec₂Def v x y ∧ !qqNRelDef p 2 ↑ltIndex v” (by simp) -lemma qqEQ_defined : 𝚺₁-Function₃ (qqEQ : V → V → V → V) via qqEQDef := by +lemma qqEQ_defined : 𝚺₁-Function₂ (qqEQ : V → V → V) via qqEQDef := by intro v; simp [qqEQDef, numeral_eq_natCast, qqEQ] -lemma qqNEQ_defined : 𝚺-[0 + 1]-Function₃ (qqNEQ : V → V → V → V) via qqNEQDef := by +lemma qqNEQ_defined : 𝚺₁-Function₂ (qqNEQ : V → V → V) via qqNEQDef := by intro v; simp [qqNEQDef, numeral_eq_natCast, qqNEQ] -lemma qqLT_defined : 𝚺-[0 + 1]-Function₃ (qqLT : V → V → V → V) via qqLTDef := by +lemma qqLT_defined : 𝚺₁-Function₂ (qqLT : V → V → V) via qqLTDef := by intro v; simp [qqLTDef, numeral_eq_natCast, qqLT] -lemma qqNLT_defined : 𝚺-[0 + 1]-Function₃ (qqNLT : V → V → V → V) via qqNLTDef := by +lemma qqNLT_defined : 𝚺₁-Function₂ (qqNLT : V → V → V) via qqNLTDef := by intro v; simp [qqNLTDef, numeral_eq_natCast, qqNLT] -instance (Γ m) : Γ-[m + 1]-Function₃ (qqEQ : V → V → V → V) := .of_sigmaOne qqEQ_defined.to_definable _ _ +instance (Γ m) : Γ-[m + 1]-Function₂ (qqEQ : V → V → V) := .of_sigmaOne qqEQ_defined.to_definable -instance (Γ m) : Γ-[m + 1]-Function₃ (qqNEQ : V → V → V → V) := .of_sigmaOne qqNEQ_defined.to_definable _ _ +instance (Γ m) : Γ-[m + 1]-Function₂ (qqNEQ : V → V → V) := .of_sigmaOne qqNEQ_defined.to_definable -instance (Γ m) : Γ-[m + 1]-Function₃ (qqLT : V → V → V → V) := .of_sigmaOne qqLT_defined.to_definable _ _ +instance (Γ m) : Γ-[m + 1]-Function₂ (qqLT : V → V → V) := .of_sigmaOne qqLT_defined.to_definable -instance (Γ m) : Γ-[m + 1]-Function₃ (qqNLT : V → V → V → V) := .of_sigmaOne qqNLT_defined.to_definable _ _ +instance (Γ m) : Γ-[m + 1]-Function₂ (qqNLT : V → V → V) := .of_sigmaOne qqNLT_defined.to_definable -@[simp] lemma eval_qqEQDef (v) : Semiformula.Evalbm V v qqEQDef.val ↔ v 0 = v 2 ^=[v 1] v 3 := qqEQ_defined.df.iff v +@[simp] lemma eval_qqEQDef (v) : Semiformula.Evalbm V v qqEQDef.val ↔ v 0 = v 1 ^= v 2 := qqEQ_defined.df.iff v -@[simp] lemma eval_qqNEQDef (v) : Semiformula.Evalbm V v qqNEQDef.val ↔ v 0 = v 2 ^≠[v 1] v 3 := qqNEQ_defined.df.iff v +@[simp] lemma eval_qqNEQDef (v) : Semiformula.Evalbm V v qqNEQDef.val ↔ v 0 = v 1 ^≠ v 2 := qqNEQ_defined.df.iff v -@[simp] lemma eval_qqLTDef (v) : Semiformula.Evalbm V v qqLTDef.val ↔ v 0 = v 2 ^<[v 1] v 3 := qqLT_defined.df.iff v +@[simp] lemma eval_qqLTDef (v) : Semiformula.Evalbm V v qqLTDef.val ↔ v 0 = v 1 ^< v 2 := qqLT_defined.df.iff v -@[simp] lemma eval_qqNLTDef (v) : Semiformula.Evalbm V v qqNLTDef.val ↔ v 0 = v 2 ^≮[v 1] v 3 := qqNLT_defined.df.iff v +@[simp] lemma eval_qqNLTDef (v) : Semiformula.Evalbm V v qqNLTDef.val ↔ v 0 = v 1 ^≮ v 2 := qqNLT_defined.df.iff v end Formalized diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean index 7dce5b1..7b00e8d 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean @@ -1,6 +1,5 @@ import Arithmetization.ISigmaOne.Metamath.Formula.Functions - noncomputable section namespace LO.Arith @@ -13,13 +12,13 @@ variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] namespace QQConj -def blueprint : VecRec.Blueprint 1 where - nil := .mkSigma “y n | !qqVerumDef y n” (by simp) - cons := .mkSigma “y p ps ih n | !qqAndDef y n p ih” (by simp) +def blueprint : VecRec.Blueprint 0 where + nil := .mkSigma “y | !qqVerumDef y” (by simp) + cons := .mkSigma “y p ps ih | !qqAndDef y p ih” (by simp) def construction : VecRec.Construction V blueprint where - nil (param) := ^⊤[param 0] - cons (param) p _ ih := p ^⋏[param 0] ih + nil _ := ^⊤ + cons _ p _ ih := p ^⋏ ih nil_defined := by intro v; simp [blueprint] cons_defined := by intro v; simp [blueprint]; rfl @@ -29,36 +28,33 @@ section qqConj open QQConj -def qqConj (n ps : V) : V := construction.result ![n] ps - -scoped notation:65 "^⋀[" n "] " ps:66 => qqConj n ps +def qqConj (ps : V) : V := construction.result ![] ps -scoped notation:65 "^⋀ " ps:66 => qqConj 0 ps +scoped notation:65 "^⋀ " ps:66 => qqConj ps -@[simp] lemma qqConj_nil (n : V) : ^⋀[n] 0 = ^⊤[n] := by simp [qqConj, construction] +@[simp] lemma qqConj_nil : ^⋀ (0 : V) = ^⊤ := by simp [qqConj, construction] -@[simp] lemma qqConj_cons (n p ps : V) : ^⋀[n] (p ∷ ps) = p ^⋏[n] (^⋀[n] ps) := by simp [qqConj, construction] +@[simp] lemma qqConj_cons (p ps : V) : ^⋀ (p ∷ ps) = p ^⋏ (^⋀ ps) := by simp [qqConj, construction] section -def _root_.LO.FirstOrder.Arith.qqConjDef : 𝚺₁.Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.qqConjDef : 𝚺₁.Semisentence 2 := blueprint.resultDef -lemma qqConj_defined : 𝚺₁-Function₂ (qqConj : V → V → V) via qqConjDef := by - intro v; simpa [qqConjDef] using construction.result_defined ![v 0, v 2, v 1] +lemma qqConj_defined : 𝚺₁-Function₁ (qqConj : V → V) via qqConjDef := construction.result_defined @[simp] lemma eval_qqConj (v) : - Semiformula.Evalbm V v qqConjDef.val ↔ v 0 = qqConj (v 1) (v 2) := qqConj_defined.df.iff v + Semiformula.Evalbm V v qqConjDef.val ↔ v 0 = qqConj (v 1) := qqConj_defined.df.iff v -instance qqConj_definable : 𝚺₁-Function₂ (qqConj : V → V → V) := qqConj_defined.to_definable +instance qqConj_definable : 𝚺₁-Function₁ (qqConj : V → V) := qqConj_defined.to_definable -instance qqConj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqConj : V → V → V) := .of_sigmaOne qqConj_definable _ _ +instance qqConj_definable' : Γ-[m + 1]-Function₁ (qqConj : V → V) := .of_sigmaOne qqConj_definable end @[simp] lemma qqConj_semiformula {n ps : V} : - L.Semiformula n (^⋀[n] ps) ↔ (∀ i < len ps, L.Semiformula n ps.[i]) := by - induction ps using cons_induction_sigma₁ + L.IsSemiformula n (^⋀ ps) ↔ (∀ i < len ps, L.IsSemiformula n ps.[i]) := by + induction ps using cons_induction_sigma1 · definability case nil => simp case cons p ps ih => @@ -73,8 +69,8 @@ lemma qqConj_semiformula {n ps : V} : by simpa using h 0 (by simp), fun i hi ↦ by simpa using h (i + 1) (by simpa using hi)⟩ -@[simp] lemma len_le_conj (n ps : V) : len ps ≤ ^⋀[n] ps := by - induction ps using cons_induction_sigma₁ +@[simp] lemma len_le_conj (ps : V) : len ps ≤ ^⋀ ps := by + induction ps using cons_induction_sigma1 · definability case nil => simp [qqVerum] case cons p ps ih => @@ -85,13 +81,13 @@ end qqConj namespace QQDisj -def blueprint : VecRec.Blueprint 1 where - nil := .mkSigma “y n | !qqFalsumDef y n” (by simp) - cons := .mkSigma “y p ps ih n | !qqOrDef y n p ih” (by simp) +def blueprint : VecRec.Blueprint 0 where + nil := .mkSigma “y | !qqFalsumDef y” (by simp) + cons := .mkSigma “y p ps ih | !qqOrDef y p ih” (by simp) def construction : VecRec.Construction V blueprint where - nil (param) := ^⊥[param 0] - cons (param) p _ ih := p ^⋎[param 0] ih + nil _ := ^⊥ + cons _ p _ ih := p ^⋎ ih nil_defined := by intro v; simp [blueprint] cons_defined := by intro v; simp [blueprint]; rfl @@ -101,36 +97,33 @@ section qqDisj open QQDisj -def qqDisj (n ps : V) : V := construction.result ![n] ps +def qqDisj (ps : V) : V := construction.result ![] ps -scoped notation:65 "^⋁[" n "] " ps:66 => qqDisj n ps +scoped notation:65 "^⋁ " ps:66 => qqDisj ps -scoped notation:65 "^⋁ " ps:66 => qqDisj 0 ps +@[simp] lemma qqDisj_nil : ^⋁ (0 : V) = ^⊥ := by simp [qqDisj, construction] -@[simp] lemma qqDisj_nil (n : V) : ^⋁[n] 0 = ^⊥[n] := by simp [qqDisj, construction] - -@[simp] lemma qqDisj_cons (n p ps : V) : ^⋁[n] (p ∷ ps) = p ^⋎[n] (^⋁[n] ps) := by simp [qqDisj, construction] +@[simp] lemma qqDisj_cons (p ps : V) : ^⋁ (p ∷ ps) = p ^⋎ (^⋁ ps) := by simp [qqDisj, construction] section -def _root_.LO.FirstOrder.Arith.qqDisjDef : 𝚺₁.Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.qqDisjDef : 𝚺₁.Semisentence 2 := blueprint.resultDef -lemma qqDisj_defined : 𝚺₁-Function₂ (qqDisj : V → V → V) via qqDisjDef := by - intro v; simpa [qqDisjDef] using construction.result_defined ![v 0, v 2, v 1] +lemma qqDisj_defined : 𝚺₁-Function₁ (qqDisj : V → V) via qqDisjDef := construction.result_defined @[simp] lemma eval_qqDisj (v) : - Semiformula.Evalbm V v qqDisjDef.val ↔ v 0 = qqDisj (v 1) (v 2) := qqDisj_defined.df.iff v + Semiformula.Evalbm V v qqDisjDef.val ↔ v 0 = qqDisj (v 1) := qqDisj_defined.df.iff v -instance qqDisj_definable : 𝚺₁-Function₂ (qqDisj : V → V → V) := qqDisj_defined.to_definable +instance qqDisj_definable : 𝚺₁-Function₁ (qqDisj : V → V) := qqDisj_defined.to_definable -instance qqDisj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqDisj : V → V → V) := .of_sigmaOne qqDisj_definable _ _ +instance qqDisj_definable' (Γ) : Γ-[m + 1]-Function₁ (qqDisj : V → V) := .of_sigmaOne qqDisj_definable end @[simp] -lemma qqDisj_semiformula {n ps : V} : - L.Semiformula n (^⋁[n] ps) ↔ (∀ i < len ps, L.Semiformula n ps.[i]) := by - induction ps using cons_induction_sigma₁ +lemma qqDisj_semiformula {ps : V} : + L.IsSemiformula n (^⋁ ps) ↔ (∀ i < len ps, L.IsSemiformula n ps.[i]) := by + induction ps using cons_induction_sigma1 · definability case nil => simp case cons p ps ih => @@ -153,58 +146,51 @@ section substItr namespace SubstItr -def blueprint : PR.Blueprint 3 where - zero := .mkSigma “y n w p | y = 0” (by simp) - succ := .mkSigma “y ih k n w p | ∃ numeral, !numeralDef numeral k ∧ ∃ v, !consDef v numeral w ∧ - ∃ sp, !(Language.lDef ℒₒᵣ).substsDef sp n v p ∧ !consDef y sp ih” (by simp) +def blueprint : PR.Blueprint 2 where + zero := .mkSigma “y w p | y = 0” (by simp) + succ := .mkSigma “y ih k w p | ∃ numeral, !numeralDef numeral k ∧ ∃ v, !consDef v numeral w ∧ + ∃ sp, !(Language.lDef ℒₒᵣ).substsDef sp v p ∧ !consDef y sp ih” (by simp) def construction : PR.Construction V blueprint where zero _ := 0 - succ param k ih := (⌜ℒₒᵣ⌝.substs (param 0) (numeral k ∷ param 1) (param 2)) ∷ ih + succ param k ih := (⌜ℒₒᵣ⌝.substs (numeral k ∷ param 0) (param 1)) ∷ ih zero_defined := by intro v; simp [blueprint] - succ_defined := by intro v; simp [blueprint, (substs_defined ⌜ℒₒᵣ⌝).df.iff]; rfl + succ_defined := by intro v; simp [blueprint, ⌜ℒₒᵣ⌝.substs_defined.df.iff]; rfl end SubstItr open SubstItr -def substItr (n w p k : V) : V := construction.result ![n, w, p] k +def substItr (w p k : V) : V := construction.result ![w, p] k -@[simp] lemma substItr_zero (n w p : V) : substItr n w p 0 = 0 := by simp [substItr, construction] +@[simp] lemma substItr_zero (w p : V) : substItr w p 0 = 0 := by simp [substItr, construction] -@[simp] lemma substItr_succ (n w p k : V) : substItr n w p (k + 1) = ⌜ℒₒᵣ⌝.substs n (numeral k ∷ w) p ∷ substItr n w p k := by simp [substItr, construction] +@[simp] lemma substItr_succ (w p k : V) : substItr w p (k + 1) = ⌜ℒₒᵣ⌝.substs (numeral k ∷ w) p ∷ substItr w p k := by simp [substItr, construction] section -def _root_.LO.FirstOrder.Arith.substItrDef : 𝚺₁.Semisentence 5 := blueprint.resultDef |>.rew (Rew.substs ![#0, #4, #1, #2, #3]) +def _root_.LO.FirstOrder.Arith.substItrDef : 𝚺₁.Semisentence 4 := blueprint.resultDef |>.rew (Rew.substs ![#0, #3, #1, #2]) -lemma substItr_defined : 𝚺₁-Function₄ (substItr : V → V → V → V → V) via substItrDef := - fun v ↦ by simp [construction.result_defined_iff, substItrDef]; rfl +lemma substItr_defined : 𝚺₁-Function₃ (substItr : V → V → V → V) via substItrDef := + fun v ↦ by simp [construction.result_defined_iff, substItrDef, substItr]; rfl @[simp] lemma substItr_defined_iff (v) : - Semiformula.Evalbm V v substItrDef.val ↔ v 0 = substItr (v 1) (v 2) (v 3) (v 4) := substItr_defined.df.iff v - -instance substItr_definable : 𝚺₁-Function₄ (substItr : V → V → V → V → V) := substItr_defined.to_definable + Semiformula.Evalbm V v substItrDef.val ↔ v 0 = substItr (v 1) (v 2) (v 3) := substItr_defined.df.iff v -@[simp, definability] instance substItr_definable' (Γ m) : Γ-[m + 1]-Function₄ (substItr : V → V → V → V → V) := - .of_sigmaOne substItr_definable _ _ +instance substItr_definable : 𝚺₁-Function₃ (substItr : V → V → V → V) := substItr_defined.to_definable -instance substItr_definable₁ (n w p : V) : 𝚺₁-Function₁ (substItr n w p) := by - simpa using substItr_definable.retractiont ![&n, &w, &p, #0] - -instance substItr_definable₁' (n w p : V) (Γ m) : Γ-[m + 1]-Function₁ (substItr n w p) := - .of_sigmaOne (substItr_definable₁ n w p) _ _ +instance substItr_definable' : Γ-[m + 1]-Function₃ (substItr : V → V → V → V) := .of_sigmaOne substItr_definable end -@[simp] lemma len_substItr (n w p k : V) : len (substItr n w p k) = k := by +@[simp] lemma len_substItr (w p k : V) : len (substItr w p k) = k := by induction k using induction_sigma1 · definability case zero => simp case succ k ih => simp [ih] -@[simp] lemma substItr_nth (n w p k : V) {i} (hi : i < k) : - (substItr n w p k).[i] = ⌜ℒₒᵣ⌝.substs n (numeral (k - (i + 1)) ∷ w) p := by +@[simp] lemma substItr_nth (w p k : V) {i} (hi : i < k) : + (substItr w p k).[i] = ⌜ℒₒᵣ⌝.substs (numeral (k - (i + 1)) ∷ w) p := by induction k using induction_sigma1 generalizing i · definability case zero => simp at hi @@ -214,67 +200,69 @@ end · simp · simp [ih (by simpa using hi)] -lemma neg_conj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) : - ⌜ℒₒᵣ⌝.neg (^⋀[m] (substItr m w p k)) = ^⋁[m] (substItr m w (⌜ℒₒᵣ⌝.neg p) k) := by +lemma neg_conj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.IsSemiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.IsSemitermVec n m w) : + ⌜ℒₒᵣ⌝.neg (^⋀ (substItr w p k)) = ^⋁ (substItr w (⌜ℒₒᵣ⌝.neg p) k) := by induction k using induction_sigma1 · definability case zero => simp case succ k ih => simp [hw] - rw [neg_and, ←substs_neg hp, ih] + rw [neg_and, ←substs_neg hp (m := m), ih] · simp [hw] - · apply Language.Semiformula.substs hp (by simp [hw]) - · simp only [qqConj_semiformula, len_substItr] + · exact Language.IsSemiformula.isUFormula <| hp.substs (by simpa [hw]) + · apply Language.IsSemiformula.isUFormula + simp only [qqConj_semiformula, len_substItr] intro i hi simp only [gt_iff_lt, hi, substItr_nth] - apply Language.Semiformula.substs hp (by simp [hw]) + apply hp.substs (by simpa [hw]) -lemma neg_disj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) : - ⌜ℒₒᵣ⌝.neg (^⋁[m] (substItr m w p k)) = ^⋀[m] (substItr m w (⌜ℒₒᵣ⌝.neg p) k) := by +lemma neg_disj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.IsSemiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.IsSemitermVec n m w) : + ⌜ℒₒᵣ⌝.neg (^⋁ (substItr w p k)) = ^⋀ (substItr w (⌜ℒₒᵣ⌝.neg p) k) := by induction k using induction_sigma1 · definability case zero => simp case succ k ih => simp [hw] - rw [neg_or, ←substs_neg hp, ih] + rw [neg_or, ←substs_neg hp (m := m), ih] · simp [hw] - · apply Language.Semiformula.substs hp (by simp [hw]) - · simp only [qqDisj_semiformula, len_substItr] + · apply Language.IsSemiformula.isUFormula <| hp.substs (by simpa [hw]) + · apply Language.IsSemiformula.isUFormula + simp only [qqDisj_semiformula, len_substItr] intro i hi simp only [gt_iff_lt, hi, substItr_nth] - apply Language.Semiformula.substs hp (by simp [hw]) + apply hp.substs (by simpa [hw]) -lemma substs_conj_substItr {n m l w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) (hv : ⌜ℒₒᵣ⌝.SemitermVec m l v) : - ⌜ℒₒᵣ⌝.substs l v (^⋀[m] (substItr m w p k)) = ^⋀[l] (substItr l (⌜ℒₒᵣ⌝.termSubstVec n m l v w) p k) := by +lemma substs_conj_substItr {n m l w p k : V} (hp : ⌜ℒₒᵣ⌝.IsSemiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.IsSemitermVec n m w) (hv : ⌜ℒₒᵣ⌝.IsSemitermVec m l v) : + ⌜ℒₒᵣ⌝.substs v (^⋀ (substItr w p k)) = ^⋀ (substItr (⌜ℒₒᵣ⌝.termSubstVec n v w) p k) := by induction k using induction_sigma1 · definability case zero => simp case succ k ih => - have hkw : ⌜ℒₒᵣ⌝.SemitermVec (n + 1) m (numeral k ∷ w) := by simp [hw] - have ha : ⌜ℒₒᵣ⌝.Semiformula m (^⋀[m] substItr m w p k) := by + have hkw : ⌜ℒₒᵣ⌝.IsSemitermVec (n + 1) m (numeral k ∷ w) := by simp [hw] + have ha : ⌜ℒₒᵣ⌝.IsSemiformula m (^⋀ substItr w p k) := by simp only [qqConj_semiformula, len_substItr] intro i hi; simpa [hi] using hp.substs (hw.cons (by simp)) simp only [substItr_succ, qqConj_cons] - rw [substs_and (hp.substs hkw) ha, + rw [substs_and (hp.substs hkw).isUFormula ha.isUFormula, substs_substs hp hv hkw, - termSubstVec_cons (by simp) hw, + termSubstVec_cons (by simp) hw.isUTerm, numeral_substs hv] simp [ih] -lemma substs_disj_substItr {n m l w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) (hv : ⌜ℒₒᵣ⌝.SemitermVec m l v) : - ⌜ℒₒᵣ⌝.substs l v (^⋁[m] (substItr m w p k)) = ^⋁[l] (substItr l (⌜ℒₒᵣ⌝.termSubstVec n m l v w) p k) := by +lemma substs_disj_substItr {n m l w p k : V} (hp : ⌜ℒₒᵣ⌝.IsSemiformula (n + 1) p) (hw : ⌜ℒₒᵣ⌝.IsSemitermVec n m w) (hv : ⌜ℒₒᵣ⌝.IsSemitermVec m l v) : + ⌜ℒₒᵣ⌝.substs v (^⋁ (substItr w p k)) = ^⋁ (substItr (⌜ℒₒᵣ⌝.termSubstVec n v w) p k) := by induction k using induction_sigma1 · definability case zero => simp case succ k ih => - have hkw : ⌜ℒₒᵣ⌝.SemitermVec (n + 1) m (numeral k ∷ w) := by simp [hw] - have ha : ⌜ℒₒᵣ⌝.Semiformula m (^⋁[m] substItr m w p k) := by + have hkw : ⌜ℒₒᵣ⌝.IsSemitermVec (n + 1) m (numeral k ∷ w) := by simp [hw] + have ha : ⌜ℒₒᵣ⌝.IsSemiformula m (^⋁ substItr w p k) := by simp only [qqDisj_semiformula, len_substItr] intro i hi; simpa [hi] using hp.substs (hw.cons (by simp)) simp only [substItr_succ, qqDisj_cons] - rw [substs_or (hp.substs hkw) ha, + rw [substs_or (hp.substs hkw).isUFormula ha.isUFormula, substs_substs hp hv hkw, - termSubstVec_cons (by simp) hw, + termSubstVec_cons (by simp) hw.isUTerm, numeral_substs hv] simp [ih] @@ -284,26 +272,25 @@ end Formalized section verums -def qqVerums (n k : V) : V := ^⋀[n] repeatVec (^⊤[n]) k +def qqVerums (k : V) : V := ^⋀ repeatVec ^⊤ k -@[simp] lemma le_qqVerums (n k : V) : k ≤ qqVerums n k := by - simpa [qqVerums] using len_le_conj n (repeatVec ^⊤[n] k) +@[simp] lemma le_qqVerums (k : V) : k ≤ qqVerums k := by + simpa [qqVerums] using len_le_conj (repeatVec ^⊤ k) section -def _root_.LO.FirstOrder.Arith.qqVerumsDef : 𝚺₁.Semisentence 3 := .mkSigma - “y n k | ∃ verum, !qqVerumDef verum n ∧ ∃ vs, !repeatVecDef vs verum k ∧ !qqConjDef y n vs” (by simp) +def _root_.LO.FirstOrder.Arith.qqVerumsDef : 𝚺₁.Semisentence 2 := .mkSigma + “y k | ∃ verum, !qqVerumDef verum ∧ ∃ vs, !repeatVecDef vs verum k ∧ !qqConjDef y vs” (by simp) -lemma qqVerums_defined : 𝚺₁-Function₂ (qqVerums : V → V → V) via qqVerumsDef := +lemma qqVerums_defined : 𝚺₁-Function₁ (qqVerums : V → V) via qqVerumsDef := fun v ↦ by simp [qqVerumsDef]; rfl @[simp] lemma qqVerums_repeatVec (v) : - Semiformula.Evalbm V v qqVerumsDef.val ↔ v 0 = qqVerums (v 1) (v 2) := qqVerums_defined.df.iff v + Semiformula.Evalbm V v qqVerumsDef.val ↔ v 0 = qqVerums (v 1) := qqVerums_defined.df.iff v -instance qqVerums_definable : 𝚺₁-Function₂ (qqVerums : V → V → V) := qqVerums_defined.to_definable +instance qqVerums_definable : 𝚺₁-Function₁ (qqVerums : V → V) := qqVerums_defined.to_definable -instance qqVerums_definable' (Γ) : Γ-[m + 1]-Function₂ (qqVerums : V → V → V) := - .of_sigmaOne qqVerums_definable _ _ +instance qqVerums_definable' : Γ-[m + 1]-Function₁ (qqVerums : V → V) := .of_sigmaOne qqVerums_definable end diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean index 24cdb5f..7511964 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean @@ -31,408 +31,414 @@ section typed_formula variable (L) -structure Language.TSemiformula (n : V) where +structure Language.Semiformula (n : V) where val : V - prop : L.Semiformula n val + prop : L.IsSemiformula n val -attribute [simp] Language.TSemiformula.prop +attribute [simp] Language.Semiformula.prop -abbrev Language.TFormula := L.TSemiformula 0 +abbrev Language.Formula := L.Semiformula 0 variable {L} -scoped instance : LogicalConnective (L.TSemiformula n) where - top := ⟨^⊤[n], by simp⟩ - bot := ⟨^⊥[n], by simp⟩ - wedge (p q) := ⟨p.val ^⋏[n] q.val, by simp⟩ - vee (p q) := ⟨p.val ^⋎[n] q.val, by simp⟩ +@[simp] lemma Language.Semiformula.isUFormula (p : L.Semiformula n) : L.IsUFormula p.val := p.prop.isUFormula + +scoped instance : LogicalConnective (L.Semiformula n) where + top := ⟨^⊤, by simp⟩ + bot := ⟨^⊥, by simp⟩ + wedge (p q) := ⟨p.val ^⋏ q.val, by simp⟩ + vee (p q) := ⟨p.val ^⋎ q.val, by simp⟩ tilde (p) := ⟨L.neg p.val, by simp⟩ - arrow (p q) := ⟨L.imp n p.val q.val, by simp⟩ + arrow (p q) := ⟨L.imp p.val q.val, by simp⟩ -def Language.TSemiformula.cast (p : L.TSemiformula n) (eq : n = n' := by simp) : L.TSemiformula n' := eq ▸ p +def Language.Semiformula.cast (p : L.Semiformula n) (eq : n = n' := by simp) : L.Semiformula n' := eq ▸ p -@[simp] lemma Language.TSemiformula.val_cast (p : L.TSemiformula n) (eq : n = n') : - (p.cast eq).val = p.val := by rcases eq; simp [Language.TSemiformula.cast] +@[simp] lemma Language.Semiformula.val_cast (p : L.Semiformula n) (eq : n = n') : + (p.cast eq).val = p.val := by rcases eq; simp [Language.Semiformula.cast] -def Language.TSemiformula.all (p : L.TSemiformula (n + 1)) : L.TSemiformula n := ⟨^∀[n] p.val, by simp⟩ +def Language.Semiformula.all (p : L.Semiformula (n + 1)) : L.Semiformula n := ⟨^∀ p.val, by simp⟩ -def Language.TSemiformula.ex (p : L.TSemiformula (n + 1)) : L.TSemiformula n := ⟨^∃[n] p.val, by simp⟩ +def Language.Semiformula.ex (p : L.Semiformula (n + 1)) : L.Semiformula n := ⟨^∃ p.val, by simp⟩ -namespace Language.TSemiformula +namespace Language.Semiformula -@[simp] lemma val_verum : (⊤ : L.TSemiformula n).val = ^⊤[n] := rfl +@[simp] lemma val_verum : (⊤ : L.Semiformula n).val = ^⊤ := rfl -@[simp] lemma val_falsum : (⊥ : L.TSemiformula n).val = ^⊥[n] := rfl +@[simp] lemma val_falsum : (⊥ : L.Semiformula n).val = ^⊥ := rfl -@[simp] lemma val_and (p q : L.TSemiformula n) : - (p ⋏ q).val = p.val ^⋏[n] q.val := rfl +@[simp] lemma val_and (p q : L.Semiformula n) : + (p ⋏ q).val = p.val ^⋏ q.val := rfl -@[simp] lemma val_or (p q : L.TSemiformula n) : - (p ⋎ q).val = p.val ^⋎[n] q.val := rfl +@[simp] lemma val_or (p q : L.Semiformula n) : + (p ⋎ q).val = p.val ^⋎ q.val := rfl -@[simp] lemma val_neg (p : L.TSemiformula n) : +@[simp] lemma val_neg (p : L.Semiformula n) : (~p).val = L.neg p.val := rfl -@[simp] lemma val_imp (p q : L.TSemiformula n) : - (p ⟶ q).val = L.imp n p.val q.val := rfl +@[simp] lemma val_imp (p q : L.Semiformula n) : + (p ⟶ q).val = L.imp p.val q.val := rfl -@[simp] lemma val_all (p : L.TSemiformula (n + 1)) : - p.all.val = ^∀[n] p.val := rfl +@[simp] lemma val_all (p : L.Semiformula (n + 1)) : + p.all.val = ^∀ p.val := rfl -@[simp] lemma val_ex (p : L.TSemiformula (n + 1)) : - p.ex.val = ^∃[n] p.val := rfl +@[simp] lemma val_ex (p : L.Semiformula (n + 1)) : + p.ex.val = ^∃ p.val := rfl -lemma val_inj {p q : L.TSemiformula n} : +lemma val_inj {p q : L.Semiformula n} : p.val = q.val ↔ p = q := by rcases p; rcases q; simp -@[ext] lemma ext {p q : L.TSemiformula n} (h : p.val = q.val) : p = q := val_inj.mp h +@[ext] lemma ext {p q : L.Semiformula n} (h : p.val = q.val) : p = q := val_inj.mp h -lemma ext_iff {p q : L.TSemiformula n} : p = q ↔ p.val = q.val := by rcases p; rcases q; simp +lemma ext_iff {p q : L.Semiformula n} : p = q ↔ p.val = q.val := by rcases p; rcases q; simp -@[simp] lemma and_inj {p₁ p₂ q₁ q₂ : L.TSemiformula n} : +@[simp] lemma and_inj {p₁ p₂ q₁ q₂ : L.Semiformula n} : p₁ ⋏ p₂ = q₁ ⋏ q₂ ↔ p₁ = q₁ ∧ p₂ = q₂ := by simp [ext_iff] -@[simp] lemma or_inj {p₁ p₂ q₁ q₂ : L.TSemiformula n} : +@[simp] lemma or_inj {p₁ p₂ q₁ q₂ : L.Semiformula n} : p₁ ⋎ p₂ = q₁ ⋎ q₂ ↔ p₁ = q₁ ∧ p₂ = q₂ := by simp [ext_iff] -@[simp] lemma all_inj {p q : L.TSemiformula (n + 1)} : +@[simp] lemma all_inj {p q : L.Semiformula (n + 1)} : p.all = q.all ↔ p = q := by simp [ext_iff] -@[simp] lemma ex_inj {p q : L.TSemiformula (n + 1)} : +@[simp] lemma ex_inj {p q : L.Semiformula (n + 1)} : p.ex = q.ex ↔ p = q := by simp [ext_iff] -@[simp] lemma neg_verum : ~(⊤ : L.TSemiformula n) = ⊥ := by ext; simp -@[simp] lemma neg_falsum : ~(⊥ : L.TSemiformula n) = ⊤ := by ext; simp -@[simp] lemma neg_and (p q : L.TSemiformula n) : ~(p ⋏ q) = ~p ⋎ ~q := by ext; simp -@[simp] lemma neg_or (p q : L.TSemiformula n) : ~(p ⋎ q) = ~p ⋏ ~q := by ext; simp -@[simp] lemma neg_all (p : L.TSemiformula (n + 1)) : ~p.all = (~p).ex := by ext; simp -@[simp] lemma neg_ex (p : L.TSemiformula (n + 1)) : ~p.ex = (~p).all := by ext; simp +@[simp] lemma neg_verum : ~(⊤ : L.Semiformula n) = ⊥ := by ext; simp +@[simp] lemma neg_falsum : ~(⊥ : L.Semiformula n) = ⊤ := by ext; simp +@[simp] lemma neg_and (p q : L.Semiformula n) : ~(p ⋏ q) = ~p ⋎ ~q := by ext; simp +@[simp] lemma neg_or (p q : L.Semiformula n) : ~(p ⋎ q) = ~p ⋏ ~q := by ext; simp +@[simp] lemma neg_all (p : L.Semiformula (n + 1)) : ~p.all = (~p).ex := by ext; simp +@[simp] lemma neg_ex (p : L.Semiformula (n + 1)) : ~p.ex = (~p).all := by ext; simp -lemma imp_def (p q : L.TSemiformula n) : p ⟶ q = ~p ⋎ q := by ext; simp [imp] +lemma imp_def (p q : L.Semiformula n) : p ⟶ q = ~p ⋎ q := by ext; simp [imp] -@[simp] lemma neg_neg (p : L.TSemiformula n) : ~~p = p := by - ext; simp [shift, Arith.neg_neg p.prop] +@[simp] lemma neg_neg (p : L.Semiformula n) : ~~p = p := by + ext; simp [shift, Language.IsUFormula.neg_neg] -def shift (p : L.TSemiformula n) : L.TSemiformula n := ⟨L.shift p.val, p.prop.shift⟩ +def shift (p : L.Semiformula n) : L.Semiformula n := ⟨L.shift p.val, p.prop.shift⟩ -def substs (p : L.TSemiformula n) (w : L.TSemitermVec n m) : L.TSemiformula m := - ⟨L.substs m w.val p.val, p.prop.substs w.prop⟩ +def substs (p : L.Semiformula n) (w : L.SemitermVec n m) : L.Semiformula m := + ⟨L.substs w.val p.val, p.prop.substs w.prop⟩ -@[simp] lemma val_shift (p : L.TSemiformula n) : p.shift.val = L.shift p.val := rfl -@[simp] lemma val_substs (p : L.TSemiformula n) (w : L.TSemitermVec n m) : (p.substs w).val = L.substs m w.val p.val := rfl +@[simp] lemma val_shift (p : L.Semiformula n) : p.shift.val = L.shift p.val := rfl +@[simp] lemma val_substs (p : L.Semiformula n) (w : L.SemitermVec n m) : (p.substs w).val = L.substs w.val p.val := rfl -@[simp] lemma shift_verum : (⊤ : L.TSemiformula n).shift = ⊤ := by ext; simp [shift] -@[simp] lemma shift_falsum : (⊥ : L.TSemiformula n).shift = ⊥ := by ext; simp [shift] -@[simp] lemma shift_and (p q : L.TSemiformula n) : (p ⋏ q).shift = p.shift ⋏ q.shift := by ext; simp [shift] -@[simp] lemma shift_or (p q : L.TSemiformula n) : (p ⋎ q).shift = p.shift ⋎ q.shift := by ext; simp [shift] -@[simp] lemma shift_all (p : L.TSemiformula (n + 1)) : p.all.shift = p.shift.all := by ext; simp [shift] -@[simp] lemma shift_ex (p : L.TSemiformula (n + 1)) : p.ex.shift = p.shift.ex := by ext; simp [shift] +@[simp] lemma shift_verum : (⊤ : L.Semiformula n).shift = ⊤ := by ext; simp [shift] +@[simp] lemma shift_falsum : (⊥ : L.Semiformula n).shift = ⊥ := by ext; simp [shift] +@[simp] lemma shift_and (p q : L.Semiformula n) : (p ⋏ q).shift = p.shift ⋏ q.shift := by ext; simp [shift] +@[simp] lemma shift_or (p q : L.Semiformula n) : (p ⋎ q).shift = p.shift ⋎ q.shift := by ext; simp [shift] +@[simp] lemma shift_all (p : L.Semiformula (n + 1)) : p.all.shift = p.shift.all := by ext; simp [shift] +@[simp] lemma shift_ex (p : L.Semiformula (n + 1)) : p.ex.shift = p.shift.ex := by ext; simp [shift] -@[simp] lemma neg_inj {p q : L.TSemiformula n} : +@[simp] lemma neg_inj {p q : L.Semiformula n} : ~p = ~q ↔ p = q := ⟨by intro h; simpa using congr_arg (~·) h, by rintro rfl; rfl⟩ -@[simp] lemma imp_inj {p₁ p₂ q₁ q₂ : L.TSemiformula n} : +@[simp] lemma imp_inj {p₁ p₂ q₁ q₂ : L.Semiformula n} : p₁ ⟶ p₂ = q₁ ⟶ q₂ ↔ p₁ = q₁ ∧ p₂ = q₂ := by simp [imp_def] -@[simp] lemma shift_neg (p : L.TSemiformula n) : (~p).shift = ~(p.shift) := by - ext; simp [shift, val_neg, TSemitermVec.prop] +@[simp] lemma shift_neg (p : L.Semiformula n) : (~p).shift = ~(p.shift) := by + ext; simp [shift, val_neg, SemitermVec.prop] rw [Arith.shift_neg p.prop] -@[simp] lemma shift_imp (p q : L.TSemiformula n) : (p ⟶ q).shift = p.shift ⟶ q.shift := by +@[simp] lemma shift_imp (p q : L.Semiformula n) : (p ⟶ q).shift = p.shift ⟶ q.shift := by simp [imp_def] -@[simp] lemma substs_verum (w : L.TSemitermVec n m) : (⊤ : L.TSemiformula n).substs w = ⊤ := by ext; simp [substs] -@[simp] lemma substs_falsum (w : L.TSemitermVec n m) : (⊥ : L.TSemiformula n).substs w = ⊥ := by ext; simp [substs] -@[simp] lemma substs_and (w : L.TSemitermVec n m) (p q : L.TSemiformula n) : +@[simp] lemma substs_verum (w : L.SemitermVec n m) : (⊤ : L.Semiformula n).substs w = ⊤ := by ext; simp [substs] +@[simp] lemma substs_falsum (w : L.SemitermVec n m) : (⊥ : L.Semiformula n).substs w = ⊥ := by ext; simp [substs] +@[simp] lemma substs_and (w : L.SemitermVec n m) (p q : L.Semiformula n) : (p ⋏ q).substs w = p.substs w ⋏ q.substs w := by ext; simp [substs] -@[simp] lemma substs_or (w : L.TSemitermVec n m) (p q : L.TSemiformula n) : +@[simp] lemma substs_or (w : L.SemitermVec n m) (p q : L.Semiformula n) : (p ⋎ q).substs w = p.substs w ⋎ q.substs w := by ext; simp [substs] -@[simp] lemma substs_all (w : L.TSemitermVec n m) (p : L.TSemiformula (n + 1)) : +@[simp] lemma substs_all (w : L.SemitermVec n m) (p : L.Semiformula (n + 1)) : p.all.substs w = (p.substs w.q).all := by - ext; simp [substs, Language.bvar, Language.qVec, Language.TSemitermVec.bShift, Language.TSemitermVec.q] -@[simp] lemma substs_ex (w : L.TSemitermVec n m) (p : L.TSemiformula (n + 1)) : + ext; simp [substs, Language.bvar, Language.qVec, Language.SemitermVec.bShift, Language.SemitermVec.q, w.prop.lh] +@[simp] lemma substs_ex (w : L.SemitermVec n m) (p : L.Semiformula (n + 1)) : p.ex.substs w = (p.substs w.q).ex := by - ext; simp [substs, Language.bvar, Language.qVec, Language.TSemitermVec.bShift, Language.TSemitermVec.q] + ext; simp [substs, Language.bvar, Language.qVec, Language.SemitermVec.bShift, Language.SemitermVec.q, w.prop.lh] -@[simp] lemma substs_neg (w : L.TSemitermVec n m) (p : L.TSemiformula n) : (~p).substs w = ~(p.substs w) := by - ext; simp only [substs, val_neg, TSemitermVec.prop, Arith.substs_neg p.prop] -@[simp] lemma substs_imp (w : L.TSemitermVec n m) (p q : L.TSemiformula n) : (p ⟶ q).substs w = p.substs w ⟶ q.substs w := by +@[simp] lemma substs_neg (w : L.SemitermVec n m) (p : L.Semiformula n) : (~p).substs w = ~(p.substs w) := by + ext; simp [substs, val_neg, SemitermVec.prop, Arith.substs_neg p.prop w.prop] +@[simp] lemma substs_imp (w : L.SemitermVec n m) (p q : L.Semiformula n) : (p ⟶ q).substs w = p.substs w ⟶ q.substs w := by simp [imp_def] -@[simp] lemma substs_imply (w : L.TSemitermVec n m) (p q : L.TSemiformula n) : (p ⟷ q).substs w = p.substs w ⟷ q.substs w := by +@[simp] lemma substs_imply (w : L.SemitermVec n m) (p q : L.Semiformula n) : (p ⟷ q).substs w = p.substs w ⟷ q.substs w := by simp [LogicalConnective.iff] -end Language.TSemiformula +end Language.Semiformula -notation p:max "^/[" w "]" => Language.TSemiformula.substs p w +notation p:max "^/[" w "]" => Language.Semiformula.substs p w -structure Language.TSemiformulaVec (n : V) where +structure Language.SemiformulaVec (n : V) where val : V - prop : ∀ i < len val, L.Semiformula n val.[i] + prop : ∀ i < len val, L.IsSemiformula n val.[i] -namespace Language.TSemiformulaVec +namespace Language.SemiformulaVec -def conj (ps : L.TSemiformulaVec n) : L.TSemiformula n := ⟨^⋀[n] ps.val, by simpa using ps.prop⟩ +def conj (ps : L.SemiformulaVec n) : L.Semiformula n := ⟨^⋀ ps.val, by simpa using ps.prop⟩ -def disj (ps : L.TSemiformulaVec n) : L.TSemiformula n := ⟨^⋁[n] ps.val, by simpa using ps.prop⟩ +def disj (ps : L.SemiformulaVec n) : L.Semiformula n := ⟨^⋁ ps.val, by simpa using ps.prop⟩ -def nth (ps : L.TSemiformulaVec n) (i : V) (hi : i < len ps.val) : L.TSemiformula n := +def nth (ps : L.SemiformulaVec n) (i : V) (hi : i < len ps.val) : L.Semiformula n := ⟨ps.val.[i], ps.prop i hi⟩ -@[simp] lemma val_conj (ps : L.TSemiformulaVec n) : ps.conj.val = ^⋀[n] ps.val := rfl +@[simp] lemma val_conj (ps : L.SemiformulaVec n) : ps.conj.val = ^⋀ ps.val := rfl -@[simp] lemma val_disj (ps : L.TSemiformulaVec n) : ps.disj.val = ^⋁[n] ps.val := rfl +@[simp] lemma val_disj (ps : L.SemiformulaVec n) : ps.disj.val = ^⋁ ps.val := rfl -@[simp] lemma val_nth (ps : L.TSemiformulaVec n) (i : V) (hi : i < len ps.val) : +@[simp] lemma val_nth (ps : L.SemiformulaVec n) (i : V) (hi : i < len ps.val) : (ps.nth i hi).val = ps.val.[i] := rfl -end Language.TSemiformulaVec +end Language.SemiformulaVec namespace Language.TSemifromula -lemma subst_eq_self {n : V} (w : L.TSemitermVec n n) (p : L.TSemiformula n) (H : ∀ i, (hi : i < n) → w.nth i hi = L.bvar i hi) : +lemma subst_eq_self {n : V} (w : L.SemitermVec n n) (p : L.Semiformula n) (H : ∀ i, (hi : i < n) → w.nth i hi = L.bvar i hi) : p^/[w] = p := by ext; simp; rw [Arith.subst_eq_self p.prop w.prop] intro i hi - simpa using congr_arg Language.TSemiterm.val (H i hi) + simpa using congr_arg Language.Semiterm.val (H i hi) -@[simp] lemma subst_eq_self₁ (p : L.TSemiformula (0 + 1)) : +@[simp] lemma subst_eq_self₁ (p : L.Semiformula (0 + 1)) : p^/[(L.bvar 0 (by simp)).sing] = p := by apply subst_eq_self simp only [zero_add, lt_one_iff_eq_zero] rintro _ rfl simp -@[simp] lemma subst_nil_eq_self (w : L.TSemitermVec 0 0) : +@[simp] lemma subst_nil_eq_self (w : L.SemitermVec 0 0) : p^/[w] = p := subst_eq_self _ _ (by simp) -lemma shift_substs {n m : V} (w : L.TSemitermVec n m) (p : L.TSemiformula n) : +lemma shift_substs {n m : V} (w : L.SemitermVec n m) (p : L.Semiformula n) : (p^/[w]).shift = p.shift^/[w.shift] := by ext; simp; rw [Arith.shift_substs p.prop w.prop] -lemma substs_substs {n m l : V} (v : L.TSemitermVec m l) (w : L.TSemitermVec n m) (p : L.TSemiformula n) : +lemma substs_substs {n m l : V} (v : L.SemitermVec m l) (w : L.SemitermVec n m) (p : L.Semiformula n) : (p^/[w])^/[v] = p^/[w.substs v] := by ext; simp; rw [Arith.substs_substs p.prop v.prop w.prop] end Language.TSemifromula end typed_formula +/- section typed_isfvfree -namespace Language.TSemiformula +namespace Language.Semiformula -def FVFree (p : L.TSemiformula n) : Prop := L.IsFVFree n p.val +def FVFree (p : L.Semiformula n) : Prop := L.IsFVFree n p.val -lemma FVFree.iff {p : L.TSemiformula n} : p.FVFree ↔ p.shift = p := by +lemma FVFree.iff {p : L.Semiformula n} : p.FVFree ↔ p.shift = p := by simp [FVFree, Language.IsFVFree, ext_iff] -@[simp] lemma Fvfree.verum : (⊤ : L.TSemiformula n).FVFree := by simp [FVFree] +@[simp] lemma Fvfree.verum : (⊤ : L.Semiformula n).FVFree := by simp [FVFree] -@[simp] lemma Fvfree.falsum : (⊥ : L.TSemiformula n).FVFree := by simp [FVFree] +@[simp] lemma Fvfree.falsum : (⊥ : L.Semiformula n).FVFree := by simp [FVFree] -@[simp] lemma Fvfree.and {p q : L.TSemiformula n} : +@[simp] lemma Fvfree.and {p q : L.Semiformula n} : (p ⋏ q).FVFree ↔ p.FVFree ∧ q.FVFree := by simp [FVFree.iff, FVFree.iff] -@[simp] lemma Fvfree.or {p q : L.TSemiformula n} : (p ⋎ q).FVFree ↔ p.FVFree ∧ q.FVFree := by +@[simp] lemma Fvfree.or {p q : L.Semiformula n} : (p ⋎ q).FVFree ↔ p.FVFree ∧ q.FVFree := by simp [FVFree.iff] -@[simp] lemma Fvfree.neg {p : L.TSemiformula n} : (~p).FVFree ↔ p.FVFree := by +@[simp] lemma Fvfree.neg {p : L.Semiformula n} : (~p).FVFree ↔ p.FVFree := by simp [FVFree.iff] -@[simp] lemma Fvfree.all {p : L.TSemiformula (n + 1)} : p.all.FVFree ↔ p.FVFree := by +@[simp] lemma Fvfree.all {p : L.Semiformula (n + 1)} : p.all.FVFree ↔ p.FVFree := by simp [FVFree.iff] -@[simp] lemma Fvfree.ex {p : L.TSemiformula (n + 1)} : p.ex.FVFree ↔ p.FVFree := by +@[simp] lemma Fvfree.ex {p : L.Semiformula (n + 1)} : p.ex.FVFree ↔ p.FVFree := by simp [FVFree.iff] -@[simp] lemma Fvfree.imp {p q : L.TSemiformula n} : (p ⟶ q).FVFree ↔ p.FVFree ∧ q.FVFree := by +@[simp] lemma Fvfree.imp {p q : L.Semiformula n} : (p ⟶ q).FVFree ↔ p.FVFree ∧ q.FVFree := by simp [FVFree.iff] -end Language.TSemiformula +end Language.Semiformula end typed_isfvfree +-/ open Formalized -def Language.TSemiterm.equals {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiformula n := ⟨t.val ^=[n] u.val, by simp [qqEQ]⟩ +def Language.Semiterm.equals {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiformula n := ⟨t.val ^= u.val, by simp [qqEQ]⟩ -def Language.TSemiterm.notEquals {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiformula n := ⟨t.val ^≠[n] u.val, by simp [qqNEQ]⟩ +def Language.Semiterm.notEquals {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiformula n := ⟨t.val ^≠ u.val, by simp [qqNEQ]⟩ -def Language.TSemiterm.lessThan {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiformula n := ⟨t.val ^<[n] u.val, by simp [qqLT]⟩ +def Language.Semiterm.lessThan {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiformula n := ⟨t.val ^< u.val, by simp [qqLT]⟩ -def Language.TSemiterm.notLessThan {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiformula n := ⟨t.val ^≮[n] u.val, by simp [qqNLT]⟩ +def Language.Semiterm.notLessThan {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiformula n := ⟨t.val ^≮ u.val, by simp [qqNLT]⟩ -scoped infix:75 " =' " => Language.TSemiterm.equals +scoped infix:75 " =' " => Language.Semiterm.equals -scoped infix:75 " ≠' " => Language.TSemiterm.notEquals +scoped infix:75 " ≠' " => Language.Semiterm.notEquals -scoped infix:75 " <' " => Language.TSemiterm.lessThan +scoped infix:75 " <' " => Language.Semiterm.lessThan -scoped infix:75 " ≮' " => Language.TSemiterm.notLessThan +scoped infix:75 " ≮' " => Language.Semiterm.notLessThan -def Language.TSemiformula.ball {n : V} (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : ⌜ℒₒᵣ⌝.TSemiformula n := +def Language.Semiformula.ball {n : V} (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : ⌜ℒₒᵣ⌝.Semiformula n := (⌜ℒₒᵣ⌝.bvar 0 ≮' t.bShift ⋎ p).all -def Language.TSemiformula.bex {n : V} (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : ⌜ℒₒᵣ⌝.TSemiformula n := +def Language.Semiformula.bex {n : V} (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : ⌜ℒₒᵣ⌝.Semiformula n := (⌜ℒₒᵣ⌝.bvar 0 <' t.bShift ⋏ p).ex namespace Formalized variable {n m : V} -@[simp] lemma val_equals {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : (t =' u).val = t.val ^=[n] u.val := rfl -@[simp] lemma val_notEquals {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : (t ≠' u).val = t.val ^≠[n] u.val := rfl -@[simp] lemma val_lessThan {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : (t <' u).val = t.val ^<[n] u.val := rfl -@[simp] lemma val_notLessThan {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : (t ≮' u).val = t.val ^≮[n] u.val := rfl +@[simp] lemma val_equals {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : (t =' u).val = t.val ^= u.val := rfl +@[simp] lemma val_notEquals {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : (t ≠' u).val = t.val ^≠ u.val := rfl +@[simp] lemma val_lessThan {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : (t <' u).val = t.val ^< u.val := rfl +@[simp] lemma val_notLessThan {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : (t ≮' u).val = t.val ^≮ u.val := rfl -@[simp] lemma equals_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma equals_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : (t₁ =' u₁) = (t₂ =' u₂) ↔ t₁ = t₂ ∧ u₁ = u₂ := by - simp [Language.TSemiformula.ext_iff, Language.TSemiterm.ext_iff, qqEQ] + simp [Language.Semiformula.ext_iff, Language.Semiterm.ext_iff, qqEQ] -@[simp] lemma notequals_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma notequals_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : (t₁ ≠' u₁) = (t₂ ≠' u₂) ↔ t₁ = t₂ ∧ u₁ = u₂ := by - simp [Language.TSemiformula.ext_iff, Language.TSemiterm.ext_iff, qqNEQ] + simp [Language.Semiformula.ext_iff, Language.Semiterm.ext_iff, qqNEQ] -@[simp] lemma lessThan_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma lessThan_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : (t₁ <' u₁) = (t₂ <' u₂) ↔ t₁ = t₂ ∧ u₁ = u₂ := by - simp [Language.TSemiformula.ext_iff, Language.TSemiterm.ext_iff, qqLT] + simp [Language.Semiformula.ext_iff, Language.Semiterm.ext_iff, qqLT] -@[simp] lemma notLessThan_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma notLessThan_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : (t₁ ≮' u₁) = (t₂ ≮' u₂) ↔ t₁ = t₂ ∧ u₁ = u₂ := by - simp [Language.TSemiformula.ext_iff, Language.TSemiterm.ext_iff, qqNLT] + simp [Language.Semiformula.ext_iff, Language.Semiterm.ext_iff, qqNLT] -@[simp] lemma neg_equals (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma neg_equals (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : ~(t₁ =' t₂) = (t₁ ≠' t₂) := by - ext; simp [Language.TSemiterm.equals, Language.TSemiterm.notEquals, qqEQ, qqNEQ] + ext; simp [Language.Semiterm.equals, Language.Semiterm.notEquals, qqEQ, qqNEQ] -@[simp] lemma neg_notEquals (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma neg_notEquals (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : ~(t₁ ≠' t₂) = (t₁ =' t₂) := by - ext; simp [Language.TSemiterm.equals, Language.TSemiterm.notEquals, qqEQ, qqNEQ] + ext; simp [Language.Semiterm.equals, Language.Semiterm.notEquals, qqEQ, qqNEQ] -@[simp] lemma neg_lessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma neg_lessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : ~(t₁ <' t₂) = (t₁ ≮' t₂) := by - ext; simp [Language.TSemiterm.lessThan, Language.TSemiterm.notLessThan, qqLT, qqNLT] + ext; simp [Language.Semiterm.lessThan, Language.Semiterm.notLessThan, qqLT, qqNLT] -@[simp] lemma neg_notLessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma neg_notLessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : ~(t₁ ≮' t₂) = (t₁ <' t₂) := by - ext; simp [Language.TSemiterm.lessThan, Language.TSemiterm.notLessThan, qqLT, qqNLT] + ext; simp [Language.Semiterm.lessThan, Language.Semiterm.notLessThan, qqLT, qqNLT] -@[simp] lemma shift_equals (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma shift_equals (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ =' t₂).shift = (t₁.shift =' t₂.shift) := by - ext; simp [Language.TSemiterm.equals, Language.TSemiterm.shift, Language.TSemiformula.shift, qqEQ] + ext; simp [Language.Semiterm.equals, Language.Semiterm.shift, Language.Semiformula.shift, qqEQ] -@[simp] lemma shift_notEquals (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma shift_notEquals (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ ≠' t₂).shift = (t₁.shift ≠' t₂.shift) := by - ext; simp [Language.TSemiterm.notEquals, Language.TSemiterm.shift, Language.TSemiformula.shift, qqNEQ] + ext; simp [Language.Semiterm.notEquals, Language.Semiterm.shift, Language.Semiformula.shift, qqNEQ] -@[simp] lemma shift_lessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma shift_lessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ <' t₂).shift = (t₁.shift <' t₂.shift) := by - ext; simp [Language.TSemiterm.lessThan, Language.TSemiterm.shift, Language.TSemiformula.shift, qqLT] + ext; simp [Language.Semiterm.lessThan, Language.Semiterm.shift, Language.Semiformula.shift, qqLT] -@[simp] lemma shift_notLessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma shift_notLessThan (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ ≮' t₂).shift = (t₁.shift ≮' t₂.shift) := by - ext; simp [Language.TSemiterm.notLessThan, Language.TSemiterm.shift, Language.TSemiformula.shift, qqNLT] + ext; simp [Language.Semiterm.notLessThan, Language.Semiterm.shift, Language.Semiformula.shift, qqNLT] -@[simp] lemma substs_equals (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma substs_equals (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ =' t₂).substs w = (t₁.substs w =' t₂.substs w) := by - ext; simp [Language.TSemiterm.equals, Language.TSemiterm.substs, Language.TSemiformula.substs, qqEQ] + ext; simp [Language.Semiterm.equals, Language.Semiterm.substs, Language.Semiformula.substs, qqEQ] -@[simp] lemma substs_notEquals (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma substs_notEquals (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ ≠' t₂).substs w = (t₁.substs w ≠' t₂.substs w) := by - ext; simp [Language.TSemiterm.notEquals, Language.TSemiterm.substs, Language.TSemiformula.substs, qqNEQ] + ext; simp [Language.Semiterm.notEquals, Language.Semiterm.substs, Language.Semiformula.substs, qqNEQ] -@[simp] lemma substs_lessThan (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma substs_lessThan (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ <' t₂).substs w = (t₁.substs w <' t₂.substs w) := by - ext; simp [Language.TSemiterm.lessThan, Language.TSemiterm.substs, Language.TSemiformula.substs, qqLT] + ext; simp [Language.Semiterm.lessThan, Language.Semiterm.substs, Language.Semiformula.substs, qqLT] -@[simp] lemma substs_notLessThan (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma substs_notLessThan (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ ≮' t₂).substs w = (t₁.substs w ≮' t₂.substs w) := by - ext; simp [Language.TSemiterm.notLessThan, Language.TSemiterm.substs, Language.TSemiformula.substs, qqNLT] + ext; simp [Language.Semiterm.notLessThan, Language.Semiterm.substs, Language.Semiformula.substs, qqNLT] -@[simp] lemma val_ball {n : V} (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : - (p.ball t).val = ^∀[n] (^#0 ^≮[n + 1] ⌜ℒₒᵣ⌝.termBShift n t.val) ^⋎[n + 1] p.val := by - simp [Language.TSemiformula.ball] +@[simp] lemma val_ball {n : V} (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : + (p.ball t).val = ^∀ (^#0 ^≮ ⌜ℒₒᵣ⌝.termBShift t.val) ^⋎ p.val := by + simp [Language.Semiformula.ball] -@[simp] lemma val_bex {n : V} (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : - (p.bex t).val = ^∃[n] (^#0 ^<[n + 1] ⌜ℒₒᵣ⌝.termBShift n t.val) ^⋏[n + 1] p.val := by - simp [Language.TSemiformula.bex] +@[simp] lemma val_bex {n : V} (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : + (p.bex t).val = ^∃ (^#0 ^< ⌜ℒₒᵣ⌝.termBShift t.val) ^⋏ p.val := by + simp [Language.Semiformula.bex] -lemma neg_ball (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +lemma neg_ball (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : ~(p.ball t) = (~p).bex t := by - ext; simp; rw [neg_all, neg_or] <;> simp [qqNLT, qqLT] + ext; simp; rw [neg_all, neg_or] <;> simp [qqNLT, qqLT, t.prop.termBShift.isUTerm] -lemma neg_bex (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +lemma neg_bex (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : ~(p.bex t) = (~p).ball t := by - ext; simp; rw [neg_ex, neg_and] <;> simp [qqNLT, qqLT] + ext; simp; rw [neg_ex, neg_and] <;> simp [qqNLT, qqLT, t.prop.termBShift.isUTerm] -@[simp] lemma shifts_ball (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +@[simp] lemma shifts_ball (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : (p.ball t).shift = p.shift.ball t.shift := by - simp [Language.TSemiformula.ball, Language.TSemiterm.bShift_shift_comm] + simp [Language.Semiformula.ball, Language.Semiterm.bShift_shift_comm] -@[simp] lemma shifts_bex (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +@[simp] lemma shifts_bex (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : (p.bex t).shift = p.shift.bex t.shift := by - simp [Language.TSemiformula.bex, Language.TSemiterm.bShift_shift_comm] + simp [Language.Semiformula.bex, Language.Semiterm.bShift_shift_comm] -@[simp] lemma substs_ball (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +@[simp] lemma substs_ball (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : (p.ball t)^/[w] = (p^/[w.q]).ball (t^ᵗ/[w]) := by - simp [Language.TSemiformula.ball] + simp [Language.Semiformula.ball] -@[simp] lemma substs_bex (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : +@[simp] lemma substs_bex (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : (p.bex t)^/[w] = (p^/[w.q]).bex (t^ᵗ/[w]) := by - simp [Language.TSemiformula.bex] + simp [Language.Semiformula.bex] -def tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : - ⌜ℒₒᵣ⌝.TSemiformulaVec m := ⟨substItr m w.val p.val k, by +def tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : + ⌜ℒₒᵣ⌝.SemiformulaVec m := ⟨substItr w.val p.val k, by intro i hi have : i < k := by simpa using hi simp only [gt_iff_lt, this, substItr_nth] - exact Language.Semiformula.substs p.prop (w.prop.cons (by simp))⟩ + exact p.prop.substs (w.prop.cons (by simp))⟩ -@[simp] lemma val_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : - (tSubstItr w p k).val = substItr m w.val p.val k := by simp [tSubstItr] +@[simp] lemma val_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : + (tSubstItr w p k).val = substItr w.val p.val k := by simp [tSubstItr] -@[simp] lemma len_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : +@[simp] lemma len_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : len (tSubstItr w p k).val = k := by simp -lemma nth_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) {i} (hi : i < k) : - (tSubstItr w p k).nth i (by simp [hi]) = p.substs (↑(k - (i + 1)) ∷ᵗ w) := by ext; simp [tSubstItr, Language.TSemiformula.substs, hi] +lemma nth_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) {i} (hi : i < k) : + (tSubstItr w p k).nth i (by simp [hi]) = p.substs (↑(k - (i + 1)) ∷ᵗ w) := by ext; simp [tSubstItr, Language.Semiformula.substs, hi] -lemma nth_tSubstItr' {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) {i} (hi : i < k) : +lemma nth_tSubstItr' {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) {i} (hi : i < k) : (tSubstItr w p k).nth (k - (i + 1)) (by simpa using sub_succ_lt_self hi) = p.substs (↑i ∷ᵗ w) := by - ext; simp [tSubstItr, Language.TSemiformula.substs, hi, sub_succ_lt_self hi, sub_succ_lt_selfs hi] + ext; simp [tSubstItr, Language.Semiformula.substs, hi, sub_succ_lt_self hi, sub_succ_lt_selfs hi] -@[simp] lemma neg_conj_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : +@[simp] lemma neg_conj_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : ~(tSubstItr w p k).conj = (tSubstItr w (~p) k).disj := by ext; simp [neg_conj_substItr p.prop w.prop] -@[simp] lemma neg_disj_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : +@[simp] lemma neg_disj_tSubstItr {n m : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : ~(tSubstItr w p k).disj = (tSubstItr w (~p) k).conj := by ext; simp [neg_disj_substItr p.prop w.prop] -@[simp] lemma substs_conj_tSubstItr {n m l : V} (v : ⌜ℒₒᵣ⌝.TSemitermVec m l) (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : +@[simp] lemma substs_conj_tSubstItr {n m l : V} (v : ⌜ℒₒᵣ⌝.SemitermVec m l) (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : (tSubstItr w p k).conj.substs v = (tSubstItr (w.substs v) p k).conj := by - ext; simp [Language.TSemiformula.substs, Language.TSemitermVec.substs] + ext; simp [Language.Semiformula.substs, Language.SemitermVec.substs] rw [substs_conj_substItr p.prop w.prop v.prop] -@[simp] lemma substs_disj_tSubstItr {n m l : V} (v : ⌜ℒₒᵣ⌝.TSemitermVec m l) (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) (k : V) : +@[simp] lemma substs_disj_tSubstItr {n m l : V} (v : ⌜ℒₒᵣ⌝.SemitermVec m l) (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) (k : V) : (tSubstItr w p k).disj.substs v = (tSubstItr (w.substs v) p k).disj := by - ext; simp [Language.TSemiformula.substs, Language.TSemitermVec.substs] + ext; simp [Language.Semiformula.substs, Language.SemitermVec.substs] rw [substs_disj_substItr p.prop w.prop v.prop] -@[simp] lemma equals_fvfree {t u : ⌜ℒₒᵣ⌝.TSemiterm n} : (t =' u).FVFree ↔ t.FVFree ∧ u.FVFree := by - simp [Language.TSemiformula.FVFree.iff, Language.TSemiterm.FVFree.iff] +/- +@[simp] lemma equals_fvfree {t u : ⌜ℒₒᵣ⌝.Semiterm n} : (t =' u).FVFree ↔ t.FVFree ∧ u.FVFree := by + simp [Language.Semiformula.FVFree.iff, Language.Semiterm.FVFree.iff] -@[simp] lemma notEquals_fvfree {t u : ⌜ℒₒᵣ⌝.TSemiterm n} : (t ≠' u).FVFree ↔ t.FVFree ∧ u.FVFree := by - simp [Language.TSemiformula.FVFree.iff, Language.TSemiterm.FVFree.iff] +@[simp] lemma notEquals_fvfree {t u : ⌜ℒₒᵣ⌝.Semiterm n} : (t ≠' u).FVFree ↔ t.FVFree ∧ u.FVFree := by + simp [Language.Semiformula.FVFree.iff, Language.Semiterm.FVFree.iff] -@[simp] lemma lessThan_fvfree {t u : ⌜ℒₒᵣ⌝.TSemiterm n} : (t <' u).FVFree ↔ t.FVFree ∧ u.FVFree := by - simp [Language.TSemiformula.FVFree.iff, Language.TSemiterm.FVFree.iff] +@[simp] lemma lessThan_fvfree {t u : ⌜ℒₒᵣ⌝.Semiterm n} : (t <' u).FVFree ↔ t.FVFree ∧ u.FVFree := by + simp [Language.Semiformula.FVFree.iff, Language.Semiterm.FVFree.iff] -@[simp] lemma notLessThan_fvfree {t u : ⌜ℒₒᵣ⌝.TSemiterm n} : (t ≮' u).FVFree ↔ t.FVFree ∧ u.FVFree := by - simp [Language.TSemiformula.FVFree.iff, Language.TSemiterm.FVFree.iff] +@[simp] lemma notLessThan_fvfree {t u : ⌜ℒₒᵣ⌝.Semiterm n} : (t ≮' u).FVFree ↔ t.FVFree ∧ u.FVFree := by + simp [Language.Semiformula.FVFree.iff, Language.Semiterm.FVFree.iff] +-/ end Formalized -lemma Language.TSemiformula.ball_eq_imp {n : V} (t : ⌜ℒₒᵣ⌝.TSemiterm n) (p : ⌜ℒₒᵣ⌝.TSemiformula (n + 1)) : - p.ball t = (⌜ℒₒᵣ⌝.bvar 0 <' t.bShift ⟶ p).all := by simp [Language.TSemiformula.ball, imp_def] +lemma Language.Semiformula.ball_eq_imp {n : V} (t : ⌜ℒₒᵣ⌝.Semiterm n) (p : ⌜ℒₒᵣ⌝.Semiformula (n + 1)) : + p.ball t = (⌜ℒₒᵣ⌝.bvar 0 <' t.bShift ⟶ p).all := by simp [Language.Semiformula.ball, imp_def] end LO.Arith diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean index 7711a2b..f146e8a 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean @@ -16,42 +16,44 @@ section derivation variable (L) -def Language.FormulaSet (s : V) : Prop := ∀ p ∈ s, L.Formula p +def Language.IsFormulaSet (s : V) : Prop := ∀ p ∈ s, L.IsFormula p variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.formulaSetDef (pL : LDef) : 𝚫₁.Semisentence 1 := .mkDelta +def _root_.LO.FirstOrder.Arith.LDef.isFormulaSetDef (pL : LDef) : 𝚫₁.Semisentence 1 := .mkDelta (.mkSigma “s | ∀ p ∈' s, !pL.isSemiformulaDef.sigma 0 p” (by simp)) (.mkPi “s | ∀ p ∈' s, !pL.isSemiformulaDef.pi 0 p” (by simp)) variable (L) -lemma formulaSet_defined : 𝚫₁-Predicate L.FormulaSet via pL.formulaSetDef := - ⟨by intro v; simp [LDef.formulaSetDef, HierarchySymbol.Semiformula.val_sigma, (semiformula_defined L).df.iff, (semiformula_defined L).proper.iff'], - by intro v; simp [LDef.formulaSetDef, HierarchySymbol.Semiformula.val_sigma, (semiformula_defined L).df.iff]; rfl⟩ +lemma Language.isFormulaSet_defined : 𝚫₁-Predicate L.IsFormulaSet via pL.isFormulaSetDef := + ⟨by intro v; simp [LDef.isFormulaSetDef, HierarchySymbol.Semiformula.val_sigma, L.isSemiformula_defined.df.iff, L.isSemiformula_defined.proper.iff'], + by intro v; simp [LDef.isFormulaSetDef, HierarchySymbol.Semiformula.val_sigma, L.isSemiformula_defined.df.iff]; rfl⟩ -@[simp] instance formulaSet_definable : 𝚫₁-Predicate L.FormulaSet := (formulaSet_defined L).to_definable +instance Language.isFormulaSet_definable : 𝚫₁-Predicate L.IsFormulaSet := L.isFormulaSet_defined.to_definable + +instance Language.isFormulaSet_definable' : Γ-[m + 1]-Predicate L.IsFormulaSet := .of_deltaOne L.isFormulaSet_definable end -@[simp] lemma Language.FormulaSet.empty : L.FormulaSet ∅ := fun p ↦ by simp +@[simp] lemma Language.IsFormulaSet.empty : L.IsFormulaSet ∅ := fun p ↦ by simp -@[simp] lemma Language.FormulaSet.singleton {p} : L.FormulaSet {p} ↔ L.Formula p := +@[simp] lemma Language.IsFormulaSet.singleton {p} : L.IsFormulaSet {p} ↔ L.IsFormula p := ⟨fun h ↦ h p (by simp), fun h p ↦ by simp only [mem_singleton_iff] rintro rfl; exact h⟩ -@[simp] lemma Language.FormulaSet.insert_iff {p s} : L.FormulaSet (insert p s) ↔ L.Formula p ∧ L.FormulaSet s := +@[simp] lemma Language.IsFormulaSet.insert_iff {p s} : L.IsFormulaSet (insert p s) ↔ L.IsFormula p ∧ L.IsFormulaSet s := ⟨fun h ↦ ⟨h p (by simp), fun q hq ↦ h q (by simp [hq])⟩, by rintro ⟨hp, hs⟩ q; simp; rintro (rfl | hqs) · exact hp · exact hs q hqs⟩ -alias ⟨Language.FormulaSet.insert, _⟩ := Language.FormulaSet.insert_iff +alias ⟨Language.IsFormulaSet.insert, _⟩ := Language.IsFormulaSet.insert_iff -@[simp] lemma Language.FormulaSet.union {s₁ s₂} : L.FormulaSet (s₁ ∪ s₂) ↔ L.FormulaSet s₁ ∧ L.FormulaSet s₂ := +@[simp] lemma Language.IsFormulaSet.union {s₁ s₂} : L.IsFormulaSet (s₁ ∪ s₂) ↔ L.IsFormulaSet s₁ ∧ L.IsFormulaSet s₂ := ⟨fun h ↦ ⟨fun p hp ↦ h p (by simp [hp]), fun p hp ↦ h p (by simp [hp])⟩, fun h p hp ↦ by rcases mem_cup_iff.mp hp with (h₁ | h₂) @@ -73,16 +75,16 @@ section setShift lemma mem_setShift_iff {s y : V} : y ∈ L.setShift s ↔ ∃ x ∈ s, y = L.shift x := Classical.choose!_spec (setShift_existsUnique L s) y -lemma Language.FormulaSet.setShift {s : V} (h : L.FormulaSet s) : L.FormulaSet (L.setShift s) := by - simp [Language.FormulaSet, mem_setShift_iff] +lemma Language.IsFormulaSet.setShift {s : V} (h : L.IsFormulaSet s) : L.IsFormulaSet (L.setShift s) := by + simp [Language.IsFormulaSet, mem_setShift_iff] rintro _ p hp rfl; exact (h p hp).shift lemma shift_mem_setShift {p s : V} (h : p ∈ s) : L.shift p ∈ L.setShift s := mem_setShift_iff.mpr ⟨p, h, rfl⟩ -@[simp] lemma Language.FormulaSet.setShift_iff {s : V} : - L.FormulaSet (L.setShift s) ↔ L.FormulaSet s := - ⟨by intro h p hp; simpa using h (L.shift p) (shift_mem_setShift hp), Language.FormulaSet.setShift⟩ +@[simp] lemma Language.IsFormulaSet.setShift_iff {s : V} : + L.IsFormulaSet (L.setShift s) ↔ L.IsFormulaSet s := + ⟨by intro h p hp; simpa using h (L.shift p) (shift_mem_setShift hp), Language.IsFormulaSet.setShift⟩ @[simp] lemma mem_setShift_union {s t : V} : L.setShift (s ∪ t) = L.setShift s ∪ L.setShift t := mem_ext <| by simp [mem_setShift_iff]; intro x @@ -120,10 +122,10 @@ def _root_.LO.FirstOrder.Arith.LDef.setShiftDef (pL : LDef) : 𝚺₁.Semisenten variable (L) -lemma setShift_defined : 𝚺₁-Function₁ L.setShift via pL.setShiftDef := by - intro v; simp [LDef.setShiftDef, setShift_graph, (shift_defined L).df.iff] +lemma Language.setShift_defined : 𝚺₁-Function₁ L.setShift via pL.setShiftDef := by + intro v; simp [LDef.setShiftDef, setShift_graph, L.shift_defined.df.iff] -@[simp, definability] instance setShift_definable : 𝚺₁-Function₁ L.setShift := (setShift_defined L).to_definable +instance Language.setShift_definable : 𝚺₁-Function₁ L.setShift := L.setShift_defined.to_definable end @@ -346,13 +348,13 @@ abbrev conseq (x : V) : V := π₁ x variable (T) def Phi (C : Set V) (d : V) : Prop := - L.FormulaSet (fstIdx d) ∧ + L.IsFormulaSet (fstIdx d) ∧ ( (∃ s p, d = axL s p ∧ p ∈ s ∧ L.neg p ∈ s) ∨ (∃ s, d = verumIntro s ∧ ^⊤ ∈ s) ∨ (∃ s p q dp dq, d = andIntro s p q dp dq ∧ p ^⋏ q ∈ s ∧ (fstIdx dp = insert p s ∧ dp ∈ C) ∧ (fstIdx dq = insert q s ∧ dq ∈ C)) ∨ (∃ s p q dpq, d = orIntro s p q dpq ∧ p ^⋎ q ∈ s ∧ fstIdx dpq = insert p (insert q s) ∧ dpq ∈ C) ∨ (∃ s p dp, d = allIntro s p dp ∧ ^∀ p ∈ s ∧ fstIdx dp = insert (L.free p) (L.setShift s) ∧ dp ∈ C) ∨ - (∃ s p t dp, d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.Term t ∧ fstIdx dp = insert (L.substs₁ t p) s ∧ dp ∈ C) ∨ + (∃ s p t dp, d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.IsTerm t ∧ fstIdx dp = insert (L.substs₁ t p) s ∧ dp ∈ C) ∨ (∃ s d', d = wkRule s d' ∧ fstIdx d' ⊆ s ∧ d' ∈ C) ∨ (∃ s d', d = shiftRule s d' ∧ s = L.setShift (fstIdx d') ∧ d' ∈ C) ∨ (∃ s p d₁ d₂, d = cutRule s p d₁ d₂ ∧ (fstIdx d₁ = insert p s ∧ d₁ ∈ C) ∧ (fstIdx d₂ = insert (L.neg p) s ∧ d₂ ∈ C)) ∨ @@ -360,7 +362,7 @@ def Phi (C : Set V) (d : V) : Prop := private lemma phi_iff (C d : V) : Phi T {x | x ∈ C} d ↔ - L.FormulaSet (fstIdx d) ∧ + L.IsFormulaSet (fstIdx d) ∧ ( (∃ s < d, ∃ p < d, d = axL s p ∧ p ∈ s ∧ L.neg p ∈ s) ∨ (∃ s < d, d = verumIntro s ∧ ^⊤ ∈ s) ∨ (∃ s < d, ∃ p < d, ∃ q < d, ∃ dp < d, ∃ dq < d, @@ -370,7 +372,7 @@ private lemma phi_iff (C d : V) : (∃ s < d, ∃ p < d, ∃ dp < d, d = allIntro s p dp ∧ ^∀ p ∈ s ∧ fstIdx dp = insert (L.free p) (L.setShift s) ∧ dp ∈ C) ∨ (∃ s < d, ∃ p < d, ∃ t < d, ∃ dp < d, - d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.Term t ∧ fstIdx dp = insert (L.substs₁ t p) s ∧ dp ∈ C) ∨ + d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.IsTerm t ∧ fstIdx dp = insert (L.substs₁ t p) s ∧ dp ∈ C) ∨ (∃ s < d, ∃ d' < d, d = wkRule s d' ∧ fstIdx d' ⊆ s ∧ d' ∈ C) ∨ (∃ s < d, ∃ d' < d, @@ -412,22 +414,22 @@ private lemma phi_iff (C d : V) : def blueprint {pL : LDef} (pT : pL.TDef) : Fixpoint.Blueprint 0 := ⟨.mkDelta (.mkSigma “d C | - (∃ fst, !fstIdxDef fst d ∧ !pL.formulaSetDef.sigma fst) ∧ + (∃ fst, !fstIdxDef fst d ∧ !pL.isFormulaSetDef.sigma fst) ∧ ( (∃ s < d, ∃ p < d, !axLDef d s p ∧ p ∈ s ∧ ∃ np, !pL.negDef np p ∧ np ∈ s) ∨ - (∃ s < d, !verumIntroDef d s ∧ ∃ vrm, !qqVerumDef vrm 0 ∧ vrm ∈ s) ∨ + (∃ s < d, !verumIntroDef d s ∧ ∃ vrm, !qqVerumDef vrm ∧ vrm ∈ s) ∨ (∃ s < d, ∃ p < d, ∃ q < d, ∃ dp < d, ∃ dq < d, - !andIntroDef d s p q dp dq ∧ (∃ and, !qqAndDef and 0 p q ∧ and ∈ s) ∧ + !andIntroDef d s p q dp dq ∧ (∃ and, !qqAndDef and p q ∧ and ∈ s) ∧ (∃ c, !fstIdxDef c dp ∧ !insertDef c p s ∧ dp ∈ C) ∧ (∃ c, !fstIdxDef c dq ∧ !insertDef c q s ∧ dq ∈ C)) ∨ (∃ s < d, ∃ p < d, ∃ q < d, ∃ dpq < d, - !orIntroDef d s p q dpq ∧ (∃ or, !qqOrDef or 0 p q ∧ or ∈ s) ∧ + !orIntroDef d s p q dpq ∧ (∃ or, !qqOrDef or p q ∧ or ∈ s) ∧ ∃ c, !fstIdxDef c dpq ∧ ∃ c', !insertDef c' q s ∧ !insertDef c p c' ∧ dpq ∈ C) ∨ (∃ s < d, ∃ p < d, ∃ dp < d, - !allIntroDef d s p dp ∧ (∃ all, !qqAllDef all 0 p ∧ all ∈ s) ∧ + !allIntroDef d s p dp ∧ (∃ all, !qqAllDef all p ∧ all ∈ s) ∧ ∃ c, !fstIdxDef c dp ∧ ∃ fp, !pL.freeDef fp p ∧ ∃ ss, !pL.setShiftDef ss s ∧ !insertDef c fp ss ∧ dp ∈ C) ∨ (∃ s < d, ∃ p < d, ∃ t < d, ∃ dp < d, - !exIntroDef d s p t dp ∧ (∃ ex, !qqExDef ex 0 p ∧ ex ∈ s) ∧ + !exIntroDef d s p t dp ∧ (∃ ex, !qqExDef ex p ∧ ex ∈ s) ∧ !pL.isSemitermDef.sigma 0 t ∧ ∃ c, !fstIdxDef c dp ∧ ∃ pt, !pL.substs₁Def pt t p ∧ !insertDef c pt s ∧ dp ∈ C) ∨ (∃ s < d, ∃ d' < d, !wkRuleDef d s d' ∧ ∃ c, !fstIdxDef c d' ∧ !bitSubsetDef c s ∧ d' ∈ C) ∨ @@ -441,22 +443,22 @@ def blueprint {pL : LDef} (pT : pL.TDef) : Fixpoint.Blueprint 0 := ⟨.mkDelta !rootDef d s p ∧ p ∈ s ∧ !pT.ch.sigma p) )” (by simp)) (.mkPi “d C | - (∀ fst, !fstIdxDef fst d → !pL.formulaSetDef.pi fst) ∧ + (∀ fst, !fstIdxDef fst d → !pL.isFormulaSetDef.pi fst) ∧ ( (∃ s < d, ∃ p < d, !axLDef d s p ∧ p ∈ s ∧ ∀ np, !pL.negDef np p → np ∈ s) ∨ - (∃ s < d, !verumIntroDef d s ∧ ∀ vrm, !qqVerumDef vrm 0 → vrm ∈ s) ∨ + (∃ s < d, !verumIntroDef d s ∧ ∀ vrm, !qqVerumDef vrm → vrm ∈ s) ∨ (∃ s < d, ∃ p < d, ∃ q < d, ∃ dp < d, ∃ dq < d, - !andIntroDef d s p q dp dq ∧ (∀ and, !qqAndDef and 0 p q → and ∈ s) ∧ + !andIntroDef d s p q dp dq ∧ (∀ and, !qqAndDef and p q → and ∈ s) ∧ (∀ c, !fstIdxDef c dp → !insertDef c p s ∧ dp ∈ C) ∧ (∀ c, !fstIdxDef c dq → !insertDef c q s ∧ dq ∈ C)) ∨ (∃ s < d, ∃ p < d, ∃ q < d, ∃ dpq < d, - !orIntroDef d s p q dpq ∧ (∀ or, !qqOrDef or 0 p q → or ∈ s) ∧ + !orIntroDef d s p q dpq ∧ (∀ or, !qqOrDef or p q → or ∈ s) ∧ ∀ c, !fstIdxDef c dpq → ∀ c', !insertDef c' q s → !insertDef c p c' ∧ dpq ∈ C) ∨ (∃ s < d, ∃ p < d, ∃ dp < d, - !allIntroDef d s p dp ∧ (∀ all, !qqAllDef all 0 p → all ∈ s) ∧ + !allIntroDef d s p dp ∧ (∀ all, !qqAllDef all p → all ∈ s) ∧ ∀ c, !fstIdxDef c dp → ∀ fp, !pL.freeDef fp p → ∀ ss, !pL.setShiftDef ss s → !insertDef c fp ss ∧ dp ∈ C) ∨ (∃ s < d, ∃ p < d, ∃ t < d, ∃ dp < d, - !exIntroDef d s p t dp ∧ (∀ ex, !qqExDef ex 0 p → ex ∈ s) ∧ + !exIntroDef d s p t dp ∧ (∀ ex, !qqExDef ex p → ex ∈ s) ∧ !pL.isSemitermDef.pi 0 t ∧ ∀ c, !fstIdxDef c dp → ∀ pt, !pL.substs₁Def pt t p → !insertDef c pt s ∧ dp ∈ C) ∨ (∃ s < d, ∃ d' < d, @@ -478,64 +480,71 @@ def construction : Fixpoint.Construction V (blueprint pT) where intro v /- simp? [blueprint, HierarchySymbol.Semiformula.val_sigma, - (formulaSet_defined L).df.iff, (formulaSet_defined L).proper.iff', - (neg_defined L).df.iff, - (free_defined L).df.iff, - (setShift_defined L).df.iff, - (isSemiterm_defined L).df.iff, (isSemiterm_defined L).proper.iff', - (substs₁_defined L).df.iff, + L.isFormulaSet_defined.df.iff, L.isFormulaSet_defined.proper.iff', + L.neg_defined.df.iff, + L.free_defined.df.iff, + L.setShift_defined.df.iff, + L.isSemiterm_defined.df.iff, + L.isSemiterm_defined.proper.iff', + L.substs₁_defined.df.iff, T.mem_defined.df.iff, T.mem_defined.proper.iff'] -/ - - simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, HierarchySymbol.Semiformula.val_sigma, - HierarchySymbol.Semiformula.sigma_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_and, + simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, + HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.sigma_mkDelta, + HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_and, Semiformula.eval_ex, Semiformula.eval_substs, Matrix.comp_vecCons', Semiterm.val_bvar, Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.cons_val_one, Matrix.vecHead, - Matrix.constant_eq_singleton, eval_fstIdxDef, (formulaSet_defined L).df.iff, + Matrix.constant_eq_singleton, eval_fstIdxDef, L.isFormulaSet_defined.df.iff, LogicalConnective.Prop.and_eq, exists_eq_left, LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, eval_axLDef, Semiformula.eval_operator₂, Structure.Mem.mem, - (neg_defined L).df.iff, eval_verumIntroDef, Semiterm.val_operator₀, - Structure.numeral_eq_numeral, ORingSymbol.zero_eq_zero, eval_qqVerumDef, - Matrix.cons_val_three, Fin.succ_one_eq_two, Matrix.cons_val_four, Matrix.cons_val_succ, - Matrix.cons_app_five, eval_andIntroDef, eval_qqAndDef, insert_defined_iff, - Matrix.cons_app_seven, Matrix.cons_app_six, eval_orIntroDef, eval_qqOrDef, eval_allIntroDef, - eval_qqAllDef, (free_defined L).df.iff, (setShift_defined L).df.iff, eval_exIntroDef, - eval_qqExDef, (isSemiterm_defined L).df.iff, (substs₁_defined L).df.iff, eval_wkRuleDef, - bitSubset_defined_iff, eval_shiftRuleDef, eval_cutRuleDef, eval_rootDef, T.mem_defined.df.iff, - LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, - Semiformula.eval_all, LogicalConnective.HomClass.map_imply, - (formulaSet_defined L).proper.iff', LogicalConnective.Prop.arrow_eq, forall_eq, - (isSemiterm_defined L).proper.iff', Structure.Eq.eq, T.mem_defined.proper.iff'], + L.neg_defined.df.iff, eval_verumIntroDef, eval_qqVerumDef, Matrix.cons_val_three, + Fin.succ_one_eq_two, Matrix.cons_val_four, Matrix.cons_val_succ, Matrix.cons_app_five, + eval_andIntroDef, eval_qqAndDef, insert_defined_iff, Matrix.cons_app_seven, + Matrix.cons_app_six, eval_orIntroDef, eval_qqOrDef, eval_allIntroDef, eval_qqAllDef, + L.free_defined.df.iff, L.setShift_defined.df.iff, eval_exIntroDef, eval_qqExDef, + Semiterm.val_operator₀, Structure.numeral_eq_numeral, ORingSymbol.zero_eq_zero, + L.isSemiterm_defined.df.iff, L.substs₁_defined.df.iff, eval_wkRuleDef, bitSubset_defined_iff, + eval_shiftRuleDef, eval_cutRuleDef, eval_rootDef, T.mem_defined.df.iff, + LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, + HierarchySymbol.Semiformula.val_mkPi, Semiformula.eval_all, + LogicalConnective.HomClass.map_imply, L.isFormulaSet_defined.proper.iff', + LogicalConnective.Prop.arrow_eq, forall_eq, L.isSemiterm_defined.proper.iff', Structure.Eq.eq, + T.mem_defined.proper.iff'] + , by intro v /- simp? [phi_iff, blueprint, HierarchySymbol.Semiformula.val_sigma, - (formulaSet_defined L).df.iff, (formulaSet_defined L).proper.iff', - (neg_defined L).df.iff, - (free_defined L).df.iff, - (setShift_defined L).df.iff, - (isSemiterm_defined L).df.iff, (isSemiterm_defined L).proper.iff', - (substs₁_defined L).df.iff, + L.isFormulaSet_defined.df.iff, + L.isFormulaSet_defined.proper.iff', + L.neg_defined.df.iff, + L.free_defined.df.iff, + L.setShift_defined.df.iff, + L.isSemiterm_defined.df.iff, + L.isSemiterm_defined.proper.iff', + L.substs₁_defined.df.iff, T.mem_defined.df.iff] -/ simp only [Fin.isValue, phi_iff, Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, - HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.val_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, - LogicalConnective.HomClass.map_and, Semiformula.eval_ex, Semiformula.eval_substs, - Matrix.comp_vecCons', Semiterm.val_bvar, Matrix.cons_val_zero, Matrix.cons_val_fin_one, - Matrix.cons_val_one, Matrix.vecHead, Matrix.constant_eq_singleton, eval_fstIdxDef, - (formulaSet_defined L).df.iff, LogicalConnective.Prop.and_eq, exists_eq_left, - LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, Matrix.cons_val_two, - Matrix.vecTail, Function.comp_apply, Fin.succ_zero_eq_one, eval_axLDef, - Semiformula.eval_operator₂, Structure.Mem.mem, (neg_defined L).df.iff, eval_verumIntroDef, + HierarchySymbol.Semiformula.val_sigma, HierarchySymbol.Semiformula.val_mkDelta, + HierarchySymbol.Semiformula.val_mkSigma, LogicalConnective.HomClass.map_and, + Semiformula.eval_ex, Semiformula.eval_substs, Matrix.comp_vecCons', Semiterm.val_bvar, + Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.cons_val_one, Matrix.vecHead, + Matrix.constant_eq_singleton, eval_fstIdxDef, L.isFormulaSet_defined.df.iff, + LogicalConnective.Prop.and_eq, exists_eq_left, LogicalConnective.HomClass.map_or, + Semiformula.eval_bexLT, Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, + Fin.succ_zero_eq_one, eval_axLDef, Semiformula.eval_operator₂, Structure.Mem.mem, + L.neg_defined.df.iff, eval_verumIntroDef, eval_qqVerumDef, Matrix.cons_val_three, + Fin.succ_one_eq_two, Matrix.cons_val_four, Matrix.cons_val_succ, Matrix.cons_app_five, + eval_andIntroDef, eval_qqAndDef, insert_defined_iff, Matrix.cons_app_seven, + Matrix.cons_app_six, eval_orIntroDef, eval_qqOrDef, eval_allIntroDef, eval_qqAllDef, + L.free_defined.df.iff, L.setShift_defined.df.iff, eval_exIntroDef, eval_qqExDef, Semiterm.val_operator₀, Structure.numeral_eq_numeral, ORingSymbol.zero_eq_zero, - eval_qqVerumDef, Matrix.cons_val_three, Fin.succ_one_eq_two, Matrix.cons_val_four, - Matrix.cons_val_succ, Matrix.cons_app_five, eval_andIntroDef, eval_qqAndDef, - insert_defined_iff, Matrix.cons_app_seven, Matrix.cons_app_six, eval_orIntroDef, eval_qqOrDef, - eval_allIntroDef, eval_qqAllDef, (free_defined L).df.iff, (setShift_defined L).df.iff, - eval_exIntroDef, eval_qqExDef, (isSemiterm_defined L).df.iff, (substs₁_defined L).df.iff, - eval_wkRuleDef, bitSubset_defined_iff, eval_shiftRuleDef, eval_cutRuleDef, eval_rootDef, - T.mem_defined.df.iff, LogicalConnective.Prop.or_eq]⟩ + L.isSemiterm_defined.df.iff, L.substs₁_defined.df.iff, eval_wkRuleDef, bitSubset_defined_iff, + eval_shiftRuleDef, eval_cutRuleDef, eval_rootDef, T.mem_defined.df.iff, + LogicalConnective.Prop.or_eq] + ⟩ monotone := by rintro C C' hC _ d ⟨hs, H⟩ refine ⟨hs, ?_⟩ @@ -589,25 +598,23 @@ section def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationDef {pL : LDef} (pT : pL.TDef) : 𝚫₁.Semisentence 1 := (blueprint pT).fixpointDefΔ₁ -lemma derivation_defined : 𝚫₁-Predicate T.Derivation via pT.derivationDef := (construction T).fixpoint_definedΔ₁ +lemma Language.Theory.derivation_defined : 𝚫₁-Predicate T.Derivation via pT.derivationDef := (construction T).fixpoint_definedΔ₁ -instance derivation_definable : 𝚫₁-Predicate T.Derivation := (derivation_defined T).to_definable +instance Language.Theory.derivation_definable : 𝚫₁-Predicate T.Derivation := T.derivation_defined.to_definable -@[simp] instance derivatin_definable' (Γ) : Γ-[m + 1]-Predicate T.Derivation := - .of_deltaOne (derivation_definable T) _ _ +instance Language.Theory.derivation_definable' : Γ-[m + 1]-Predicate T.Derivation := T.derivation_definable.of_deltaOne def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationOfDef {pL : LDef} (pT : pL.TDef) : 𝚫₁.Semisentence 2 := .mkDelta (.mkSigma “d s | !fstIdxDef s d ∧ !pT.derivationDef.sigma d” (by simp)) (.mkPi “d s | !fstIdxDef s d ∧ !pT.derivationDef.pi d” (by simp)) -lemma derivationOf_defined : 𝚫₁-Relation T.DerivationOf via pT.derivationOfDef := - ⟨by intro v; simp [LDef.TDef.derivationOfDef, HierarchySymbol.Semiformula.val_sigma, (derivation_defined T).proper.iff'], - by intro v; simp [LDef.TDef.derivationOfDef, HierarchySymbol.Semiformula.val_sigma, (derivation_defined T).df.iff, eq_comm (b := fstIdx (v 0))]; rfl⟩ +lemma Language.Theory.derivationOf_defined : 𝚫₁-Relation T.DerivationOf via pT.derivationOfDef := + ⟨by intro v; simp [LDef.TDef.derivationOfDef, HierarchySymbol.Semiformula.val_sigma, T.derivation_defined.proper.iff'], + by intro v; simp [LDef.TDef.derivationOfDef, HierarchySymbol.Semiformula.val_sigma, T.derivation_defined.df.iff, eq_comm (b := fstIdx (v 0))]; rfl⟩ -instance derivationOf_definable : 𝚫₁-Relation T.DerivationOf := (derivationOf_defined T).to_definable +instance Language.Theory.derivationOf_definable : 𝚫₁-Relation T.DerivationOf := T.derivationOf_defined.to_definable -@[simp] instance derivatinOf_definable' (Γ) : Γ-[m + 1]-Relation T.DerivationOf := - .of_deltaOne (derivationOf_definable T) _ _ +instance Language.Theory.derivationOf_definable' : Γ-[m + 1]-Relation T.DerivationOf := T.derivationOf_definable.of_deltaOne def _root_.LO.FirstOrder.Arith.LDef.TDef.derivableDef {pL : LDef} (pT : pL.TDef) : 𝚺₁.Semisentence 1 := .mkSigma “s | ∃ d, !pT.derivationOfDef.sigma d s” (by simp) @@ -637,13 +644,13 @@ variable {T} lemma Language.Theory.Derivation.case_iff {d : V} : T.Derivation d ↔ - L.FormulaSet (fstIdx d) ∧ + L.IsFormulaSet (fstIdx d) ∧ ( (∃ s p, d = axL s p ∧ p ∈ s ∧ L.neg p ∈ s) ∨ (∃ s, d = verumIntro s ∧ ^⊤ ∈ s) ∨ (∃ s p q dp dq, d = andIntro s p q dp dq ∧ p ^⋏ q ∈ s ∧ T.DerivationOf dp (insert p s) ∧ T.DerivationOf dq (insert q s)) ∨ (∃ s p q dpq, d = orIntro s p q dpq ∧ p ^⋎ q ∈ s ∧ T.DerivationOf dpq (insert p (insert q s))) ∨ (∃ s p dp, d = allIntro s p dp ∧ ^∀ p ∈ s ∧ T.DerivationOf dp (insert (L.free p) (L.setShift s))) ∨ - (∃ s p t dp, d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.Term t ∧ T.DerivationOf dp (insert (L.substs₁ t p) s)) ∨ + (∃ s p t dp, d = exIntro s p t dp ∧ ^∃ p ∈ s ∧ L.IsTerm t ∧ T.DerivationOf dp (insert (L.substs₁ t p) s)) ∨ (∃ s d', d = wkRule s d' ∧ fstIdx d' ⊆ s ∧ T.Derivation d') ∨ (∃ s d', d = shiftRule s d' ∧ s = L.setShift (fstIdx d') ∧ T.Derivation d') ∨ (∃ s p d₁ d₂, d = cutRule s p d₁ d₂ ∧ T.DerivationOf d₁ (insert p s) ∧ T.DerivationOf d₂ (insert (L.neg p) s)) ∨ @@ -652,44 +659,44 @@ lemma Language.Theory.Derivation.case_iff {d : V} : alias ⟨Language.Theory.Derivation.case, Language.Theory.Derivation.mk⟩ := Language.Theory.Derivation.case_iff -lemma Language.Theory.DerivationOf.formulaSet {d s : V} (h : T.DerivationOf d s) : L.FormulaSet s := by +lemma Language.Theory.DerivationOf.isFormulaSet {d s : V} (h : T.DerivationOf d s) : L.IsFormulaSet s := by simpa [h.1] using h.2.case.1 -lemma Language.Theory.Derivation.axL {s p : V} (hs : L.FormulaSet s) (h : p ∈ s) (hn : L.neg p ∈ s) : T.Derivation (axL s p) := +lemma Language.Theory.Derivation.axL {s p : V} (hs : L.IsFormulaSet s) (h : p ∈ s) (hn : L.neg p ∈ s) : T.Derivation (axL s p) := Language.Theory.Derivation.mk ⟨by simpa using hs, Or.inl ⟨s, p, rfl, h, hn⟩⟩ -lemma Language.Theory.Derivation.verumIntro {s : V} (hs : L.FormulaSet s) (h : ^⊤ ∈ s) : +lemma Language.Theory.Derivation.verumIntro {s : V} (hs : L.IsFormulaSet s) (h : ^⊤ ∈ s) : T.Derivation (verumIntro s) := Language.Theory.Derivation.mk ⟨by simpa using hs, Or.inr <| Or.inl ⟨s, rfl, h⟩⟩ lemma Language.Theory.Derivation.andIntro {s p q dp dq : V} (h : p ^⋏ q ∈ s) (hdp : T.DerivationOf dp (insert p s)) (hdq : T.DerivationOf dq (insert q s)) : T.Derivation (andIntro s p q dp dq) := - Language.Theory.Derivation.mk ⟨by simp; intro r hr; exact hdp.formulaSet r (by simp [hr]), + Language.Theory.Derivation.mk ⟨by simp; intro r hr; exact hdp.isFormulaSet r (by simp [hr]), Or.inr <| Or.inr <| Or.inl ⟨s, p, q, dp, dq, rfl, h, hdp, hdq⟩⟩ lemma Language.Theory.Derivation.orIntro {s p q dpq : V} (h : p ^⋎ q ∈ s) (hdpq : T.DerivationOf dpq (insert p (insert q s))) : T.Derivation (orIntro s p q dpq) := - Language.Theory.Derivation.mk ⟨by simp; intro r hr; exact hdpq.formulaSet r (by simp [hr]), + Language.Theory.Derivation.mk ⟨by simp; intro r hr; exact hdpq.isFormulaSet r (by simp [hr]), Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨s, p, q, dpq, rfl, h, hdpq⟩⟩ lemma Language.Theory.Derivation.allIntro {s p dp : V} (h : ^∀ p ∈ s) (hdp : T.DerivationOf dp (insert (L.free p) (L.setShift s))) : T.Derivation (allIntro s p dp) := Language.Theory.Derivation.mk - ⟨by simp; intro q hq; simpa using hdp.formulaSet (L.shift q) (by simp [shift_mem_setShift hq]), + ⟨by simp; intro q hq; simpa using hdp.isFormulaSet (L.shift q) (by simp [shift_mem_setShift hq]), Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨s, p, dp, rfl, h, hdp⟩⟩ lemma Language.Theory.Derivation.exIntro {s p t dp : V} - (h : ^∃ p ∈ s) (ht : L.Term t) + (h : ^∃ p ∈ s) (ht : L.IsTerm t) (hdp : T.DerivationOf dp (insert (L.substs₁ t p) s)) : T.Derivation (exIntro s p t dp) := Language.Theory.Derivation.mk - ⟨by simp; intro q hq; exact hdp.formulaSet q (by simp [hq]), + ⟨by simp; intro q hq; exact hdp.isFormulaSet q (by simp [hq]), Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨s, p, t, dp, rfl, h, ht, hdp⟩⟩ -lemma Language.Theory.Derivation.wkRule {s s' d : V} (hs : L.FormulaSet s) +lemma Language.Theory.Derivation.wkRule {s s' d : V} (hs : L.IsFormulaSet s) (h : s' ⊆ s) (hd : T.DerivationOf d s') : T.Derivation (wkRule s d) := Language.Theory.Derivation.mk ⟨by simpa using hs, @@ -698,7 +705,7 @@ lemma Language.Theory.Derivation.wkRule {s s' d : V} (hs : L.FormulaSet s) lemma Language.Theory.Derivation.shiftRule {s d : V} (hd : T.DerivationOf d s) : T.Derivation (shiftRule (L.setShift s) d) := Language.Theory.Derivation.mk - ⟨by simp [hd.formulaSet], + ⟨by simp [hd.isFormulaSet], Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨L.setShift s, d, rfl, by simp [hd.1], hd.2⟩⟩ lemma Language.Theory.Derivation.cutRule {s p d₁ d₂ : V} @@ -706,10 +713,10 @@ lemma Language.Theory.Derivation.cutRule {s p d₁ d₂ : V} (hd₂ : T.DerivationOf d₂ (insert (L.neg p) s)) : T.Derivation (cutRule s p d₁ d₂) := Language.Theory.Derivation.mk - ⟨by simp; intro q hq; exact hd₁.formulaSet q (by simp [hq]), + ⟨by simp; intro q hq; exact hd₁.isFormulaSet q (by simp [hq]), Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inr <| Or.inl ⟨s, p, d₁, d₂, rfl, hd₁, hd₂⟩⟩ -lemma Language.Theory.Derivation.root {s p : V} (hs : L.FormulaSet s) (hp : p ∈ s) (hT : p ∈ T) : +lemma Language.Theory.Derivation.root {s p : V} (hs : L.IsFormulaSet s) (hp : p ∈ s) (hT : p ∈ T) : T.Derivation (root s p) := Language.Theory.Derivation.mk ⟨by simpa using hs, @@ -717,13 +724,13 @@ lemma Language.Theory.Derivation.root {s p : V} (hs : L.FormulaSet s) (hp : p namespace Language.Theory.Derivable -lemma formulaSet {s : V} (h : T.Derivable s) : L.FormulaSet s := by - rcases h with ⟨d, hd⟩; exact hd.formulaSet +lemma isFormulaSet {s : V} (h : T.Derivable s) : L.IsFormulaSet s := by + rcases h with ⟨d, hd⟩; exact hd.isFormulaSet -lemma em {s : V} (hs : L.FormulaSet s) (p) (h : p ∈ s) (hn : L.neg p ∈ s) : +lemma em {s : V} (hs : L.IsFormulaSet s) (p) (h : p ∈ s) (hn : L.neg p ∈ s) : T.Derivable s := ⟨axL s p, by simp, Language.Theory.Derivation.axL hs h hn⟩ -lemma verum {s : V} (hs : L.FormulaSet s) (h : ^⊤ ∈ s) : +lemma verum {s : V} (hs : L.IsFormulaSet s) (h : ^⊤ ∈ s) : T.Derivable s := ⟨verumIntro s, by simp, Language.Theory.Derivation.verumIntro hs h⟩ lemma and_m {s p q : V} (h : p ^⋏ q ∈ s) (hp : T.Derivable (insert p s)) (hq : T.Derivable (insert q s)) : @@ -741,12 +748,12 @@ lemma all_m {s p : V} (h : ^∀ p ∈ s) (hp : T.Derivable (insert (L.free p) (L rcases hp with ⟨dp, hdp⟩ exact ⟨allIntro s p dp, by simp, Language.Theory.Derivation.allIntro h hdp⟩ -lemma ex_m {s p t : V} (h : ^∃ p ∈ s) (ht : L.Term t) (hp : T.Derivable (insert (L.substs₁ t p) s)) : +lemma ex_m {s p t : V} (h : ^∃ p ∈ s) (ht : L.IsTerm t) (hp : T.Derivable (insert (L.substs₁ t p) s)) : T.Derivable s := by rcases hp with ⟨dp, hdp⟩ exact ⟨exIntro s p t dp, by simp, Language.Theory.Derivation.exIntro h ht hdp⟩ -lemma wk {s s' : V} (hs : L.FormulaSet s) (h : s' ⊆ s) (hd : T.Derivable s') : +lemma wk {s s' : V} (hs : L.IsFormulaSet s) (h : s' ⊆ s) (hd : T.Derivable s') : T.Derivable s := by rcases hd with ⟨d, hd⟩ exact ⟨wkRule s d, by simp, Language.Theory.Derivation.wkRule hs h hd⟩ @@ -766,24 +773,24 @@ lemma cut {s : V} (p) (hd₁ : T.Derivable (insert p s)) (hd₂ : T.Derivable (i rcases hd₁ with ⟨d₁, hd₁⟩; rcases hd₂ with ⟨d₂, hd₂⟩ exact ⟨cutRule s p d₁ d₂, by simp, Language.Theory.Derivation.cutRule hd₁ hd₂⟩ -lemma by_axm {s : V} (hs : L.FormulaSet s) (p) (hp : p ∈ s) (hT : p ∈ T) : +lemma by_axm {s : V} (hs : L.IsFormulaSet s) (p) (hp : p ∈ s) (hT : p ∈ T) : T.Derivable s := by exact ⟨Arith.root s p, by simp, Language.Theory.Derivation.root hs hp hT⟩ lemma and {s p q : V} (hp : T.Derivable (insert p s)) (hq : T.Derivable (insert q s)) : T.Derivable (insert (p ^⋏ q) s) := and_m (p := p) (q := q) (by simp) - (wk (by simp [hp.formulaSet.insert, hq.formulaSet.insert]) (insert_subset_insert_of_subset _ <| by simp) hp) - (wk (by simp [hp.formulaSet.insert, hq.formulaSet.insert]) (insert_subset_insert_of_subset _ <| by simp) hq) + (wk (by simp [hp.isFormulaSet.insert, hq.isFormulaSet.insert]) (insert_subset_insert_of_subset _ <| by simp) hp) + (wk (by simp [hp.isFormulaSet.insert, hq.isFormulaSet.insert]) (insert_subset_insert_of_subset _ <| by simp) hq) lemma or {s p q : V} (hpq : T.Derivable (insert p (insert q s))) : T.Derivable (insert (p ^⋎ q) s) := or_m (p := p) (q := q) (by simp) - (wk (by simp [hpq.formulaSet.insert, hpq.formulaSet.insert.2.insert]) + (wk (by simp [hpq.isFormulaSet.insert, hpq.isFormulaSet.insert.2.insert]) (insert_subset_insert_of_subset _ <| insert_subset_insert_of_subset _ <| by simp) hpq) /-- Crucial inducion for formalized $\Sigma_1$-completeness. -/ -lemma conj (ps : V) {s} (hs : L.FormulaSet s) +lemma conj (ps : V) {s} (hs : L.IsFormulaSet s) (ds : ∀ i < len ps, T.Derivable (insert ps.[i] s)) : T.Derivable (insert (^⋀ ps) s) := by have : ∀ k ≤ len ps, T.Derivable (insert (^⋀ (takeLast ps k)) s) := by intro k hk @@ -810,9 +817,9 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( case zero => intro s' _ ss hs' refine wk ?_ ?_ d - · simp [by simpa using d.formulaSet] + · simp [by simpa using d.isFormulaSet] intro x hx - exact d.formulaSet x (by simp [ss hx]) + exact d.isFormulaSet x (by simp [ss hx]) · intro x simp only [mem_cup_iff, mem_vecToSet_iff, takeLast_zero, qqDisj_nil, mem_bitInsert_iff] rintro (⟨i, hi, rfl⟩ | hx) @@ -839,22 +846,22 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( exact ofSetEq (by intro x; simp [s'']; tauto) this simpa using this (len ps) (by rfl) ∅ (by simp [emptyset_def]) (by simp) (by simp) -lemma disj (ps s : V) {i} (hps : ∀ i < len ps, L.Formula ps.[i]) +lemma disj (ps s : V) {i} (hps : ∀ i < len ps, L.IsFormula ps.[i]) (hi : i < len ps) (d : T.Derivable (insert ps.[i] s)) : T.Derivable (insert (^⋁ ps) s) := disjDistr ps s <| wk - (by simp [by simpa using d.formulaSet]; intro x hx; rcases mem_vecToSet_iff.mp hx with ⟨i, hi, rfl⟩; exact hps i hi) + (by simp [by simpa using d.isFormulaSet]; intro x hx; rcases mem_vecToSet_iff.mp hx with ⟨i, hi, rfl⟩; exact hps i hi) (by intro x; simp only [mem_bitInsert_iff, mem_cup_iff] rintro (rfl | hx) · left; exact mem_vecToSet_iff.mpr ⟨i, hi, rfl⟩ · right; exact hx) d -lemma all {p : V} (hp : L.Semiformula 1 p) (dp : T.Derivable (insert (L.free p) (L.setShift s))) : T.Derivable (insert (^∀ p) s) := - all_m (p := p) (by simp) (wk (by simp [hp, by simpa using dp.formulaSet]) (by intro x; simp; tauto) dp) +lemma all {p : V} (hp : L.IsSemiformula 1 p) (dp : T.Derivable (insert (L.free p) (L.setShift s))) : T.Derivable (insert (^∀ p) s) := + all_m (p := p) (by simp) (wk (by simp [hp, by simpa using dp.isFormulaSet]) (by intro x; simp; tauto) dp) -lemma ex {p t : V} (hp : L.Semiformula 1 p) (ht : L.Term t) +lemma ex {p t : V} (hp : L.IsSemiformula 1 p) (ht : L.IsTerm t) (dp : T.Derivable (insert (L.substs₁ t p) s)) : T.Derivable (insert (^∃ p) s) := - ex_m (p := p) (by simp) ht (wk (by simp [ht, hp, by simpa using dp.formulaSet]) (by intro x; simp; tauto) dp) + ex_m (p := p) (by simp) ht (wk (by simp [ht, hp, by simpa using dp.isFormulaSet]) (by intro x; simp; tauto) dp) end Language.Theory.Derivable diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean index 6ba0082..c80cd6b 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean @@ -20,30 +20,30 @@ variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] section typed_formula -abbrev Language.TSemiformula.substs₁ (p : L.TSemiformula (0 + 1)) (t : L.TTerm) : L.TFormula := p.substs t.sing +abbrev Language.Semiformula.substs₁ (p : L.Semiformula (0 + 1)) (t : L.Term) : L.Formula := p.substs t.sing -abbrev Language.TSemiformula.free (p : L.TSemiformula (0 + 1)) : L.TFormula := p.shift.substs₁ (L.fvar 0) +abbrev Language.Semiformula.free (p : L.Semiformula (0 + 1)) : L.Formula := p.shift.substs₁ (L.fvar 0) -@[simp] lemma Language.TSemiformula.val_substs₁ (p : L.TSemiformula (0 + 1)) (t : L.TTerm) : - (p.substs₁ t).val = L.substs 0 ?[t.val] p.val := by simp [substs₁, substs] +@[simp] lemma Language.Semiformula.val_substs₁ (p : L.Semiformula (0 + 1)) (t : L.Term) : + (p.substs₁ t).val = L.substs ?[t.val] p.val := by simp [substs₁, substs] -@[simp] lemma Language.TSemiformula.val_free (p : L.TSemiformula (0 + 1)) : - p.free.val = L.substs 0 ?[^&0] (L.shift p.val) := by simp [free, substs₁, substs, shift, fvar] +@[simp] lemma Language.Semiformula.val_free (p : L.Semiformula (0 + 1)) : + p.free.val = L.substs ?[^&0] (L.shift p.val) := by simp [free, substs₁, substs, shift, fvar] -@[simp] lemma substs₁_neg (p : L.TSemiformula (0 + 1)) (t : L.TTerm) : - (~p).substs₁ t = ~(p.substs₁ t) := by simp [Language.TSemiformula.substs₁] +@[simp] lemma substs₁_neg (p : L.Semiformula (0 + 1)) (t : L.Term) : + (~p).substs₁ t = ~(p.substs₁ t) := by simp [Language.Semiformula.substs₁] -@[simp] lemma substs₁_all (p : L.TSemiformula (0 + 1 + 1)) (t : L.TTerm) : - p.all.substs₁ t = (p.substs t.sing.q).all := by simp [Language.TSemiformula.substs₁] +@[simp] lemma substs₁_all (p : L.Semiformula (0 + 1 + 1)) (t : L.Term) : + p.all.substs₁ t = (p.substs t.sing.q).all := by simp [Language.Semiformula.substs₁] -@[simp] lemma substs₁_ex (p : L.TSemiformula (0 + 1 + 1)) (t : L.TTerm) : - p.ex.substs₁ t = (p.substs t.sing.q).ex := by simp [Language.TSemiformula.substs₁] +@[simp] lemma substs₁_ex (p : L.Semiformula (0 + 1 + 1)) (t : L.Term) : + p.ex.substs₁ t = (p.substs t.sing.q).ex := by simp [Language.Semiformula.substs₁] end typed_formula section typed_theory -abbrev Language.Theory.tmem (p : L.TFormula) (T : L.Theory) : Prop := p.val ∈ T +abbrev Language.Theory.tmem (p : L.Formula) (T : L.Theory) : Prop := p.val ∈ T scoped infix:50 " ∈' " => Language.Theory.tmem @@ -55,7 +55,7 @@ variable (L) structure Language.Sequent where val : V - val_formulaSet : L.FormulaSet val + val_formulaSet : L.IsFormulaSet val attribute [simp] Language.Sequent.val_formulaSet @@ -63,13 +63,13 @@ variable {L} instance : EmptyCollection L.Sequent := ⟨⟨∅, by simp⟩⟩ -instance : Singleton L.TFormula L.Sequent := ⟨fun p ↦ ⟨{p.val}, by simp⟩⟩ +instance : Singleton L.Formula L.Sequent := ⟨fun p ↦ ⟨{p.val}, by simp⟩⟩ -instance : Insert L.TFormula L.Sequent := ⟨fun p Γ ↦ ⟨insert p.val Γ.val, by simp⟩⟩ +instance : Insert L.Formula L.Sequent := ⟨fun p Γ ↦ ⟨insert p.val Γ.val, by simp⟩⟩ instance : Union L.Sequent := ⟨fun Γ Δ ↦ ⟨Γ.val ∪ Δ.val, by simp⟩⟩ -instance : Membership L.TFormula L.Sequent := ⟨(·.val ∈ ·.val)⟩ +instance : Membership L.Formula L.Sequent := ⟨(·.val ∈ ·.val)⟩ instance : HasSubset L.Sequent := ⟨(·.val ⊆ ·.val)⟩ @@ -77,7 +77,7 @@ scoped infixr:50 " ⫽ " => Insert.insert namespace Language.Sequent -variable {Γ Δ : L.Sequent} {p q : L.TFormula} +variable {Γ Δ : L.Sequent} {p q : L.Formula} lemma mem_iff : p ∈ Γ ↔ p.val ∈ Γ.val := iff_of_eq rfl @@ -85,19 +85,19 @@ lemma subset_iff : Γ ⊆ Δ ↔ Γ.val ⊆ Δ.val := iff_of_eq rfl @[simp] lemma val_empty : (∅ : L.Sequent).val = ∅ := rfl -@[simp] lemma val_singleton (p : L.TFormula) : ({p} : L.Sequent).val = {p.val} := rfl +@[simp] lemma val_singleton (p : L.Formula) : ({p} : L.Sequent).val = {p.val} := rfl -@[simp] lemma val_insert (p : L.TFormula) (Γ : L.Sequent) : (insert p Γ).val = insert p.val Γ.val := rfl +@[simp] lemma val_insert (p : L.Formula) (Γ : L.Sequent) : (insert p Γ).val = insert p.val Γ.val := rfl @[simp] lemma val_union (Γ Δ : L.Sequent) : (Γ ∪ Δ).val = Γ.val ∪ Δ.val := rfl -@[simp] lemma not_mem_empty (p : L.TFormula) : p ∉ (∅ : L.Sequent) := by simp [mem_iff] +@[simp] lemma not_mem_empty (p : L.Formula) : p ∉ (∅ : L.Sequent) := by simp [mem_iff] -@[simp] lemma mem_singleton_iff : p ∈ ({q} : L.Sequent) ↔ p = q := by simp [mem_iff, Language.TSemiformula.val_inj] +@[simp] lemma mem_singleton_iff : p ∈ ({q} : L.Sequent) ↔ p = q := by simp [mem_iff, Language.Semiformula.val_inj] -@[simp] lemma mem_insert_iff : p ∈ insert q Γ ↔ p = q ∨ p ∈ Γ := by simp [mem_iff, Language.TSemiformula.val_inj] +@[simp] lemma mem_insert_iff : p ∈ insert q Γ ↔ p = q ∨ p ∈ Γ := by simp [mem_iff, Language.Semiformula.val_inj] -@[simp] lemma mem_union_iff : p ∈ Γ ∪ Δ ↔ p ∈ Γ ∨ p ∈ Δ := by simp [mem_iff, Language.TSemiformula.val_inj] +@[simp] lemma mem_union_iff : p ∈ Γ ∪ Δ ↔ p ∈ Γ ∨ p ∈ Δ := by simp [mem_iff, Language.Semiformula.val_inj] @[ext] lemma ext (h : ∀ x, x ∈ Γ ↔ x ∈ Δ) : Γ = Δ := by rcases Γ with ⟨Γ, hΓ⟩; rcases Δ with ⟨Δ, hΔ⟩; simp @@ -137,9 +137,9 @@ structure Language.Theory.TDerivation (T : Language.TTheory L) (Γ : L.Sequent) scoped infix:45 " ⊢¹ " => Language.Theory.TDerivation -def Language.Theory.TProof (T : Language.TTheory L) (p : L.TFormula) := T ⊢¹ insert p ∅ +def Language.Theory.TProof (T : Language.TTheory L) (p : L.Formula) := T ⊢¹ insert p ∅ -instance : System L.TFormula L.TTheory := ⟨Language.Theory.TProof⟩ +instance : System L.Formula L.TTheory := ⟨Language.Theory.TProof⟩ variable {T : L.TTheory} @@ -150,7 +150,7 @@ def Language.Theory.Derivable.toTDerivation (Γ : L.Sequent) (h : T.thy.Derivabl lemma Language.Theory.TDerivation.toDerivable {Γ : L.Sequent} (d : T ⊢¹ Γ) : T.thy.Derivable Γ.val := ⟨d.derivation, d.derivationOf⟩ -lemma Language.Theory.TProvable.iff_provable {σ : L.TFormula} : +lemma Language.Theory.TProvable.iff_provable {σ : L.Formula} : T ⊢! σ ↔ T.thy.Provable σ.val := by constructor · intro b @@ -164,7 +164,7 @@ def Language.Theory.TProof.toTDerivation {p} (d : T ⊢ p) : T ⊢¹ insert p namespace Language.Theory.TDerivation -variable {Γ Δ : L.Sequent} {p q p₀ p₁ p₂ p₃ p₄ : L.TFormula} +variable {Γ Δ : L.Sequent} {p q p₀ p₁ p₂ p₃ p₄ : L.Formula} def byAxm (p) (h : p ∈' T.thy) (hΓ : p ∈ Γ) : T ⊢¹ Γ := Language.Theory.Derivable.toTDerivation _ @@ -185,11 +185,11 @@ def and (dp : T ⊢¹ insert p Γ) (dq : T ⊢¹ insert q Γ) : T ⊢¹ insert ( def or (dpq : T ⊢¹ insert p (insert q Γ)) : T ⊢¹ insert (p ⋎ q) Γ := Language.Theory.Derivable.toTDerivation _ <| by simpa using Language.Theory.Derivable.or (by simpa using dpq.toDerivable) -def all {p : L.TSemiformula (0 + 1)} (dp : T ⊢¹ insert p.free Γ.shift) : T ⊢¹ insert p.all Γ := +def all {p : L.Semiformula (0 + 1)} (dp : T ⊢¹ insert p.free Γ.shift) : T ⊢¹ insert p.all Γ := Language.Theory.Derivable.toTDerivation _ <| by simpa using Language.Theory.Derivable.all (by simpa using p.prop) (by simpa using dp.toDerivable) -def ex {p : L.TSemiformula (0 + 1)} (t : L.TTerm) (dp : T ⊢¹ insert (p.substs₁ t) Γ) : T ⊢¹ insert p.ex Γ := +def ex {p : L.Semiformula (0 + 1)} (t : L.Term) (dp : T ⊢¹ insert (p.substs₁ t) Γ) : T ⊢¹ insert p.ex Γ := Language.Theory.Derivable.toTDerivation _ <| by simpa using Language.Theory.Derivable.ex (by simpa using p.prop) t.prop (by simpa using dp.toDerivable) @@ -208,12 +208,12 @@ def cut (d₁ : T ⊢¹ insert p Γ) (d₂ : T ⊢¹ insert (~p) Γ) : T ⊢¹ def cut' (d₁ : T ⊢¹ insert p Γ) (d₂ : T ⊢¹ insert (~p) Δ) : T ⊢¹ Γ ∪ Δ := cut (p := p) (d₁.wk (by intro x; simp; tauto)) (d₂.wk (by intro x; simp; tauto)) -def conj (ps : L.TSemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢¹ insert (ps.nth i hi) Γ) : T ⊢¹ insert ps.conj Γ := by +def conj (ps : L.SemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢¹ insert (ps.nth i hi) Γ) : T ⊢¹ insert ps.conj Γ := by have : ∀ i < len ps.val, T.thy.Derivable (insert (ps.val.[i]) Γ.val) := by intro i hi; simpa using (ds i hi).toDerivable have : T.thy.Derivable (insert (^⋀ ps.val) Γ.val) := Language.Theory.Derivable.conj ps.val (by simp) this exact Language.Theory.Derivable.toTDerivation _ (by simpa using this) -def disj (ps : L.TSemiformulaVec 0) {i} (hi : i < len ps.val) +def disj (ps : L.SemiformulaVec 0) {i} (hi : i < len ps.val) (d : T ⊢¹ insert (ps.nth i hi) Γ) : T ⊢¹ insert ps.disj Γ := by have : T.thy.Derivable (insert (^⋁ ps.val) Γ.val) := Language.Theory.Derivable.disj ps.val Γ.val ps.prop hi (by simpa using d.toDerivable) @@ -222,7 +222,7 @@ def disj (ps : L.TSemiformulaVec 0) {i} (hi : i < len ps.val) def modusPonens (dpq : T ⊢¹ insert (p ⟶ q) Γ) (dp : T ⊢¹ insert p Γ) : T ⊢¹ insert q Γ := by let d : T ⊢¹ insert (p ⟶ q) (insert q Γ) := dpq.wk (insert_subset_insert_of_subset _ <| by simp) let b : T ⊢¹ insert (~(p ⟶ q)) (insert q Γ) := by - simp only [TSemiformula.imp_def, TSemiformula.neg_or, TSemiformula.neg_neg] + simp only [Semiformula.imp_def, Semiformula.neg_or, Semiformula.neg_neg] exact and (dp.wk (insert_subset_insert_of_subset _ <| by simp)) (em q (by simp) (by simp)) exact cut d b @@ -241,14 +241,14 @@ def rotate₃ (d : T ⊢¹ p₀ ⫽ p₁ ⫽ p₂ ⫽ p₃ ⫽ Γ) : T ⊢¹ p def orInv (d : T ⊢¹ p ⋎ q ⫽ Γ) : T ⊢¹ p ⫽ q ⫽ Γ := by have b : T ⊢¹ p ⋎ q ⫽ p ⫽ q ⫽ Γ := wk d (by intro x; simp; tauto) have : T ⊢¹ ~(p ⋎ q) ⫽ p ⫽ q ⫽ Γ := by - simp only [TSemiformula.neg_or] + simp only [Semiformula.neg_or] apply and (em p) (em q) exact cut b this -def specialize {p : L.TSemiformula (0 + 1)} (b : T ⊢¹ p.all ⫽ Γ) (t : L.TTerm) : T ⊢¹ p.substs₁ t ⫽ Γ := by +def specialize {p : L.Semiformula (0 + 1)} (b : T ⊢¹ p.all ⫽ Γ) (t : L.Term) : T ⊢¹ p.substs₁ t ⫽ Γ := by apply TDerivation.cut (p := p.all) · exact (TDerivation.wk b <| by intro x; simp; tauto) - · rw [TSemiformula.neg_all] + · rw [Semiformula.neg_all] apply TDerivation.ex t apply TDerivation.em (p.substs₁ t) @@ -256,18 +256,18 @@ end Language.Theory.TDerivation namespace Language.Theory.TProof -variable {T : L.TTheory} {p q : L.TFormula} +variable {T : L.TTheory} {p q : L.Formula} /-- Condition D2 -/ def modusPonens (d : T ⊢ p ⟶ q) (b : T ⊢ p) : T ⊢ q := TDerivation.modusPonens d b -def byAxm {p : L.TFormula} (h : p ∈' T.thy) : T ⊢ p := TDerivation.byAxm p h (by simp) +def byAxm {p : L.Formula} (h : p ∈' T.thy) : T ⊢ p := TDerivation.byAxm p h (by simp) instance : System.ModusPonens T := ⟨modusPonens⟩ instance : System.NegationEquiv T where neg_equiv p := by - simp [Axioms.NegEquiv, LO.LogicalConnective.iff, TSemiformula.imp_def] + simp [Axioms.NegEquiv, LO.LogicalConnective.iff, Semiformula.imp_def] apply TDerivation.and · apply TDerivation.or apply TDerivation.rotate₁ @@ -281,13 +281,13 @@ instance : System.NegationEquiv T where instance : System.Minimal T where verum := TDerivation.toTProof <| TDerivation.verum imply₁ (p q) := by - simp only [Axioms.Imply₁, TSemiformula.imp_def] + simp only [Axioms.Imply₁, Semiformula.imp_def] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or exact TDerivation.em p imply₂ (p q r) := by - simp only [Axioms.Imply₂, TSemiformula.imp_def, TSemiformula.neg_or, TSemiformula.neg_neg] + simp only [Axioms.Imply₂, Semiformula.imp_def, Semiformula.neg_or, Semiformula.neg_neg] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or @@ -303,17 +303,17 @@ instance : System.Minimal T where · exact TDerivation.em q · exact TDerivation.em r and₁ (p q) := by - simp only [Axioms.AndElim₁, TSemiformula.imp_def, TSemiformula.neg_and] + simp only [Axioms.AndElim₁, Semiformula.imp_def, Semiformula.neg_and] apply TDerivation.or apply TDerivation.or exact TDerivation.em p and₂ (p q) := by - simp only [Axioms.AndElim₂, TSemiformula.imp_def, TSemiformula.neg_and] + simp only [Axioms.AndElim₂, Semiformula.imp_def, Semiformula.neg_and] apply TDerivation.or apply TDerivation.or exact TDerivation.em q and₃ (p q) := by - simp only [Axioms.AndInst, TSemiformula.imp_def] + simp only [Axioms.AndInst, Semiformula.imp_def] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or @@ -322,19 +322,19 @@ instance : System.Minimal T where · exact TDerivation.em p · exact TDerivation.em q or₁ (p q) := by - simp only [Axioms.OrInst₁, TSemiformula.imp_def] + simp only [Axioms.OrInst₁, Semiformula.imp_def] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or exact TDerivation.em p or₂ (p q) := by - simp [Axioms.OrInst₂, TSemiformula.imp_def] + simp [Axioms.OrInst₂, Semiformula.imp_def] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or exact TDerivation.em q or₃ (p q r) := by - simp [Axioms.OrElim, TSemiformula.imp_def] + simp [Axioms.OrElim, Semiformula.imp_def] apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.or @@ -352,76 +352,76 @@ instance : System.Minimal T where instance : System.Classical T where dne p := by - simp [Axioms.DNE, TSemiformula.imp_def] + simp [Axioms.DNE, Semiformula.imp_def] apply TDerivation.or exact TDerivation.em p -def exIntro (p : L.TSemiformula (0 + 1)) (t : L.TTerm) (b : T ⊢ p.substs₁ t) : T ⊢ p.ex := TDerivation.ex t b +def exIntro (p : L.Semiformula (0 + 1)) (t : L.Term) (b : T ⊢ p.substs₁ t) : T ⊢ p.ex := TDerivation.ex t b -lemma ex_intro! (p : L.TSemiformula (0 + 1)) (t : L.TTerm) (b : T ⊢! p.substs₁ t) : T ⊢! p.ex := ⟨exIntro _ t b.get⟩ +lemma ex_intro! (p : L.Semiformula (0 + 1)) (t : L.Term) (b : T ⊢! p.substs₁ t) : T ⊢! p.ex := ⟨exIntro _ t b.get⟩ -def specialize {p : L.TSemiformula (0 + 1)} (b : T ⊢ p.all) (t : L.TTerm) : T ⊢ p.substs₁ t := TDerivation.specialize b t +def specialize {p : L.Semiformula (0 + 1)} (b : T ⊢ p.all) (t : L.Term) : T ⊢ p.substs₁ t := TDerivation.specialize b t -lemma specialize! {p : L.TSemiformula (0 + 1)} (b : T ⊢! p.all) (t : L.TTerm) : T ⊢! p.substs₁ t := ⟨TDerivation.specialize b.get t⟩ +lemma specialize! {p : L.Semiformula (0 + 1)} (b : T ⊢! p.all) (t : L.Term) : T ⊢! p.substs₁ t := ⟨TDerivation.specialize b.get t⟩ -def conj (ps : L.TSemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth i hi) : T ⊢ ps.conj := TDerivation.conj ps ds +def conj (ps : L.SemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth i hi) : T ⊢ ps.conj := TDerivation.conj ps ds -lemma conj! (ps : L.TSemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢! ps.nth i hi) : T ⊢! ps.conj := ⟨conj ps fun i hi ↦ (ds i hi).get⟩ +lemma conj! (ps : L.SemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢! ps.nth i hi) : T ⊢! ps.conj := ⟨conj ps fun i hi ↦ (ds i hi).get⟩ -def conj' (ps : L.TSemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth (len ps.val - (i + 1)) (sub_succ_lt_self hi)) : T ⊢ ps.conj := +def conj' (ps : L.SemiformulaVec 0) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth (len ps.val - (i + 1)) (sub_succ_lt_self hi)) : T ⊢ ps.conj := TDerivation.conj ps <| fun i hi ↦ by simpa [sub_succ_lt_selfs hi] using ds (len ps.val - (i + 1)) (by simp [tsub_lt_iff_left (succ_le_iff_lt.mpr hi)]) -def conjOr' (ps : L.TSemiformulaVec 0) (q) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth (len ps.val - (i + 1)) (sub_succ_lt_self hi) ⋎ q) : T ⊢ ps.conj ⋎ q := +def conjOr' (ps : L.SemiformulaVec 0) (q) (ds : ∀ i, (hi : i < len ps.val) → T ⊢ ps.nth (len ps.val - (i + 1)) (sub_succ_lt_self hi) ⋎ q) : T ⊢ ps.conj ⋎ q := TDerivation.or <| TDerivation.conj ps <| fun i hi ↦ by simpa [sub_succ_lt_selfs hi] using TDerivation.orInv (ds (len ps.val - (i + 1)) (by simp [tsub_lt_iff_left (succ_le_iff_lt.mpr hi)])) -def disj (ps : L.TSemiformulaVec 0) {i} (hi : i < len ps.val) (d : T ⊢ ps.nth i hi) : T ⊢ ps.disj := +def disj (ps : L.SemiformulaVec 0) {i} (hi : i < len ps.val) (d : T ⊢ ps.nth i hi) : T ⊢ ps.disj := TDerivation.disj ps hi d -def shift {p : L.TFormula} (d : T ⊢ p) : T ⊢ p.shift := by simpa using TDerivation.shift d +def shift {p : L.Formula} (d : T ⊢ p) : T ⊢ p.shift := by simpa using TDerivation.shift d -lemma shift! {p : L.TFormula} (d : T ⊢! p) : T ⊢! p.shift := ⟨by simpa using TDerivation.shift d.get⟩ +lemma shift! {p : L.Formula} (d : T ⊢! p) : T ⊢! p.shift := ⟨by simpa using TDerivation.shift d.get⟩ -def all {p : L.TSemiformula (0 + 1)} (dp : T ⊢ p.free) : T ⊢ p.all := TDerivation.all (by simpa using dp) +def all {p : L.Semiformula (0 + 1)} (dp : T ⊢ p.free) : T ⊢ p.all := TDerivation.all (by simpa using dp) -lemma all! {p : L.TSemiformula (0 + 1)} (dp : T ⊢! p.free) : T ⊢! p.all := ⟨all dp.get⟩ +lemma all! {p : L.Semiformula (0 + 1)} (dp : T ⊢! p.free) : T ⊢! p.all := ⟨all dp.get⟩ -def generalizeAux {C : L.TFormula} {p : L.TSemiformula (0 + 1)} (dp : T ⊢ C.shift ⟶ p.free) : T ⊢ C ⟶ p.all := by - rw [TSemiformula.imp_def] at dp ⊢ +def generalizeAux {C : L.Formula} {p : L.Semiformula (0 + 1)} (dp : T ⊢ C.shift ⟶ p.free) : T ⊢ C ⟶ p.all := by + rw [Semiformula.imp_def] at dp ⊢ apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.all exact TDerivation.wk (TDerivation.orInv dp) (by intro x; simp; tauto) -lemma conj_shift (Γ : List L.TFormula) : (⋀Γ).shift = ⋀(Γ.map .shift) := by +lemma conj_shift (Γ : List L.Formula) : (⋀Γ).shift = ⋀(Γ.map .shift) := by induction Γ using List.induction_with_singleton case hnil => simp case hsingle => simp [List.conj₂] case hcons p ps hps ih => simp [hps, ih] -def generalize {Γ} {p : L.TSemiformula (0 + 1)} (d : Γ.map .shift ⊢[T] p.free) : Γ ⊢[T] p.all := by +def generalize {Γ} {p : L.Semiformula (0 + 1)} (d : Γ.map .shift ⊢[T] p.free) : Γ ⊢[T] p.all := by apply System.FiniteContext.ofDef apply generalizeAux simpa [conj_shift] using System.FiniteContext.toDef d -lemma generalize! {Γ} {p : L.TSemiformula (0 + 1)} (d : Γ.map .shift ⊢[T]! p.free) : Γ ⊢[T]! p.all := ⟨generalize d.get⟩ +lemma generalize! {Γ} {p : L.Semiformula (0 + 1)} (d : Γ.map .shift ⊢[T]! p.free) : Γ ⊢[T]! p.all := ⟨generalize d.get⟩ -def specializeWithCtxAux {C : L.TFormula} {p : L.TSemiformula (0 + 1)} (d : T ⊢ C ⟶ p.all) (t : L.TTerm) : T ⊢ C ⟶ p.substs₁ t := by - rw [TSemiformula.imp_def] at d ⊢ +def specializeWithCtxAux {C : L.Formula} {p : L.Semiformula (0 + 1)} (d : T ⊢ C ⟶ p.all) (t : L.Term) : T ⊢ C ⟶ p.substs₁ t := by + rw [Semiformula.imp_def] at d ⊢ apply TDerivation.or apply TDerivation.rotate₁ apply TDerivation.specialize exact TDerivation.wk (TDerivation.orInv d) (by intro x; simp; tauto) -def specializeWithCtx {Γ} {p : L.TSemiformula (0 + 1)} (d : Γ ⊢[T] p.all) (t) : Γ ⊢[T] p.substs₁ t := specializeWithCtxAux d t +def specializeWithCtx {Γ} {p : L.Semiformula (0 + 1)} (d : Γ ⊢[T] p.all) (t) : Γ ⊢[T] p.substs₁ t := specializeWithCtxAux d t -lemma specialize_with_ctx! {Γ} {p : L.TSemiformula (0 + 1)} (d : Γ ⊢[T]! p.all) (t) : Γ ⊢[T]! p.substs₁ t := ⟨specializeWithCtx d.get t⟩ +lemma specialize_with_ctx! {Γ} {p : L.Semiformula (0 + 1)} (d : Γ ⊢[T]! p.all) (t) : Γ ⊢[T]! p.substs₁ t := ⟨specializeWithCtx d.get t⟩ -def ex {p : L.TSemiformula (0 + 1)} (t) (dp : T ⊢ p.substs₁ t) : T ⊢ p.ex := TDerivation.ex t (by simpa using dp) +def ex {p : L.Semiformula (0 + 1)} (t) (dp : T ⊢ p.substs₁ t) : T ⊢ p.ex := TDerivation.ex t (by simpa using dp) -lemma ex! {p : L.TSemiformula (0 + 1)} (t) (dp : T ⊢! p.substs₁ t) : T ⊢! p.ex := ⟨ex t dp.get⟩ +lemma ex! {p : L.Semiformula (0 + 1)} (t) (dp : T ⊢! p.substs₁ t) : T ⊢! p.ex := ⟨ex t dp.get⟩ end Language.Theory.TProof diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean index b7890b0..04f3919 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean @@ -90,47 +90,47 @@ namespace FormalizedTerm variable (L) -def Phi (n : V) (C : Set V) (t : V) : Prop := - (∃ z < n, t = ^#z) ∨ (∃ x, t = ^&x) ∨ (∃ k f v : V, L.Func k f ∧ k = len v ∧ (∀ i < k, v.[i] ∈ C) ∧ t = ^func k f v) +def Phi (C : Set V) (t : V) : Prop := + (∃ z, t = ^#z) ∨ (∃ x, t = ^&x) ∨ (∃ k f v : V, L.Func k f ∧ k = len v ∧ (∀ i < k, v.[i] ∈ C) ∧ t = ^func k f v) -private lemma phi_iff (n : V) (C : V) (t : V) : - Phi L n {x | x ∈ C} t ↔ - (∃ z < n, t = ^#z) ∨ +private lemma phi_iff (C : V) (t : V) : + Phi L {x | x ∈ C} t ↔ + (∃ z < t, t = ^#z) ∨ (∃ x < t, t = ^&x) ∨ (∃ k < t, ∃ f < t, ∃ v < t, L.Func k f ∧ k = len v ∧ (∀ i < k, v.[i] ∈ C) ∧ t = ^func k f v) where mp := by - rintro (⟨z, hz, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hkf, hk, hv, rfl⟩) - · left; exact ⟨z, hz, rfl⟩ + rintro (⟨z, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hkf, hk, hv, rfl⟩) + · left; exact ⟨z, lt_succ_iff_le.mpr <| by simp, rfl⟩ · right; left exact ⟨x, lt_succ_iff_le.mpr <| by simp, rfl⟩ · right; right exact ⟨k, by simp, f, by simp, v, by simp, hkf, hk, hv, rfl⟩ mpr := by unfold Phi - rintro (⟨z, hz, rfl⟩ | ⟨x, _, rfl⟩ | ⟨k, _, f, _, v, _, hkf, hk, hv, rfl⟩) - · left; exact ⟨z, hz, rfl⟩ + rintro (⟨z, _, rfl⟩ | ⟨x, _, rfl⟩ | ⟨k, _, f, _, v, _, hkf, hk, hv, rfl⟩) + · left; exact ⟨z, rfl⟩ · right; left; exact ⟨x, rfl⟩ · right; right; exact ⟨k, f, v, hkf, hk, hv, rfl⟩ -def blueprint (pL : LDef) : Fixpoint.Blueprint 1 where +def blueprint (pL : LDef) : Fixpoint.Blueprint 0 where core := .mkDelta - (.mkSigma “t C n | - (∃ z < n, !qqBvarDef t z) ∨ + (.mkSigma “t C | + (∃ z < t, !qqBvarDef t z) ∨ (∃ x < t, !qqFvarDef t x) ∨ (∃ k < t, ∃ f < t, ∃ v < t, !pL.func k f ∧ !lenDef k v ∧ (∀ i < k, ∃ u, !nthDef u v i ∧ u ∈ C) ∧ !qqFuncDef t k f v)” (by simp)) - (.mkPi “t C n | - (∃ z < n, !qqBvarDef t z) ∨ + (.mkPi “t C | + (∃ z < t, !qqBvarDef t z) ∨ (∃ x < t, !qqFvarDef t x) ∨ (∃ k < t, ∃ f < t, ∃ v < t, !pL.func k f ∧ (∀ l, !lenDef l v → k = l) ∧ (∀ i < k, ∀ u, !nthDef u v i → u ∈ C) ∧ !qqFuncDef t k f v)” (by simp)) def construction : Fixpoint.Construction V (blueprint pL) where - Φ := fun n ↦ Phi L (n 0) + Φ := fun _ ↦ Phi L defined := ⟨by intro v; simp [blueprint], by intro v; simp [blueprint, phi_iff, Language.Defined.eval_func (L := L)]⟩ monotone := by - rintro C C' hC v x (h | h | ⟨k, f, v, hkf, hk, h, rfl⟩) + rintro C C' hC _ x (h | h | ⟨k, f, v, hkf, hk, h, rfl⟩) · exact Or.inl h · exact Or.inr <| Or.inl h · exact Or.inr <| Or.inr ⟨k, f, v, hkf, hk, fun i hi ↦ hC (h i hi), rfl⟩ @@ -149,132 +149,120 @@ open FormalizedTerm variable (L) -def Language.Semiterm (n : V) : V → Prop := (construction L).Fixpoint ![n] +def Language.IsUTerm : V → Prop := (construction L).Fixpoint ![] -abbrev Language.Term : V → Prop := L.Semiterm 0 +def _root_.LO.FirstOrder.Arith.LDef.isUTermDef (pL : LDef) : 𝚫₁.Semisentence 1 := (blueprint pL).fixpointDefΔ₁ -def _root_.LO.FirstOrder.Arith.LDef.isSemitermDef (pL : LDef) : 𝚫₁.Semisentence 2 := (blueprint pL).fixpointDefΔ₁.rew (Rew.substs ![#1, #0]) +lemma Language.isUTerm_defined : 𝚫₁-Predicate L.IsUTerm via pL.isUTermDef := (construction L).fixpoint_definedΔ₁ -lemma isSemiterm_defined : 𝚫₁-Relation L.Semiterm via pL.isSemitermDef := - ⟨HierarchySymbol.Semiformula.ProperOn.rew (construction L).fixpoint_definedΔ₁.proper _, - by intro v; simp [LDef.isSemitermDef, (construction L).eval_fixpointDefΔ₁]; rfl⟩ +@[simp] lemma eval_isUTermDef (v) : + Semiformula.Evalbm V v pL.isUTermDef.val ↔ L.IsUTerm (v 0) := L.isUTerm_defined.df.iff v -@[simp] lemma eval_isSemitermDef (v) : - Semiformula.Evalbm V v pL.isSemitermDef.val ↔ L.Semiterm (v 0) (v 1) := (isSemiterm_defined L).df.iff v +instance Language.isUTerm_definable : 𝚫₁-Predicate L.IsUTerm := L.isUTerm_defined.to_definable -instance isSemitermDef_definable : 𝚫₁-Relation (L.Semiterm) := (isSemiterm_defined L).to_definable +instance isUTermDef_definable' (Γ) : Γ-[m + 1]-Predicate L.IsUTerm := L.isUTerm_definable.of_deltaOne -@[simp, definability] instance isSemitermDef_definable' (Γ) : Γ-[m + 1]-Relation (L.Semiterm) := - .of_deltaOne (isSemitermDef_definable L) _ _ - -def Language.SemitermVec (n m w : V) : Prop := n = len w ∧ ∀ i < n, L.Semiterm m (w.[i]) +def Language.IsUTermVec (n w : V) : Prop := n = len w ∧ ∀ i < n, L.IsUTerm w.[i] variable {L} -protected lemma Language.SemitermVec.lh {n m w : V} (h : L.SemitermVec n m w) : n = len w := h.1 +protected lemma Language.IsUTermVec.lh {n w : V} (h : L.IsUTermVec n w) : n = len w := h.1 -lemma Language.SemitermVec.prop {n m w : V} (h : L.SemitermVec n m w) {i} : i < n → L.Semiterm m w.[i] := h.2 i +lemma Language.IsUTermVec.nth {n w : V} (h : L.IsUTermVec n w) {i} : i < n → L.IsUTerm w.[i] := h.2 i -@[simp] lemma Language.SemitermVec.empty (m : V) : L.SemitermVec 0 m 0 := ⟨by simp, by simp⟩ +@[simp] lemma Language.IsUTermVec.empty : L.IsUTermVec 0 0 := ⟨by simp, by simp⟩ -@[simp] lemma Language.SemitermVec.empty_iff (m : V) : L.SemitermVec 0 m v ↔ v = 0 := by +@[simp] lemma Language.IsUTermVec.empty_iff : L.IsUTermVec 0 v ↔ v = 0 := by constructor · intro h; exact len_zero_iff_eq_nil.mp h.lh.symm · rintro rfl; simp -lemma Language.SemitermVec.two_iff {m v : V} : L.SemitermVec 2 m v ↔ ∃ t₁ t₂, L.Semiterm m t₁ ∧ L.Semiterm m t₂ ∧ v = ?[t₁, t₂] := by +lemma Language.IsUTermVec.two_iff {v : V} : L.IsUTermVec 2 v ↔ ∃ t₁ t₂, L.IsUTerm t₁ ∧ L.IsUTerm t₂ ∧ v = ?[t₁, t₂] := by constructor · intro h rcases eq_doubleton_of_len_eq_two.mp h.lh.symm with ⟨t₁, t₂, rfl⟩ - exact ⟨t₁, t₂, by simpa using h.prop (show 0 < 2 by simp), by simpa using h.prop (show 1 < 2 by simp), rfl⟩ + exact ⟨t₁, t₂, by simpa using h.nth (show 0 < 2 by simp), by simpa using h.nth (show 1 < 2 by simp), rfl⟩ · rintro ⟨t₁, t₂, h₁, h₂, rfl⟩; exact ⟨by simp [one_add_one_eq_two], by simp [lt_two_iff_le_one, le_one_iff_eq_zero_or_one, h₁, h₂]⟩ -@[simp] lemma Language.SemitermVec.cons {n m w t : V} (h : L.SemitermVec n m w) (ht : L.Semiterm m t) : L.SemitermVec (n + 1) m (t ∷ w) := +@[simp] lemma Language.IsUTermVec.cons {n w t : V} (h : L.IsUTermVec n w) (ht : L.IsUTerm t) : L.IsUTermVec (n + 1) (t ∷ w) := ⟨by simp [h.lh], fun i hi ↦ by rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simpa - · simpa using h.prop (by simpa using hi)⟩ + · simpa using h.nth (by simpa using hi)⟩ -@[simp] lemma Language.SemitermVec.cons₁_iff {m t : V} : - L.SemitermVec 1 m ?[t] ↔ L.Semiterm m t := by +@[simp] lemma Language.IsUTermVec.cons₁_iff {t : V} : + L.IsUTermVec 1 ?[t] ↔ L.IsUTerm t := by constructor - · intro h; simpa using h.prop (i := 0) (by simp) - · intro h; simpa using (Language.SemitermVec.empty m).cons h + · intro h; simpa using h.nth (i := 0) (by simp) + · intro h; simpa using Language.IsUTermVec.empty.cons h -@[simp] lemma Language.SemitermVec.mkSeq₂_iff {m t₁ t₂ : V} : - L.SemitermVec 2 m ?[t₁, t₂] ↔ L.Semiterm m t₁ ∧ L.Semiterm m t₂ := by +@[simp] lemma Language.IsUTermVec.mkSeq₂_iff {t₁ t₂ : V} : + L.IsUTermVec 2 ?[t₁, t₂] ↔ L.IsUTerm t₁ ∧ L.IsUTerm t₂ := by constructor - · intro h; exact ⟨by simpa using h.prop (i := 0) (by simp), by simpa using h.prop (i := 1) (by simp)⟩ + · intro h; exact ⟨by simpa using h.nth (i := 0) (by simp), by simpa using h.nth (i := 1) (by simp)⟩ · rintro ⟨h₁, h₂⟩ - simpa [one_add_one_eq_two] using (Language.SemitermVec.cons₁_iff.mpr h₂).cons h₁ + simpa [one_add_one_eq_two] using (Language.IsUTermVec.cons₁_iff.mpr h₂).cons h₁ section -def _root_.LO.FirstOrder.Arith.LDef.semitermVecDef (pL : LDef) : 𝚫₁.Semisentence 3 := .mkDelta +def _root_.LO.FirstOrder.Arith.LDef.isUTermVecDef (pL : LDef) : 𝚫₁.Semisentence 2 := .mkDelta (.mkSigma - “n m w | !lenDef n w ∧ ∀ i < n, ∃ u, !nthDef u w i ∧ !pL.isSemitermDef.sigma m u” + “n w | !lenDef n w ∧ ∀ i < n, ∃ u, !nthDef u w i ∧ !pL.isUTermDef.sigma u” (by simp)) (.mkPi - “n m w | (∀ l, !lenDef l w → n = l) ∧ ∀ i < n, ∀ u, !nthDef u w i → !pL.isSemitermDef.pi m u” + “n w | (∀ l, !lenDef l w → n = l) ∧ ∀ i < n, ∀ u, !nthDef u w i → !pL.isUTermDef.pi u” (by simp)) variable (L) -lemma semitermVec_defined : 𝚫₁-Relation₃ L.SemitermVec via pL.semitermVecDef := - ⟨by intro v; simp [LDef.semitermVecDef, HierarchySymbol.Semiformula.val_sigma, eval_isSemitermDef L, (isSemiterm_defined L).proper.iff'], - by intro v; simp [LDef.semitermVecDef, HierarchySymbol.Semiformula.val_sigma, eval_isSemitermDef L, Language.SemitermVec]⟩ +lemma Language.isUTermVec_defined : 𝚫₁-Relation L.IsUTermVec via pL.isUTermVecDef := + ⟨by intro v; simp [LDef.isUTermVecDef, HierarchySymbol.Semiformula.val_sigma, eval_isUTermDef L, L.isUTerm_defined.proper.iff'], + by intro v; simp [LDef.isUTermVecDef, HierarchySymbol.Semiformula.val_sigma, eval_isUTermDef L, Language.IsUTermVec]⟩ -@[simp] lemma eval_semitermVecDef (v) : - Semiformula.Evalbm V v pL.semitermVecDef.val ↔ L.SemitermVec (v 0) (v 1) (v 2) := (semitermVec_defined L).df.iff v +@[simp] lemma eval_isUTermVecDef (v) : + Semiformula.Evalbm V v pL.isUTermVecDef.val ↔ L.IsUTermVec (v 0) (v 1) := L.isUTermVec_defined.df.iff v -instance semitermVecDef_definable : 𝚫₁-Relation₃ (L.SemitermVec) := (semitermVec_defined L).to_definable +instance Language.isUTermVecDef_definable : 𝚫₁-Relation (L.IsUTermVec) := L.isUTermVec_defined.to_definable -@[simp, definability] instance semitermVecDef_definable' (Γ) : Γ-[m + 1]-Relation₃ (L.SemitermVec) := - .of_deltaOne (semitermVecDef_definable L) _ _ +instance Language.isUTermVecDef_definable' (Γ) : Γ-[m + 1]-Relation L.IsUTermVec := L.isUTermVecDef_definable.of_deltaOne end -variable {n : V} - -lemma Language.Semiterm.case_iff {t : V} : - L.Semiterm n t ↔ - (∃ z < n, t = ^#z) ∨ +lemma Language.IsUTerm.case_iff {t : V} : + L.IsUTerm t ↔ + (∃ z, t = ^#z) ∨ (∃ x, t = ^&x) ∨ - (∃ k f v : V, L.Func k f ∧ L.SemitermVec k n v ∧ t = ^func k f v) := by - simpa [construction, Phi, SemitermVec, and_assoc] using (construction L).case + (∃ k f v : V, L.Func k f ∧ L.IsUTermVec k v ∧ t = ^func k f v) := by + simpa [construction, Phi, IsUTermVec, and_assoc] using (construction L).case -alias ⟨Language.Semiterm.case, Language.Semiterm.mk⟩ := Language.Semiterm.case_iff +alias ⟨Language.IsUTerm.case, Language.IsUTerm.mk⟩ := Language.IsUTerm.case_iff -@[simp] lemma Language.Semiterm.bvar {z : V} : L.Semiterm n ^#z ↔ z < n := - ⟨by intro h - rcases h.case with (⟨z', hz, hzz'⟩ | ⟨x, h⟩ | ⟨k, f, v, _, _, h⟩) - · rcases (show z = z' from by simpa using hzz'); exact hz - · simp [qqBvar, qqFvar] at h - · simp [qqBvar, qqFunc] at h, - fun hz ↦ Language.Semiterm.mk (Or.inl ⟨z, hz, rfl⟩)⟩ +@[simp] lemma Language.IsUTerm.bvar {z : V} : L.IsUTerm ^#z := + Language.IsUTerm.mk (Or.inl ⟨z, rfl⟩) -@[simp] lemma Language.Semiterm.fvar (x : V) : L.Semiterm n ^&x := Language.Semiterm.mk (Or.inr <| Or.inl ⟨x, rfl⟩) +@[simp] lemma Language.IsUTerm.fvar (x : V) : L.IsUTerm ^&x := + Language.IsUTerm.mk (Or.inr <| Or.inl ⟨x, rfl⟩) -@[simp] lemma Language.Semiterm.func_iff {k f v : V} : - L.Semiterm n (^func k f v) ↔ L.Func k f ∧ L.SemitermVec k n v := +@[simp] lemma Language.IsUTerm.func_iff {k f v : V} : + L.IsUTerm (^func k f v) ↔ L.Func k f ∧ L.IsUTermVec k v := ⟨by intro h - rcases h.case with (⟨_, _, h⟩ | ⟨x, h⟩ | ⟨k', f', w, hkf, ⟨hk, hv⟩, h⟩) + rcases h.case with (⟨_, h⟩ | ⟨x, h⟩ | ⟨k', f', w, hkf, ⟨hk, hv⟩, h⟩) · simp [qqFunc, qqBvar] at h · simp [qqFunc, qqFvar] at h · rcases (show k = k' ∧ f = f' ∧ v = w by simpa [qqFunc] using h) with ⟨rfl, rfl, rfl⟩ exact ⟨hkf, hk, hv⟩, - by rintro ⟨hkf, hk, hv⟩; exact Language.Semiterm.mk (Or.inr <| Or.inr ⟨k, f, v, hkf, ⟨hk, hv⟩, rfl⟩)⟩ - -lemma Language.Semiterm.func {k f v : V} (hkf : L.Func k f) (hv : L.SemitermVec k n v) : - L.Semiterm n (^func k f v) := Language.Semiterm.func_iff.mpr ⟨hkf, hv⟩ - -lemma Language.Semiterm.induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) - (hbvar : ∀ z < n, P (^#z)) (hfvar : ∀ x, P (^&x)) - (hfunc : ∀ k f v, L.Func k f → L.SemitermVec k n v → (∀ i < k, P v.[i]) → P (^func k f v)) : - ∀ t, L.Semiterm n t → P t := - (construction L).induction (v := ![n]) hP (by - rintro C hC x (⟨z, hz, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hkf, hk, h, rfl⟩) - · exact hbvar z hz + by rintro ⟨hkf, hk, hv⟩; exact Language.IsUTerm.mk (Or.inr <| Or.inr ⟨k, f, v, hkf, ⟨hk, hv⟩, rfl⟩)⟩ + +lemma Language.IsUTerm.func {k f v : V} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + L.IsUTerm (^func k f v) := Language.IsUTerm.func_iff.mpr ⟨hkf, hv⟩ + +lemma Language.IsUTerm.induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) + (hbvar : ∀ z, P (^#z)) (hfvar : ∀ x, P (^&x)) + (hfunc : ∀ k f v, L.Func k f → L.IsUTermVec k v → (∀ i < k, P v.[i]) → P (^func k f v)) : + ∀ t, L.IsUTerm t → P t := + (construction L).induction (v := ![]) hP (by + rintro C hC x (⟨z, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hkf, hk, h, rfl⟩) + · exact hbvar z · exact hfvar x · exact hfunc k f v hkf ⟨hk, fun i hi ↦ hC _ (h i hi) |>.1⟩ (fun i hi ↦ hC _ (h i hi) |>.2)) @@ -283,54 +271,54 @@ end term namespace Language.TermRec structure Blueprint (pL : LDef) (arity : ℕ) where - bvar : 𝚺₁.Semisentence (arity + 3) - fvar : 𝚺₁.Semisentence (arity + 3) - func : 𝚺₁.Semisentence (arity + 6) + bvar : 𝚺₁.Semisentence (arity + 2) + fvar : 𝚺₁.Semisentence (arity + 2) + func : 𝚺₁.Semisentence (arity + 5) namespace Blueprint variable (β : Blueprint pL arity) -def blueprint : Fixpoint.Blueprint (arity + 1) := ⟨.mkDelta - (.mkSigma “pr C n | - ∃ t <⁺ pr, ∃ y <⁺ pr, !pairDef pr t y ∧ !pL.isSemitermDef.sigma n t ∧ - ( (∃ z < t, !qqBvarDef t z ∧ !β.bvar y n z ⋯) ∨ - (∃ x < t, !qqFvarDef t x ∧ !β.fvar y n x ⋯) ∨ +def blueprint : Fixpoint.Blueprint arity := ⟨.mkDelta + (.mkSigma “pr C | + ∃ t <⁺ pr, ∃ y <⁺ pr, !pairDef pr t y ∧ !pL.isUTermDef.sigma t ∧ + ( (∃ z < t, !qqBvarDef t z ∧ !β.bvar y z ⋯) ∨ + (∃ x < t, !qqFvarDef t x ∧ !β.fvar y x ⋯) ∨ (∃ k < t, ∃ f < t, ∃ v < t, ∃ rv, !repeatVecDef rv C k ∧ ∃ w <⁺ rv, (!lenDef k w ∧ ∀ i < k, ∃ vi, !nthDef vi v i ∧ ∃ v'i, !nthDef v'i w i ∧ :⟪vi, v'i⟫:∈ C) ∧ - !qqFuncDef t k f v ∧ !β.func y n k f v w ⋯) )” + !qqFuncDef t k f v ∧ !β.func y k f v w ⋯) )” (by simp)) - (.mkPi “pr C n | - ∃ t <⁺ pr, ∃ y <⁺ pr, !pairDef pr t y ∧ !pL.isSemitermDef.pi n t ∧ - ( (∃ z < t, !qqBvarDef t z ∧ !β.bvar.graphDelta.pi y n z ⋯) ∨ - (∃ x < t, !qqFvarDef t x ∧ !β.fvar.graphDelta.pi y n x ⋯) ∨ + (.mkPi “pr C | + ∃ t <⁺ pr, ∃ y <⁺ pr, !pairDef pr t y ∧ !pL.isUTermDef.pi t ∧ + ( (∃ z < t, !qqBvarDef t z ∧ !β.bvar.graphDelta.pi y z ⋯) ∨ + (∃ x < t, !qqFvarDef t x ∧ !β.fvar.graphDelta.pi y x ⋯) ∨ (∃ k < t, ∃ f < t, ∃ v < t, ∀ rv, !repeatVecDef rv C k → ∃ w <⁺ rv, ((∀ l, !lenDef l w → k = l) ∧ ∀ i < k, ∀ vi, !nthDef vi v i → ∀ v'i, !nthDef v'i w i → :⟪vi, v'i⟫:∈ C) ∧ - !qqFuncDef t k f v ∧ !β.func.graphDelta.pi y n k f v w ⋯) )” + !qqFuncDef t k f v ∧ !β.func.graphDelta.pi y k f v w ⋯) )” (by simp))⟩ -def graph : 𝚺₁.Semisentence (arity + 3) := .mkSigma - “n t y | ∃ pr <⁺ (t + y + 1)², !pairDef pr t y ∧ !β.blueprint.fixpointDef pr n ⋯” (by simp) +def graph : 𝚺₁.Semisentence (arity + 2) := .mkSigma + “t y | ∃ pr <⁺ (t + y + 1)², !pairDef pr t y ∧ !β.blueprint.fixpointDef pr ⋯” (by simp) -def result : 𝚺₁.Semisentence (arity + 3) := .mkSigma - “y n t | (!pL.isSemitermDef.pi n t → !β.graph n t y ⋯) ∧ (¬!pL.isSemitermDef.sigma n t → y = 0)” (by simp) +def result : 𝚺₁.Semisentence (arity + 2) := .mkSigma + “y t | (!pL.isUTermDef.pi t → !β.graph t y ⋯) ∧ (¬!pL.isUTermDef.sigma t → y = 0)” (by simp) -def resultVec : 𝚺₁.Semisentence (arity + 4) := .mkSigma - “w' k n w | - (!pL.semitermVecDef.pi k n w → !lenDef k w' ∧ ∀ i < k, ∃ z, !nthDef z w i ∧ ∃ z', !nthDef z' w' i ∧ !β.graph.val n z z' ⋯) ∧ - (¬!pL.semitermVecDef.sigma k n w → w' = 0)” (by simp) +def resultVec : 𝚺₁.Semisentence (arity + 3) := .mkSigma + “w' k w | + (!pL.isUTermVecDef.pi k w → !lenDef k w' ∧ ∀ i < k, ∃ z, !nthDef z w i ∧ ∃ z', !nthDef z' w' i ∧ !β.graph.val z z' ⋯) ∧ + (¬!pL.isUTermVecDef.sigma k w → w' = 0)” (by simp) end Blueprint variable (V) structure Construction (L : Arith.Language V) {k : ℕ} (φ : Blueprint pL k) where - bvar : (Fin k → V) → V → V → V - fvar : (Fin k → V) → V → V → V - func : (Fin k → V) → V → V → V → V → V → V - bvar_defined : 𝚺₁.DefinedFunction (fun v ↦ bvar (v ·.succ.succ) (v 0) (v 1)) φ.bvar - fvar_defined : 𝚺₁.DefinedFunction (fun v ↦ fvar (v ·.succ.succ) (v 0) (v 1)) φ.fvar - func_defined : 𝚺₁.DefinedFunction (fun v ↦ func (v ·.succ.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3) (v 4)) φ.func + bvar : (Fin k → V) → V → V + fvar : (Fin k → V) → V → V + func : (Fin k → V) → V → V → V → V → V + bvar_defined : 𝚺₁.DefinedFunction (fun v ↦ bvar (v ·.succ) (v 0)) φ.bvar + fvar_defined : 𝚺₁.DefinedFunction (fun v ↦ fvar (v ·.succ) (v 0)) φ.fvar + func_defined : 𝚺₁.DefinedFunction (fun v ↦ func (v ·.succ.succ.succ.succ) (v 0) (v 1) (v 2) (v 3)) φ.func variable {V} @@ -338,11 +326,11 @@ namespace Construction variable {arity : ℕ} {β : Blueprint pL arity} (c : Construction V L β) -def Phi (param : Fin arity → V) (n : V) (C : Set V) (pr : V) : Prop := - L.Semiterm n (π₁ pr) ∧ - ( (∃ z, pr = ⟪^#z, c.bvar param n z⟫) ∨ - (∃ x, pr = ⟪^&x, c.fvar param n x⟫) ∨ - (∃ k f v w, (k = len w ∧ ∀ i < k, ⟪v.[i], w.[i]⟫ ∈ C) ∧ pr = ⟪^func k f v, c.func param n k f v w⟫) ) +def Phi (param : Fin arity → V) (C : Set V) (pr : V) : Prop := + L.IsUTerm (π₁ pr) ∧ + ( (∃ z, pr = ⟪^#z, c.bvar param z⟫) ∨ + (∃ x, pr = ⟪^&x, c.fvar param x⟫) ∨ + (∃ k f v w, (k = len w ∧ ∀ i < k, ⟪v.[i], w.[i]⟫ ∈ C) ∧ pr = ⟪^func k f v, c.func param k f v w⟫) ) lemma seq_bound {k s m : V} (Ss : Seq s) (hk : k = lh s) (hs : ∀ i z, ⟪i, z⟫ ∈ s → z < m) : s < exp ((k + m + 1)^2) := lt_exp_iff.mpr <| fun p hp ↦ by @@ -351,14 +339,14 @@ lemma seq_bound {k s m : V} (Ss : Seq s) (hk : k = lh s) (hs : ∀ i z, ⟪i, z pair_lt_pair (Ss.lt_lh_of_mem (show ⟪π₁ p, π₂ p⟫ ∈ s by simpa using hp)) (hs (π₁ p) (π₂ p) (by simpa using hp)) exact lt_of_lt_of_le this (by simp) -private lemma phi_iff (param : Fin arity → V) (n C pr : V) : - c.Phi param n {x | x ∈ C} pr ↔ - ∃ t ≤ pr, ∃ y ≤ pr, pr = ⟪t, y⟫ ∧ L.Semiterm n t ∧ - ( (∃ z < t, t = ^#z ∧ y = c.bvar param n z) ∨ - (∃ x < t, t = ^&x ∧ y = c.fvar param n x) ∨ +private lemma phi_iff (param : Fin arity → V) (C pr : V) : + c.Phi param {x | x ∈ C} pr ↔ + ∃ t ≤ pr, ∃ y ≤ pr, pr = ⟪t, y⟫ ∧ L.IsUTerm t ∧ + ( (∃ z < t, t = ^#z ∧ y = c.bvar param z) ∨ + (∃ x < t, t = ^&x ∧ y = c.fvar param x) ∨ (∃ k < t, ∃ f < t, ∃ v < t, ∃ w ≤ repeatVec C k, (k = len w ∧ ∀ i < k, ⟪v.[i], w.[i]⟫ ∈ C) ∧ - t = ^func k f v ∧ y = c.func param n k f v w) ) := by + t = ^func k f v ∧ y = c.func param k f v w) ) := by constructor · rintro ⟨ht, H⟩ refine ⟨π₁ pr, by simp, π₂ pr, by simp, by simp, ht, ?_⟩ @@ -384,12 +372,12 @@ private lemma phi_iff (param : Fin arity → V) (n C pr : V) : @[simp] lemma cons_app_11 {n : ℕ} (a : α) (s : Fin n.succ.succ.succ.succ.succ.succ.succ.succ.succ.succ.succ → α) : (a :> s) 11 = s 10 := rfl def construction : Fixpoint.Construction V β.blueprint where - Φ := fun v ↦ c.Phi (v ·.succ) (v 0) + Φ := c.Phi defined := ⟨by intro v /- simp? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, - eval_isSemitermDef L, (isSemiterm_defined L).proper.iff', + eval_isUTermDef L, (isUTerm_defined L).proper.iff', c.bvar_defined.iff, c.bvar_defined.graph_delta.proper.iff', c.fvar_defined.iff, c.fvar_defined.graph_delta.proper.iff', c.func_defined.iff, c.func_defined.graph_delta.proper.iff'] @@ -401,14 +389,14 @@ def construction : Fixpoint.Construction V β.blueprint where Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Matrix.cons_val_succ, Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, pair_defined_iff, Fin.isValue, Fin.succ_zero_eq_one, - Matrix.cons_val_four, eval_isSemitermDef L, LogicalConnective.HomClass.map_or, + Matrix.cons_val_four, eval_isUTermDef L, LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, eval_qqBvarDef, Matrix.cons_app_five, c.bvar_defined.iff, LogicalConnective.Prop.and_eq, eval_qqFvarDef, c.fvar_defined.iff, Matrix.cons_val_three, Semiformula.eval_ex, Matrix.cons_app_seven, Matrix.cons_app_six, eval_repeatVec, eval_lenDef, Semiformula.eval_ballLT, eval_nthDef, Semiformula.eval_operator₃, cons_app_11, cons_app_10, cons_app_9, Matrix.cons_app_eight, eval_memRel, exists_eq_left, eval_qqFuncDef, Fin.succ_one_eq_two, c.func_defined.iff, LogicalConnective.Prop.or_eq, - HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, (isSemiterm_defined L).proper.iff', + HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, (isUTerm_defined L).proper.iff', c.bvar_defined.graph_delta.proper.iff', HierarchySymbol.Semiformula.graphDelta_val, c.fvar_defined.graph_delta.proper.iff', Semiformula.eval_all, LogicalConnective.HomClass.map_imply, Semiformula.eval_operator₂, Structure.Eq.eq, @@ -416,25 +404,26 @@ def construction : Fixpoint.Construction V β.blueprint where , by intro v /- - simpa? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_isSemitermDef L, + simpa? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_isUTermDef L, c.bvar_defined.iff, c.fvar_defined.iff, c.func_defined.iff] - using c.phi_iff _ _ _ _ + using c.phi_iff _ _ _ -/ - simpa only [Nat.succ_eq_add_one, BinderNotation.finSuccItr_one, Fin.succ_zero_eq_one, - Blueprint.blueprint, Nat.reduceAdd, HierarchySymbol.Semiformula.val_sigma, Nat.add_zero, - HierarchySymbol.Semiformula.val_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, Semiformula.eval_bexLTSucc', - Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, + simpa only [Nat.succ_eq_add_one, BinderNotation.finSuccItr_one, Blueprint.blueprint, + Nat.reduceAdd, HierarchySymbol.Semiformula.val_sigma, + HierarchySymbol.Semiformula.val_mkDelta, HierarchySymbol.Semiformula.val_mkSigma, + Semiformula.eval_bexLTSucc', Semiterm.val_bvar, Matrix.cons_val_one, Matrix.vecHead, LogicalConnective.HomClass.map_and, Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Matrix.cons_val_succ, Matrix.cons_val_zero, Matrix.cons_val_fin_one, Matrix.constant_eq_singleton, - pair_defined_iff, Fin.isValue, Matrix.cons_val_four, eval_isSemitermDef L, + pair_defined_iff, Fin.isValue, Fin.succ_zero_eq_one, eval_isUTermDef L, LogicalConnective.HomClass.map_or, Semiformula.eval_bexLT, eval_qqBvarDef, - Matrix.cons_app_five, c.bvar_defined.iff, LogicalConnective.Prop.and_eq, eval_qqFvarDef, - c.fvar_defined.iff, Matrix.cons_val_three, Semiformula.eval_ex, Matrix.cons_app_seven, - Matrix.cons_app_six, eval_repeatVec, eval_lenDef, Semiformula.eval_ballLT, eval_nthDef, - Semiformula.eval_operator₃, cons_app_11, cons_app_10, cons_app_9, Matrix.cons_app_eight, - eval_memRel, exists_eq_left, eval_qqFuncDef, Fin.succ_one_eq_two, c.func_defined.iff, - LogicalConnective.Prop.or_eq] using c.phi_iff _ _ _ _⟩ + c.bvar_defined.iff, LogicalConnective.Prop.and_eq, eval_qqFvarDef, c.fvar_defined.iff, + Matrix.cons_val_three, Semiformula.eval_ex, Matrix.cons_app_seven, Matrix.cons_app_six, + Matrix.cons_app_five, Matrix.cons_val_four, eval_repeatVec, eval_lenDef, + Semiformula.eval_ballLT, eval_nthDef, Semiformula.eval_operator₃, cons_app_11, + cons_app_10, cons_app_9, Matrix.cons_app_eight, eval_memRel, exists_eq_left, + eval_qqFuncDef, Fin.succ_one_eq_two, c.func_defined.iff, + LogicalConnective.Prop.or_eq] using c.phi_iff _ _ _⟩ monotone := by unfold Phi rintro C C' hC v pr ⟨ht, H⟩ @@ -453,54 +442,54 @@ instance : c.construction.Finite where ⟨k, f, v, w, ⟨hk, fun i hi ↦ ⟨hw i hi, lt_succ_iff_le.mpr <| pair_le_pair (by simp) (by simp)⟩⟩, rfl⟩⟩ -def Graph (param : Fin arity → V) (n x y : V) : Prop := c.construction.Fixpoint (n :> param) ⟪x, y⟫ +def Graph (param : Fin arity → V) (x y : V) : Prop := c.construction.Fixpoint param ⟪x, y⟫ variable {param : Fin arity → V} {n : V} variable {c} lemma Graph.case_iff {t y : V} : - c.Graph param n t y ↔ - L.Semiterm n t ∧ - ( (∃ z, t = ^#z ∧ y = c.bvar param n z) ∨ - (∃ x, t = ^&x ∧ y = c.fvar param n x) ∨ - (∃ k f v w, (k = len w ∧ ∀ i < k, c.Graph param n v.[i] w.[i]) ∧ - t = ^func k f v ∧ y = c.func param n k f v w) ) := + c.Graph param t y ↔ + L.IsUTerm t ∧ + ( (∃ z, t = ^#z ∧ y = c.bvar param z) ∨ + (∃ x, t = ^&x ∧ y = c.fvar param x) ∨ + (∃ k f v w, (k = len w ∧ ∀ i < k, c.Graph param v.[i] w.[i]) ∧ + t = ^func k f v ∧ y = c.func param k f v w) ) := Iff.trans c.construction.case (by apply and_congr (by simp); simp; rfl) variable (c) -lemma graph_defined : 𝚺₁.Defined (fun v ↦ c.Graph (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) β.graph := by +lemma graph_defined : 𝚺₁.Defined (fun v ↦ c.Graph (v ·.succ.succ) (v 0) (v 1)) β.graph := by intro v simp [Blueprint.graph, c.construction.fixpoint_defined.iff] constructor - · intro h; exact ⟨⟪v 1, v 2⟫, by simp, rfl, h⟩ + · intro h; exact ⟨⟪v 0, v 1⟫, by simp, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h @[simp] lemma eval_graphDef (v) : - Semiformula.Evalbm V v β.graph.val ↔ c.Graph (v ·.succ.succ.succ) (v 0) (v 1) (v 2) := (graph_defined c).df.iff v + Semiformula.Evalbm V v β.graph.val ↔ c.Graph (v ·.succ.succ) (v 0) (v 1) := (graph_defined c).df.iff v -instance termSubst_definable : 𝚺₁.Boldface (fun v ↦ c.Graph (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) := +instance graph_definable : 𝚺₁.Boldface (fun v ↦ c.Graph (v ·.succ.succ) (v 0) (v 1)) := (graph_defined c).to_definable -instance termSubst_definable₂ (param n) : 𝚺-[0 + 1]-Relation (c.Graph param n) := by - simpa using HierarchySymbol.Boldface.retractiont (n := 2) (termSubst_definable c) (&n :> #0 :> #1 :> fun i ↦ &(param i)) +instance graph_definable₂ (param) : 𝚺-[0 + 1]-Relation (c.Graph param) := by + simpa using HierarchySymbol.Boldface.retractiont (n := 2) (graph_definable c) (#0 :> #1 :> fun i ↦ &(param i)) -lemma graph_dom_isSemiterm {t y} : - c.Graph param n t y → L.Semiterm n t := fun h ↦ Graph.case_iff.mp h |>.1 +lemma graph_dom_isUTerm {t y} : + c.Graph param t y → L.IsUTerm t := fun h ↦ Graph.case_iff.mp h |>.1 -lemma graph_bvar_iff {z} (hz : z < n) : - c.Graph param n ^#z y ↔ y = c.bvar param n z := by +lemma graph_bvar_iff {z} : + c.Graph param ^#z y ↔ y = c.bvar param z := by constructor · intro H rcases Graph.case_iff.mp H with ⟨_, (⟨_, h, rfl⟩ | ⟨_, h, _⟩ | ⟨_, _, _, _, _, h, _⟩)⟩ · simp at h; rcases h; rfl · simp [qqBvar, qqFvar] at h · simp [qqBvar, qqFunc] at h - · rintro rfl; exact Graph.case_iff.mpr ⟨by simp [hz], Or.inl ⟨z, by simp⟩⟩ + · rintro rfl; exact Graph.case_iff.mpr ⟨by simp, Or.inl ⟨z, by simp⟩⟩ lemma graph_fvar_iff (x) : - c.Graph param n ^&x y ↔ y = c.fvar param n x := by + c.Graph param ^&x y ↔ y = c.fvar param x := by constructor · intro H rcases Graph.case_iff.mp H with ⟨_, (⟨_, h, _⟩ | ⟨_, h, rfl⟩ | ⟨_, _, _, _, _, h, _⟩)⟩ @@ -509,15 +498,15 @@ lemma graph_fvar_iff (x) : · simp [qqFvar, qqFunc] at h · rintro rfl; exact Graph.case_iff.mpr ⟨by simp, Or.inr <| Or.inl ⟨x, by simp⟩⟩ -lemma graph_func {n k f v w} (hkr : L.Func k f) (hv : L.SemitermVec k n v) - (hkw : k = len w) (hw : ∀ i < k, c.Graph param n v.[i] w.[i]) : - c.Graph param n (^func k f v) (c.func param n k f v w) := by +lemma graph_func {k f v w} (hkr : L.Func k f) (hv : L.IsUTermVec k v) + (hkw : k = len w) (hw : ∀ i < k, c.Graph param v.[i] w.[i]) : + c.Graph param (^func k f v) (c.func param k f v w) := by exact Graph.case_iff.mpr ⟨by simp [hkr, hv], Or.inr <| Or.inr ⟨k, f, v, w, ⟨hkw, hw⟩, by simp⟩⟩ -lemma graph_func_inv {n k f v y} : - c.Graph param n (^func k f v) y → ∃ w, - (k = len w ∧ ∀ i < k, c.Graph param n v.[i] w.[i]) ∧ - y = c.func param n k f v w := by +lemma graph_func_inv {k f v y} : + c.Graph param (^func k f v) y → ∃ w, + (k = len w ∧ ∀ i < k, c.Graph param v.[i] w.[i]) ∧ + y = c.func param k f v w := by intro H rcases Graph.case_iff.mp H with ⟨_, (⟨_, h, _⟩ | ⟨_, h, rfl⟩ | ⟨k, f, v, w, hw, h, rfl⟩)⟩ · simp [qqFunc, qqBvar] at h @@ -527,26 +516,26 @@ lemma graph_func_inv {n k f v y} : variable {c} (param n) -lemma graph_exists {t : V} : L.Semiterm n t → ∃ y, c.Graph param n t y := by - apply Language.Semiterm.induction 𝚺 (P := fun t ↦ ∃ y, c.Graph param n t y) (by definability) +lemma graph_exists {t : V} : L.IsUTerm t → ∃ y, c.Graph param t y := by + apply Language.IsUTerm.induction 𝚺 (P := fun t ↦ ∃ y, c.Graph param t y) (by definability) case hbvar => - intro z hz; exact ⟨c.bvar param n z, by simp [c.graph_bvar_iff hz]⟩ + intro z; exact ⟨c.bvar param z, by simp [c.graph_bvar_iff]⟩ case hfvar => - intro x; exact ⟨c.fvar param n x, by simp [c.graph_fvar_iff]⟩ + intro x; exact ⟨c.fvar param x, by simp [c.graph_fvar_iff]⟩ case hfunc => intro k f v hkf hv ih - have : ∃ w, len w = k ∧ ∀ i < k, c.Graph param n v.[i] w.[i] := sigmaOne_skolem_vec (by definability) ih + have : ∃ w, len w = k ∧ ∀ i < k, c.Graph param v.[i] w.[i] := sigmaOne_skolem_vec (by definability) ih rcases this with ⟨w, hwk, hvw⟩ - exact ⟨c.func param n k f v w, c.graph_func hkf hv (Eq.symm hwk) hvw⟩ + exact ⟨c.func param k f v w, c.graph_func hkf hv (Eq.symm hwk) hvw⟩ -lemma graph_unique {t y₁ y₂ : V} : c.Graph param n t y₁ → c.Graph param n t y₂ → y₁ = y₂ := by +lemma graph_unique {t y₁ y₂ : V} : c.Graph param t y₁ → c.Graph param t y₂ → y₁ = y₂ := by revert y₁ y₂ - suffices L.Semiterm n t → ∀ y₁ y₂, c.Graph param n t y₁ → c.Graph param n t y₂ → y₁ = y₂ - by intro u₁ u₂ h₁ h₂; exact this (c.graph_dom_isSemiterm h₁) u₁ u₂ h₁ h₂ + suffices L.IsUTerm t → ∀ y₁ y₂, c.Graph param t y₁ → c.Graph param t y₂ → y₁ = y₂ + by intro u₁ u₂ h₁ h₂; exact this (c.graph_dom_isUTerm h₁) u₁ u₂ h₁ h₂ intro ht - apply Language.Semiterm.induction 𝚷 ?_ ?_ ?_ ?_ t ht + apply Language.IsUTerm.induction 𝚷 ?_ ?_ ?_ ?_ t ht · definability - · intro z hz; simp [c.graph_bvar_iff hz] + · intro z; simp [c.graph_bvar_iff] · intro x; simp [c.graph_fvar_iff] · intro k f v _ _ ih y₁ y₂ h₁ h₂ rcases c.graph_func_inv h₁ with ⟨w₁, ⟨hk₁, hv₁⟩, rfl⟩ @@ -559,81 +548,81 @@ lemma graph_unique {t y₁ y₂ : V} : c.Graph param n t y₁ → c.Graph param variable (c) -lemma graph_existsUnique {t : V} (ht : L.Semiterm n t) : ∃! y, c.Graph param n t y := by - rcases graph_exists param n ht with ⟨y, hy⟩ - exact ExistsUnique.intro y hy (fun y' h' ↦ graph_unique param n h' hy) +lemma graph_existsUnique {t : V} (ht : L.IsUTerm t) : ∃! y, c.Graph param t y := by + rcases graph_exists param ht with ⟨y, hy⟩ + exact ExistsUnique.intro y hy (fun y' h' ↦ graph_unique param h' hy) lemma graph_existsUnique_total (t : V) : ∃! y, - (L.Semiterm n t → c.Graph param n t y) ∧ (¬L.Semiterm n t → y = 0) := by - by_cases ht : L.Semiterm n t <;> simp [ht]; exact c.graph_existsUnique _ _ ht + (L.IsUTerm t → c.Graph param t y) ∧ (¬L.IsUTerm t → y = 0) := by + by_cases ht : L.IsUTerm t <;> simp [ht]; exact c.graph_existsUnique _ ht -def result (t : V) : V := Classical.choose! (c.graph_existsUnique_total param n t) +def result (t : V) : V := Classical.choose! (c.graph_existsUnique_total param t) -def result_prop {t : V} (ht : L.Semiterm n t) : c.Graph param n t (c.result param n t) := - Classical.choose!_spec (c.graph_existsUnique_total param n t) |>.1 ht +def result_prop {t : V} (ht : L.IsUTerm t) : c.Graph param t (c.result param t) := + Classical.choose!_spec (c.graph_existsUnique_total param t) |>.1 ht -def result_prop_not {t : V} (ht : ¬L.Semiterm n t) : c.result param n t = 0 := - Classical.choose!_spec (c.graph_existsUnique_total param n t) |>.2 ht +def result_prop_not {t : V} (ht : ¬L.IsUTerm t) : c.result param t = 0 := + Classical.choose!_spec (c.graph_existsUnique_total param t) |>.2 ht -variable {c param n} +variable {c param} -lemma result_eq_of_graph {t y} (ht : L.Semiterm n t) (h : c.Graph param n t y) : - c.result param n t = y := Eq.symm <| Classical.choose_uniq (c.graph_existsUnique_total param n t) (by simp [h, ht]) +lemma result_eq_of_graph {t y} (ht : L.IsUTerm t) (h : c.Graph param t y) : + c.result param t = y := Eq.symm <| Classical.choose_uniq (c.graph_existsUnique_total param t) (by simp [h, ht]) -@[simp] lemma result_bvar {z} (hz : z < n) : c.result param n (^#z) = c.bvar param n z := - c.result_eq_of_graph (by simp [hz]) (by simp [c.graph_bvar_iff hz]) +@[simp] lemma result_bvar (z) : c.result param (^#z) = c.bvar param z := + c.result_eq_of_graph (by simp) (by simp [c.graph_bvar_iff]) -@[simp] lemma result_fvar (x) : c.result param n (^&x) = c.fvar param n x := +@[simp] lemma result_fvar (x) : c.result param (^&x) = c.fvar param x := c.result_eq_of_graph (by simp) (by simp [c.graph_fvar_iff]) -lemma result_func {k f v w} (hkf : L.Func k f) (hv : L.SemitermVec k n v) - (hkw : k = len w) (hw : ∀ i < k, c.result param n v.[i] = w.[i]) : - c.result param n (^func k f v) = c.func param n k f v w := +lemma result_func {k f v w} (hkf : L.Func k f) (hv : L.IsUTermVec k v) + (hkw : k = len w) (hw : ∀ i < k, c.result param v.[i] = w.[i]) : + c.result param (^func k f v) = c.func param k f v w := c.result_eq_of_graph (by simp [hkf, hv]) (c.graph_func hkf hv hkw (fun i hi ↦ by - simpa [hw i hi] using c.result_prop param n (hv.prop hi))) + simpa [hw i hi] using c.result_prop param (hv.nth hi))) section vec -lemma graph_existsUnique_vec {k n w : V} (hw : L.SemitermVec k n w) : - ∃! w' : V, k = len w' ∧ ∀ i < k, c.Graph param n w.[i] w'.[i] := by - have : ∀ i < k, ∃ y, c.Graph param n w.[i] y := by - intro i hi; exact ⟨c.result param n w.[i], c.result_prop param n (hw.prop hi)⟩ +lemma graph_existsUnique_vec {k w : V} (hw : L.IsUTermVec k w) : + ∃! w' : V, k = len w' ∧ ∀ i < k, c.Graph param w.[i] w'.[i] := by + have : ∀ i < k, ∃ y, c.Graph param w.[i] y := by + intro i hi; exact ⟨c.result param w.[i], c.result_prop param (hw.nth hi)⟩ rcases sigmaOne_skolem_vec (by definability) this with ⟨w', hw'k, hw'⟩ refine ExistsUnique.intro w' ⟨hw'k.symm, hw'⟩ ?_ intro w'' ⟨hkw'', hw''⟩ refine nth_ext (by simp [hw'k, ←hkw'']) (by intro i hi; - exact c.graph_unique param n (hw'' i (by simpa [hkw''] using hi)) (hw' i (by simpa [hkw''] using hi))) + exact c.graph_unique param (hw'' i (by simpa [hkw''] using hi)) (hw' i (by simpa [hkw''] using hi))) variable (c param) -lemma graph_existsUnique_vec_total (k n w : V) : ∃! w', - (L.SemitermVec k n w → k = len w' ∧ ∀ i < k, c.Graph param n w.[i] w'.[i]) ∧ - (¬L.SemitermVec k n w → w' = 0) := by - by_cases h : L.SemitermVec k n w <;> simp [h]; exact c.graph_existsUnique_vec h +lemma graph_existsUnique_vec_total (k w : V) : ∃! w', + (L.IsUTermVec k w → k = len w' ∧ ∀ i < k, c.Graph param w.[i] w'.[i]) ∧ + (¬L.IsUTermVec k w → w' = 0) := by + by_cases h : L.IsUTermVec k w <;> simp [h]; exact c.graph_existsUnique_vec h -def resultVec (k n w : V) : V := Classical.choose! (c.graph_existsUnique_vec_total param k n w) +def resultVec (k w : V) : V := Classical.choose! (c.graph_existsUnique_vec_total param k w) -@[simp] lemma resultVec_lh {k n w : V} (hw : L.SemitermVec k n w) : len (c.resultVec param k n w) = k := - Eq.symm <| Classical.choose!_spec (c.graph_existsUnique_vec_total param k n w) |>.1 hw |>.1 +@[simp] lemma resultVec_lh {k w : V} (hw : L.IsUTermVec k w) : len (c.resultVec param k w) = k := + Eq.symm <| Classical.choose!_spec (c.graph_existsUnique_vec_total param k w) |>.1 hw |>.1 -lemma graph_of_mem_resultVec {k n w : V} (hw : L.SemitermVec k n w) {i : V} (hi : i < k) : - c.Graph param n w.[i] (c.resultVec param k n w).[i] := - Classical.choose!_spec (c.graph_existsUnique_vec_total param k n w) |>.1 hw |>.2 i hi +lemma graph_of_mem_resultVec {k w : V} (hw : L.IsUTermVec k w) {i : V} (hi : i < k) : + c.Graph param w.[i] (c.resultVec param k w).[i] := + Classical.choose!_spec (c.graph_existsUnique_vec_total param k w) |>.1 hw |>.2 i hi -lemma nth_resultVec {k n w i : V} (hw : L.SemitermVec k n w) (hi : i < k) : - (c.resultVec param k n w).[i] = c.result param n w.[i] := - c.result_eq_of_graph (hw.prop hi) (c.graph_of_mem_resultVec param hw hi) |>.symm +lemma nth_resultVec {k w i : V} (hw : L.IsUTermVec k w) (hi : i < k) : + (c.resultVec param k w).[i] = c.result param w.[i] := + c.result_eq_of_graph (hw.nth hi) (c.graph_of_mem_resultVec param hw hi) |>.symm -@[simp] def resultVec_of_not {k n w : V} (hw : ¬L.SemitermVec k n w) : c.resultVec param k n w = 0 := - Classical.choose!_spec (c.graph_existsUnique_vec_total param k n w) |>.2 hw +@[simp] def resultVec_of_not {k w : V} (hw : ¬L.IsUTermVec k w) : c.resultVec param k w = 0 := + Classical.choose!_spec (c.graph_existsUnique_vec_total param k w) |>.2 hw -@[simp] lemma resultVec_nil (n : V) : - c.resultVec param 0 n 0 = 0 := len_zero_iff_eq_nil.mp (by simp) +@[simp] lemma resultVec_nil : + c.resultVec param 0 0 = 0 := len_zero_iff_eq_nil.mp (by simp) -lemma resultVec_cons {k n w t : V} (hw : L.SemitermVec k n w) (ht : L.Semiterm n t) : - c.resultVec param (k + 1) n (t ∷ w) = c.result param n t ∷ c.resultVec param k n w := +lemma resultVec_cons {k w t : V} (hw : L.IsUTermVec k w) (ht : L.IsUTerm t) : + c.resultVec param (k + 1) (t ∷ w) = c.result param t ∷ c.resultVec param k w := nth_ext (by simp [hw, hw.cons ht]) (by intro i hi have hi : i < k + 1 := by simpa [hw.cons ht, resultVec_lh] using hi @@ -646,35 +635,35 @@ end vec variable (c) -@[simp] lemma result_func' {k f v} (hkf : L.Func k f) (hv : L.SemitermVec k n v) : - c.result param n (^func k f v) = c.func param n k f v (c.resultVec param k n v) := +@[simp] lemma result_func' {k f v} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + c.result param (^func k f v) = c.func param k f v (c.resultVec param k v) := c.result_func hkf hv (by simp [hv]) (fun i hi ↦ c.nth_resultVec param hv hi |>.symm) section -lemma result_defined : 𝚺₁.DefinedFunction (fun v ↦ c.result (v ·.succ.succ) (v 0) (v 1)) β.result := by +lemma result_defined : 𝚺₁.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0)) β.result := by intro v - simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, (isSemiterm_defined L).proper.iff', - eval_isSemitermDef L, c.eval_graphDef, result, Classical.choose!_eq_iff] + simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, (isUTerm_defined L).proper.iff', + eval_isUTermDef L, c.eval_graphDef, result, Classical.choose!_eq_iff] rfl @[simp] lemma result_graphDef (v) : - Semiformula.Evalbm V v β.result.val ↔ v 0 = c.result (v ·.succ.succ.succ) (v 1) (v 2) := (result_defined c).df.iff v + Semiformula.Evalbm V v β.result.val ↔ v 0 = c.result (v ·.succ.succ) (v 1) := (result_defined c).df.iff v -private lemma resultVec_graph {w' k n w} : - w' = c.resultVec param k n w ↔ - ( (L.SemitermVec k n w → k = len w' ∧ ∀ i < k, c.Graph param n w.[i] w'.[i]) ∧ - (¬L.SemitermVec k n w → w' = 0) ) := - Classical.choose!_eq_iff (c.graph_existsUnique_vec_total param k n w) +private lemma resultVec_graph {w' k w} : + w' = c.resultVec param k w ↔ + ( (L.IsUTermVec k w → k = len w' ∧ ∀ i < k, c.Graph param w.[i] w'.[i]) ∧ + (¬L.IsUTermVec k w → w' = 0) ) := + Classical.choose!_eq_iff (c.graph_existsUnique_vec_total param k w) -lemma resultVec_defined : 𝚺₁.DefinedFunction (fun v ↦ c.resultVec (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) β.resultVec := by +lemma resultVec_defined : 𝚺₁.DefinedFunction (fun v ↦ c.resultVec (v ·.succ.succ) (v 0) (v 1)) β.resultVec := by intro v - simpa [Blueprint.resultVec, HierarchySymbol.Semiformula.val_sigma, (semitermVec_defined L).proper.iff', - eval_semitermVecDef L, c.eval_graphDef] using c.resultVec_graph + simpa [Blueprint.resultVec, HierarchySymbol.Semiformula.val_sigma, L.isUTermVec_defined.proper.iff', + eval_isUTermVecDef L, c.eval_graphDef] using c.resultVec_graph -lemma eval_resultVec (v : Fin (arity + 4) → V) : +lemma eval_resultVec (v : Fin (arity + 3) → V) : Semiformula.Evalbm V v β.resultVec.val ↔ - v 0 = c.resultVec (v ·.succ.succ.succ.succ) (v 1) (v 2) (v 3) := c.resultVec_defined.df.iff v + v 0 = c.resultVec (v ·.succ.succ.succ) (v 1) (v 2) := c.resultVec_defined.df.iff v end @@ -682,6 +671,262 @@ end Construction end Language.TermRec +namespace IsUTerm.BV + +def blueprint (pL : LDef) : Language.TermRec.Blueprint pL 0 where + bvar := .mkSigma “y z | y = z + 1” (by simp) + fvar := .mkSigma “y x | y = 0” (by simp) + func := .mkSigma “y k f v v' | !listMaxDef y v'” (by simp) + +variable (L) + +def construction : Language.TermRec.Construction V L (blueprint pL) where + bvar (_ z) := z + 1 + fvar (_ _) := 0 + func (_ _ _ _ v') := listMax v' + bvar_defined := by intro v; simp [blueprint] + fvar_defined := by intro v; simp [blueprint] + func_defined := by intro v; simp [blueprint]; rfl + +end IsUTerm.BV + +section bv + +open IsUTerm.BV + +variable (L) + +def Language.termBV (t : V) : V := (construction L).result ![] t + +def Language.termBVVec (k v : V) : V := (construction L).resultVec ![] k v + +variable {L} + +@[simp] lemma Language.termBV_bvar (z) : + L.termBV ^#z = z + 1 := by simp [Language.termBV, construction] + +@[simp] lemma Language.termBV_fvar (x) : + L.termBV ^&x = 0 := by simp [Language.termBV, construction] + +@[simp] lemma Language.termBV_func {k f v} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + L.termBV (^func k f v) = listMax (L.termBVVec k v) := by + simp [Language.termBV, construction, hkf, hv]; rfl + +@[simp] lemma Language.len_termBVVec {v} (hv : L.IsUTermVec k v) : + len (L.termBVVec k v) = k := (construction L).resultVec_lh _ hv + +@[simp] lemma Language.nth_termBVVec {v} (hv : L.IsUTermVec k v) {i} (hi : i < k) : + (L.termBVVec k v).[i] = L.termBV v.[i] := (construction L).nth_resultVec _ hv hi + +@[simp] lemma Language.termBVVec_nil : L.termBVVec 0 0 = 0 := (construction L).resultVec_nil _ + +lemma Language.termBVVec_cons {k t ts : V} (ht : L.IsUTerm t) (hts : L.IsUTermVec k ts) : + L.termBVVec (k + 1) (t ∷ ts) = L.termBV t ∷ L.termBVVec k ts := + (construction L).resultVec_cons ![] hts ht + +section + +variable (L) + +def _root_.LO.FirstOrder.Arith.LDef.termBVDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result + +def _root_.LO.FirstOrder.Arith.LDef.termBVVecDef (pL : LDef) : 𝚺₁.Semisentence 3 := (blueprint pL).resultVec + +lemma Language.termBV_defined : 𝚺₁-Function₁ L.termBV via pL.termBVDef := (construction L).result_defined + +instance Language.termBV_definable : 𝚺₁-Function₁ L.termBV := L.termBV_defined.to_definable + +instance Language.termBV_definable' : Γ-[k + 1]-Function₁ L.termBV := L.termBV_definable.of_sigmaOne + +lemma Language.termBVVec_defined : 𝚺₁.DefinedFunction (fun v ↦ L.termBVVec (v 0) (v 1)) pL.termBVVecDef := + (construction L).resultVec_defined + +instance Language.termBVVec_definable : 𝚺₁-Function₂ L.termBVVec := L.termBVVec_defined.to_definable + +instance Language.termBVVec_definable' : Γ-[i + 1]-Function₂ L.termBVVec := L.termBVVec_definable.of_sigmaOne + +end + +end bv + +section isSemiterm + +variable (L) + +def Language.IsSemiterm (n t : V) : Prop := L.IsUTerm t ∧ L.termBV t ≤ n + +def Language.IsSemitermVec (k n v : V) : Prop := L.IsUTermVec k v ∧ ∀ i < k, L.termBV v.[i] ≤ n + +abbrev Language.IsTerm (t : V) : Prop := L.IsSemiterm 0 t + +abbrev Language.IsTermVec (k v : V) : Prop := L.IsSemitermVec k 0 v + +variable {L} + +lemma Language.IsSemiterm.isUTerm {n t : V} (h : L.IsSemiterm n t) : L.IsUTerm t := h.1 + +lemma Language.IsSemiterm.bv {n t : V} (h : L.IsSemiterm n t) : L.termBV t ≤ n := h.2 + +@[simp] lemma Language.IsSemitermVec.isUTerm {k n v : V} (h : L.IsSemitermVec k n v) : L.IsUTermVec k v := h.1 + +lemma Language.IsSemitermVec.bv {k n v : V} (h : L.IsSemitermVec k n v) {i} (hi : i < k) : L.termBV v.[i] ≤ n := h.2 i hi + +lemma Language.IsSemitermVec.lh {k n v : V} (h : L.IsSemitermVec k n v) : len v = k := h.isUTerm.lh.symm + +lemma Language.IsSemitermVec.nth {k n v : V} (h : L.IsSemitermVec k n v) {i} (hi : i < k) : + L.IsSemiterm n v.[i] := ⟨h.isUTerm.nth hi, h.bv hi⟩ + +lemma Language.IsUTerm.isSemiterm {t : V} (h : L.IsUTerm t) : L.IsSemiterm (L.termBV t) t := ⟨h, by simp⟩ + +lemma Language.IsUTermVec.isSemitermVec {k v : V} (h : L.IsUTermVec k v) : L.IsSemitermVec k (listMax (L.termBVVec k v)) v := + ⟨h, fun i hi ↦ le_trans (by rw [Language.nth_termBVVec h hi]) (nth_le_listMax (i := i) (by simp [h, hi]))⟩ + +lemma Language.IsSemitermVec.iff {k n v : V} : L.IsSemitermVec k n v ↔ (len v = k ∧ ∀ i < k, L.IsSemiterm n v.[i]) := by + constructor + · intro h; exact ⟨h.lh, fun i hi ↦ h.nth hi⟩ + · intro ⟨hk, h⟩; exact ⟨⟨hk.symm, fun i hi ↦ h i hi |>.isUTerm⟩, fun i hi ↦ h i hi |>.bv⟩ + +@[simp] lemma Language.IsSemiterm.bvar {n z} : + L.IsSemiterm n ^#z ↔ z < n := by simp [Language.IsSemiterm, succ_le_iff_lt] + +@[simp] lemma Language.IsSemiterm.fvar (n x) : + L.IsSemiterm n ^&x := by simp [Language.IsSemiterm] + +@[simp] lemma Language.IsSemiterm.func {n k f v} : + L.IsSemiterm n (^func k f v) ↔ L.Func k f ∧ L.IsSemitermVec k n v := by + simp only [IsSemiterm, IsUTerm.func_iff] + constructor + · rintro ⟨⟨hf, hbv⟩, hv⟩ + exact ⟨hf, hbv, by + intro i hi + have : listMax (L.termBVVec k v) ≤ n := by simpa [termBV_func, hf, hbv] using hv + simpa [hbv, hi] using listMaxss_le_iff.mp this i (by simp [hbv, hi])⟩ + · rintro ⟨hf, h⟩ + exact ⟨⟨hf, h.isUTerm⟩, by + simp only [hf, h.isUTerm, termBV_func] + apply listMaxss_le_iff.mpr + intro i hi + have hi : i < k := by simpa [hf, h.isUTerm] using hi + simpa [hf, h.isUTerm, hi] using (h.nth hi).bv⟩ + +@[simp] lemma Language.IsSemitermVec.empty (n) : L.IsSemitermVec 0 n 0 := ⟨by simp, by simp⟩ + +@[simp] lemma Language.IsSemitermVec.empty_iff {n : V} : L.IsSemitermVec 0 n v ↔ v = 0 := by + constructor + · intro h; exact len_zero_iff_eq_nil.mp h.lh + · rintro rfl; simp + +@[simp] lemma Language.IsSemitermVec.cons_iff {n w t : V} : + L.IsSemitermVec (k + 1) n (t ∷ w) ↔ L.IsSemiterm n t ∧ L.IsSemitermVec k n w := by + constructor + · intro h + exact ⟨by simpa using h.nth (i := 0) (by simp), + Language.IsSemitermVec.iff.mpr ⟨by simpa using h.lh, fun i hi ↦ by simpa using h.nth (show i + 1 < k + 1 by simp [hi])⟩⟩ + · rintro ⟨ht, hw⟩ + exact ⟨hw.isUTerm.cons ht.isUTerm, by + intro i hi + rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) + · simp [ht.bv] + · simpa using hw.nth (by simpa using hi) |>.bv⟩ + +lemma Language.SemitermVec.cons {n m w t : V} (h : L.IsSemitermVec n m w) (ht : L.IsSemiterm m t) : L.IsSemitermVec (n + 1) m (t ∷ w) := + Language.IsSemitermVec.cons_iff.mpr ⟨ht, h⟩ + +@[simp] lemma Language.IsSemitermVec.singleton {n t : V} : + L.IsSemitermVec 1 n ?[t] ↔ L.IsSemiterm n t := by + rw [show (1 : V) = 0 + 1 by simp, Language.IsSemitermVec.cons_iff]; simp + +@[simp] lemma Language.IsSemitermVec.doubleton {n t₁ t₂ : V} : + L.IsSemitermVec 2 n ?[t₁, t₂] ↔ L.IsSemiterm n t₁ ∧ L.IsSemiterm n t₂ := by + rw [show (2 : V) = 1 + 1 by simp [one_add_one_eq_two], Language.IsSemitermVec.cons_iff]; simp + +section + +variable (L) + +def _root_.LO.FirstOrder.Arith.LDef.isSemitermDef (pL : LDef) : 𝚫₁.Semisentence 2 := .mkDelta + (.mkSigma “n p | !pL.isUTermDef.sigma p ∧ ∃ b, !pL.termBVDef b p ∧ b ≤ n” (by simp)) + (.mkPi “n p | !pL.isUTermDef.pi p ∧ ∀ b, !pL.termBVDef b p → b ≤ n” (by simp)) + +lemma Language.isSemiterm_defined : 𝚫₁-Relation L.IsSemiterm via pL.isSemitermDef where + left := by + intro v + simp [Arith.LDef.isSemitermDef, HierarchySymbol.Semiformula.val_sigma, + L.isUTerm_defined.df.iff, L.isUTerm_defined.proper.iff', + L.termBV_defined.df.iff] + right := by + intro v + simp [Arith.LDef.isSemitermDef, HierarchySymbol.Semiformula.val_sigma, + L.isUTerm_defined.df.iff, L.termBV_defined.df.iff]; rfl + +instance Language.isSemiterm_definable : 𝚫₁-Relation L.IsSemiterm := L.isSemiterm_defined.to_definable + +instance Language.isSemiterm_defined' (Γ m) : Γ-[m + 1]-Relation L.IsSemiterm := L.isSemiterm_definable.of_deltaOne + +def _root_.LO.FirstOrder.Arith.LDef.isSemitermVecDef (pL : LDef) : 𝚫₁.Semisentence 3 := .mkDelta + (.mkSigma “k n ps | !pL.isUTermVecDef.sigma k ps ∧ ∀ i < k, ∃ p, !nthDef p ps i ∧ ∃ b, !pL.termBVDef b p ∧ b ≤ n” (by simp)) + (.mkPi “k n ps | !pL.isUTermVecDef.pi k ps ∧ ∀ i < k, ∀ p, !nthDef p ps i → ∀ b, !pL.termBVDef b p → b ≤ n” (by simp)) + +lemma Language.isSemitermVec_defined : 𝚫₁-Relation₃ L.IsSemitermVec via pL.isSemitermVecDef where + left := by + intro v + simp [Arith.LDef.isSemitermVecDef, HierarchySymbol.Semiformula.val_sigma, + L.isUTermVec_defined.df.iff, L.isUTermVec_defined.proper.iff', + L.termBV_defined.df.iff] + right := by + intro v + simp [Arith.LDef.isSemitermVecDef, HierarchySymbol.Semiformula.val_sigma, + L.isUTermVec_defined.df.iff, L.termBV_defined.df.iff]; rfl + +instance Language.isSemitermVec_definable : 𝚫₁-Relation₃ L.IsSemitermVec := L.isSemitermVec_defined.to_definable + +instance Language.isSemitermVec_defined' (Γ m) : Γ-[m + 1]-Relation₃ L.IsSemitermVec := L.isSemitermVec_definable.of_deltaOne + +end + +lemma Language.IsSemiterm.case_iff {n t : V} : + L.IsSemiterm n t ↔ + (∃ z < n, t = ^#z) ∨ + (∃ x, t = ^&x) ∨ + (∃ k f v : V, L.Func k f ∧ L.IsSemitermVec k n v ∧ t = ^func k f v) := by + constructor + · intro h + rcases h.isUTerm.case with (⟨z, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hf, _, rfl⟩) + · left; exact ⟨z, by simpa [succ_le_iff_lt] using h.bv, rfl⟩ + · right; left; exact ⟨x, rfl⟩ + · right; right; exact ⟨k, f, v, hf, by simp_all, rfl⟩ + · rintro (⟨z, hz, rfl⟩ | ⟨x, rfl⟩ | ⟨k, f, v, hf, hv, rfl⟩) + · simp [hz] + · simp + · simp [hf, hv] + +alias ⟨Language.IsSemiterm.case, Language.IsSemiterm.mk⟩ := Language.IsSemiterm.case_iff + +lemma Language.IsSemiterm.induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) + (hbvar : ∀ z < n, P (^#z)) (hfvar : ∀ x, P (^&x)) + (hfunc : ∀ k f v, L.Func k f → L.IsSemitermVec k n v → (∀ i < k, P v.[i]) → P (^func k f v)) : + ∀ t, L.IsSemiterm n t → P t := by + intro t ht + suffices L.IsSemiterm n t → P t by exact this ht + apply Language.IsUTerm.induction Γ ?_ ?_ ?_ ?_ t ht.isUTerm + · definability + · intro z; simp only [bvar]; exact hbvar z + · intro x _; exact hfvar x + · intro k f v hf _ ih h + have hv : L.IsSemitermVec k n v := by simp_all + exact hfunc k f v hf hv (fun i hi ↦ ih i hi (hv.nth hi)) + +@[simp] lemma Language.IsSemitermVec.nil (k : V): L.IsSemitermVec 0 k 0 := ⟨by simp, by simp⟩ + +@[simp] lemma Language.IsSemitermVec.cons {k n w t : V} (h : L.IsSemitermVec n k w) (ht : L.IsSemiterm k t) : L.IsSemitermVec (n + 1) k (t ∷ w) := + ⟨h.isUTerm.cons ht.isUTerm, by + intro i hi + rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) + · simp [ht.bv] + · simp [h.bv (by simpa using hi)]⟩ + +end isSemiterm + end LO.Arith end diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean index ed17782..6fce960 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean @@ -14,17 +14,17 @@ variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] namespace TermSubst -def blueprint (pL : LDef) : Language.TermRec.Blueprint pL 2 where - bvar := .mkSigma “y n z m w | !nthDef y w z” (by simp) - fvar := .mkSigma “y n x m w | !qqFvarDef y x” (by simp) - func := .mkSigma “y n k f v v' m w | !qqFuncDef y k f v'” (by simp) +def blueprint (pL : LDef) : Language.TermRec.Blueprint pL 1 where + bvar := .mkSigma “y z w | !nthDef y w z” (by simp) + fvar := .mkSigma “y x w | !qqFvarDef y x” (by simp) + func := .mkSigma “y k f v v' w | !qqFuncDef y k f v'” (by simp) variable (L) def construction : Language.TermRec.Construction V L (blueprint pL) where - bvar (param _ z) := (param 1).[z] - fvar (_ _ x) := ^&x - func (_ _ k f _ v') := ^func k f v' + bvar (param z) := (param 1).[z] + fvar (_ x) := ^&x + func (_ k f _ v') := ^func k f v' bvar_defined := by intro v; simp [blueprint]; rfl fvar_defined := by intro v; simp [blueprint] func_defined := by intro v; simp [blueprint]; rfl @@ -37,153 +37,143 @@ open TermSubst variable (L) -def Language.termSubst (n m w t : V) : V := (construction L).result ![m, w] n t +def Language.termSubst (w t : V) : V := (construction L).result ![w] t -def Language.termSubstVec (k n m w v : V) : V := (construction L).resultVec ![m, w] k n v +def Language.termSubstVec (k w v : V) : V := (construction L).resultVec ![w] k v variable {L} variable {n m w : V} -@[simp] lemma termSubst_bvar {z} (hz : z < n) : - L.termSubst n m w ^#z = w.[z] := by simp [Language.termSubst, hz, construction] +@[simp] lemma Language.termSubst_bvar (z) : + L.termSubst w ^#z = w.[z] := by simp [Language.termSubst, construction] -@[simp] lemma termSubst_fvar (x) : - L.termSubst n m w ^&x = ^&x := by simp [Language.termSubst, construction] +@[simp] lemma Language.termSubst_fvar (x) : + L.termSubst w ^&x = ^&x := by simp [Language.termSubst, construction] -@[simp] lemma termSubst_func {k f v} (hkf : L.Func k f) (hv : L.SemitermVec k n v) : - L.termSubst n m w (^func k f v) = ^func k f (L.termSubstVec k n m w v) := by +@[simp] lemma Language.termSubst_func {k f v} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + L.termSubst w (^func k f v) = ^func k f (L.termSubstVec k w v) := by simp [Language.termSubst, construction, hkf, hv]; rfl section -def _root_.LO.FirstOrder.Arith.LDef.termSubstDef (pL : LDef) : 𝚺₁.Semisentence 5 := (blueprint pL).result.rew <| Rew.substs ![#0, #1, #4, #2, #3] +def _root_.LO.FirstOrder.Arith.LDef.termSubstDef (pL : LDef) : 𝚺₁.Semisentence 3 := (blueprint pL).result.rew <| Rew.substs ![#0, #2, #1] -def _root_.LO.FirstOrder.Arith.LDef.termSubstVecDef (pL : LDef) : 𝚺₁.Semisentence 6 := (blueprint pL).resultVec.rew <| Rew.substs ![#0, #1, #2, #5, #3, #4] +def _root_.LO.FirstOrder.Arith.LDef.termSubstVecDef (pL : LDef) : 𝚺₁.Semisentence 4 := (blueprint pL).resultVec.rew <| Rew.substs ![#0, #1, #3, #2] variable (L) -lemma termSubst_defined : 𝚺₁.DefinedFunction (fun v ↦ L.termSubst (v 0) (v 1) (v 2) (v 3)) pL.termSubstDef := by - intro v; simpa [LDef.termSubstDef, Language.termSubst] using (construction L).result_defined ![v 0, v 1, v 4, v 2, v 3] +lemma Language.termSubst_defined : 𝚺₁-Function₂ L.termSubst via pL.termSubstDef := by + intro v; simpa [LDef.termSubstDef, Language.termSubst] using (construction L).result_defined ![v 0, v 2, v 1] -@[simp] lemma eval_termSubstDef (v : Fin 5 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termSubstDef ↔ v 0 = L.termSubst (v 1) (v 2) (v 3) (v 4) := (termSubst_defined L).df.iff v +instance Language.termSubst_definable : 𝚺₁-Function₂ L.termSubst := (termSubst_defined L).to_definable -instance termSubst_definable : 𝚺₁.BoldfaceFunction (fun v : Fin 4 → V ↦ L.termSubst (v 0) (v 1) (v 2) (v 3)) := - (termSubst_defined L).to_definable +instance Language.termSubst_definable' : Γ-[k + 1]-Function₂ L.termSubst := L.termSubst_definable.of_sigmaOne -instance termSubst_definable₂ (n m : V) : 𝚺₁-Function₂ (L.termSubst n m) := by - simpa using HierarchySymbol.BoldfaceFunction.retractiont (n := 2) (termSubst_definable L) ![&n, &m, #0, #1] +lemma Language.termSubstVec_defined : 𝚺₁-Function₃ L.termSubstVec via pL.termSubstVecDef := by + intro v; simpa [LDef.termSubstVecDef, Language.termSubstVec] using (construction L).resultVec_defined ![v 0, v 1, v 3, v 2] -@[simp, definability] instance termSubst_definable₂' (Γ k) (n m : V) : Γ-[k + 1]-Function₂ (L.termSubst n m) := - .of_sigmaOne (termSubst_definable₂ L n m) _ _ +instance Language.termSubstVec_definable : 𝚺₁-Function₃ L.termSubstVec := L.termSubstVec_defined.to_definable -lemma termSubstVec_defined : 𝚺₁.DefinedFunction (fun v ↦ L.termSubstVec (v 0) (v 1) (v 2) (v 3) (v 4)) pL.termSubstVecDef := by - intro v; simpa [LDef.termSubstVecDef, Language.termSubstVec] using (construction L).resultVec_defined ![v 0, v 1, v 2, v 5, v 3, v 4] - -@[simp] lemma eval_termSubstVecDef (v : Fin 6 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termSubstVecDef ↔ v 0 = L.termSubstVec (v 1) (v 2) (v 3) (v 4) (v 5) := (termSubstVec_defined L).df.iff v - -instance termSubstVec_definable : 𝚺₁.BoldfaceFunction (fun v : Fin 5 → V ↦ L.termSubstVec (v 0) (v 1) (v 2) (v 3) (v 4)) := - (termSubstVec_defined L).to_definable - -instance termSubstVec_definable₂ (k n m : V) : 𝚺₁-Function₂ (L.termSubstVec k n m) := by - simpa using HierarchySymbol.BoldfaceFunction.retractiont (n := 2) (termSubstVec_definable L) ![&k, &n, &m, #0, #1] - -@[simp, definability] instance termSubstVec_definable₂' (Γ i) (k n m : V) : Γ-[i + 1]-Function₂ (L.termSubstVec k n m) := - .of_sigmaOne (termSubstVec_definable₂ L k n m) _ _ +instance Language.termSubstVec_definable' : Γ-[i + 1]-Function₃ L.termSubstVec := L.termSubstVec_definable.of_sigmaOne end -@[simp] lemma len_termSubstVec {k n ts : V} (hts : L.SemitermVec k n ts) : - len (L.termSubstVec k n m w ts) = k := (construction L).resultVec_lh _ hts +@[simp] lemma len_termSubstVec {k ts : V} (hts : L.IsUTermVec k ts) : + len (L.termSubstVec k w ts) = k := (construction L).resultVec_lh _ hts -@[simp] lemma nth_termSubstVec {k n ts i : V} (hts : L.SemitermVec k n ts) (hi : i < k) : - (L.termSubstVec k n m w ts).[i] = L.termSubst n m w ts.[i] := +@[simp] lemma nth_termSubstVec {k ts i : V} (hts : L.IsUTermVec k ts) (hi : i < k) : + (L.termSubstVec k w ts).[i] = L.termSubst w ts.[i] := (construction L).nth_resultVec _ hts hi -@[simp] lemma termSubstVec_nil (n : V) : L.termSubstVec 0 n m w 0 = 0 := - (construction L).resultVec_nil _ _ +@[simp] lemma termSubstVec_nil (w : V) : L.termSubstVec 0 w 0 = 0 := + (construction L).resultVec_nil _ -lemma termSubstVec_cons {k n t ts : V} (ht : L.Semiterm n t) (hts : L.SemitermVec k n ts) : - L.termSubstVec (k + 1) n m w (t ∷ ts) = L.termSubst n m w t ∷ L.termSubstVec k n m w ts := - (construction L).resultVec_cons ![m, w] hts ht +lemma termSubstVec_cons {k t ts : V} (ht : L.IsUTerm t) (hts : L.IsUTermVec k ts) : + L.termSubstVec (k + 1) w (t ∷ ts) = L.termSubst w t ∷ L.termSubstVec k w ts := + (construction L).resultVec_cons ![w] hts ht -@[simp] lemma termSubstVec_cons₁ {n t : V} (ht : L.Semiterm n t) : - L.termSubstVec 1 n m w ?[t] = ?[L.termSubst n m w t] := by +@[simp] lemma termSubstVec_cons₁ {t : V} (ht : L.IsUTerm t) : + L.termSubstVec 1 w ?[t] = ?[L.termSubst w t] := by rw [show (1 : V) = 0 + 1 by simp, termSubstVec_cons] <;> simp [*] -@[simp] lemma termSubstVec_cons₂ {n t₁ t₂ : V} (ht₁ : L.Semiterm n t₁) (ht₂ : L.Semiterm n t₂) : - L.termSubstVec 2 n m w ?[t₁, t₂] = ?[L.termSubst n m w t₁, L.termSubst n m w t₂] := by +@[simp] lemma termSubstVec_cons₂ {t₁ t₂ : V} (ht₁ : L.IsUTerm t₁) (ht₂ : L.IsUTerm t₂) : + L.termSubstVec 2 w ?[t₁, t₂] = ?[L.termSubst w t₁, L.termSubst w t₂] := by rw [show (2 : V) = 0 + 1 + 1 by simp [one_add_one_eq_two], termSubstVec_cons] <;> simp [*] -@[simp] lemma termSubst_rng_semiterm {t} (hw : L.SemitermVec n m w) (ht : L.Semiterm n t) : L.Semiterm m (L.termSubst n m w t) := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +@[simp] lemma Language.IsSemitermVec.termSubst {t} (hw : L.IsSemitermVec n m w) (ht : L.IsSemiterm n t) : L.IsSemiterm m (L.termSubst w t) := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability - · intro z hz; simp [hz, hw.prop hz] + · intro z hz; simpa using hw.nth hz · intro x; simp · intro k f v hkf hv ih - simp only [hkf, hv, termSubst_func, Language.Semiterm.func_iff, true_and] - exact ⟨by simp [Language.termSubstVec, hv], fun i hi ↦ by - rw [nth_termSubstVec hv hi] - exact ih i hi⟩ - -@[simp] lemma Language.SemitermVec.termSubstVec {k n m v} (hw : L.SemitermVec n m w) (hv : L.SemitermVec k n v) : - L.SemitermVec k m (L.termSubstVec k n m w v) := - ⟨by simp [Language.termSubstVec, hv], fun i hi ↦ by - rw [nth_termSubstVec hv hi] - exact termSubst_rng_semiterm hw (hv.prop hi)⟩ - -@[simp] lemma substs_nil {t} (ht : L.Semiterm 0 t) : L.termSubst 0 0 0 t = t := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + simp only [hkf, hv.isUTerm, Language.termSubst_func, Language.IsSemiterm.func, true_and] + exact Language.IsSemitermVec.iff.mpr + ⟨by simp [hv.isUTerm], fun i hi ↦ by rw [nth_termSubstVec hv.isUTerm hi]; exact ih i hi⟩ + +@[simp] lemma Language.IsUTermVec.termSubst {t} (hw : L.IsUTermVec n w) (ht : L.IsSemiterm n t) : L.IsUTerm (L.termSubst w t) := + Language.IsSemitermVec.termSubst hw.isSemitermVec ht |>.isUTerm + +@[simp] lemma Language.IsSemitermVec.termSubstVec {k n m v} (hw : L.IsSemitermVec n m w) (hv : L.IsSemitermVec k n v) : + L.IsSemitermVec k m (L.termSubstVec k w v) := Language.IsSemitermVec.iff.mpr + ⟨by simp [hv.isUTerm], fun i hi ↦ by rw [nth_termSubstVec hv.isUTerm hi]; exact hw.termSubst (hv.nth hi)⟩ + +@[simp] lemma substs_nil {t} (ht : L.IsSemiterm 0 t) : L.termSubst 0 t = t := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z; simp · intro x; simp · intro k f v hf hv ih - simp only [hf, hv, termSubst_func, qqFunc_inj, true_and] - apply nth_ext' k (by simp [hv]) (by simp [hv.1]) + simp only [hf, hv.isUTerm, Language.termSubst_func, qqFunc_inj, true_and] + apply nth_ext' k (by simp [hv.isUTerm]) (by simp [hv.lh]) intro i hi - simp [nth_termSubstVec hv hi, ih i hi] + simp [nth_termSubstVec hv.isUTerm hi, ih i hi] -lemma termSubst_termSubst {l n m w v t : V} (hv : L.SemitermVec l n v) (ht : L.Semiterm l t) : - L.termSubst n m w (L.termSubst l n v t) = L.termSubst l m (L.termSubstVec l n m w v) t := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +lemma termSubst_termSubst {l n w v t : V} (hv : L.IsSemitermVec l n v) (ht : L.IsSemiterm l t) : + L.termSubst w (L.termSubst v t) = L.termSubst (L.termSubstVec l w v) t := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability - · intro z hz; simp [hz, hv] + · intro z hz + rw [Language.termSubst_bvar z, Language.termSubst_bvar z, nth_termSubstVec hv.isUTerm hz] · intro x; simp [hv] · intro k f ts hf hts ih - simp only [termSubst_func, Language.SemitermVec.termSubstVec, qqFunc_inj, true_and, hf, hts, hv] - apply nth_ext' k (by simp [hv, hts]) (by simp [hts]) + rw [Language.termSubst_func hf hts.isUTerm, + Language.termSubst_func hf (hv.termSubstVec hts).isUTerm, + Language.termSubst_func hf hts.isUTerm] + simp only [qqFunc_inj, true_and] + apply nth_ext' k (by rw [len_termSubstVec (hv.termSubstVec hts).isUTerm]) (by rw [len_termSubstVec hts.isUTerm]) intro i hi - rw [nth_termSubstVec (hv.termSubstVec hts) hi, nth_termSubstVec hts hi, nth_termSubstVec hts hi, ih i hi] + rw [nth_termSubstVec (hv.termSubstVec hts).isUTerm hi, + nth_termSubstVec hts.isUTerm hi, nth_termSubstVec hts.isUTerm hi, ih i hi] -lemma termSubst_eq_self {n m w t : V} (ht : L.Semiterm n t) (H : ∀ i < n, w.[i] = ^#i) : - L.termSubst n m w t = t := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +lemma termSubst_eq_self {n w t : V} (ht : L.IsSemiterm n t) (H : ∀ i < n, w.[i] = ^#i) : + L.termSubst w t = t := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z hz; simp [hz, H] · intro x; simp · intro k f v hf hv ih - simp only [termSubst_func, qqFunc_inj, true_and, hf, hv] - apply nth_ext' k (by simp [*]) (by simp [hv.1]) + rw [Language.termSubst_func hf hv.isUTerm] + simp only [qqFunc_inj, true_and] + apply nth_ext' k (by rw [len_termSubstVec hv.isUTerm]) (by simp [hv.lh]) intro i hi - rw [nth_termSubstVec hv hi, ih i hi] + rw [nth_termSubstVec hv.isUTerm hi, ih i hi] end termSubst namespace TermShift def blueprint (pL : LDef) : Language.TermRec.Blueprint pL 0 where - bvar := .mkSigma “y n z m w | !qqBvarDef y z” (by simp) - fvar := .mkSigma “y n x m w | !qqFvarDef y (x + 1)” (by simp) - func := .mkSigma “y n k f v v' m w | !qqFuncDef y k f v'” (by simp) + bvar := .mkSigma “y z | !qqBvarDef y z” (by simp) + fvar := .mkSigma “y x | !qqFvarDef y (x + 1)” (by simp) + func := .mkSigma “y k f v v' | !qqFuncDef y k f v'” (by simp) variable (L) def construction : Language.TermRec.Construction V L (blueprint pL) where - bvar (_ _ z) := ^#z - fvar (_ _ x) := ^&(x + 1) - func (_ _ k f _ v') := ^func k f v' + bvar (_ z) := ^#z + fvar (_ x) := ^&(x + 1) + func (_ k f _ v') := ^func k f v' bvar_defined := by intro v; simp [blueprint] fvar_defined := by intro v; simp [blueprint] func_defined := by intro v; simp [blueprint]; rfl @@ -196,111 +186,133 @@ open TermShift variable (L) -def Language.termShift (n t : V) : V := (construction L).result ![] n t +def Language.termShift (t : V) : V := (construction L).result ![] t -def Language.termShiftVec (k n v : V) : V := (construction L).resultVec ![] k n v +def Language.termShiftVec (k v : V) : V := (construction L).resultVec ![] k v variable {L} variable {n : V} -@[simp] lemma termShift_bvar {z} (hz : z < n) : - L.termShift n ^#z = ^#z := by simp [Language.termShift, hz, construction] +@[simp] lemma Language.termShift_bvar (z) : + L.termShift ^#z = ^#z := by simp [Language.termShift, construction] -@[simp] lemma termShift_fvar (x) : - L.termShift n ^&x = ^&(x + 1) := by simp [Language.termShift, construction] +@[simp] lemma Language.termShift_fvar (x) : + L.termShift ^&x = ^&(x + 1) := by simp [Language.termShift, construction] -@[simp] lemma termShift_func {k f v} (hkf : L.Func k f) (hv : L.SemitermVec k n v) : - L.termShift n (^func k f v) = ^func k f (L.termShiftVec k n v) := by +@[simp] lemma Language.termShift_func {k f v} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + L.termShift (^func k f v) = ^func k f (L.termShiftVec k v) := by simp [Language.termShift, construction, hkf, hv]; rfl section -def _root_.LO.FirstOrder.Arith.LDef.termShiftDef (pL : LDef) : 𝚺₁.Semisentence 3 := - (blueprint pL).result +def _root_.LO.FirstOrder.Arith.LDef.termShiftDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result -def _root_.LO.FirstOrder.Arith.LDef.termShiftVecDef (pL : LDef) : 𝚺₁.Semisentence 4 := (blueprint pL).resultVec +def _root_.LO.FirstOrder.Arith.LDef.termShiftVecDef (pL : LDef) : 𝚺₁.Semisentence 3 := (blueprint pL).resultVec variable (L) -lemma termShift_defined : 𝚺₁-Function₂ L.termShift via pL.termShiftDef := by +lemma Language.termShift_defined : 𝚺₁-Function₁ L.termShift via pL.termShiftDef := by intro v; simpa [LDef.termShiftDef, Language.termShift] using (construction L).result_defined v -@[simp] lemma eval_termShiftDef (v : Fin 3 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termShiftDef ↔ v 0 = L.termShift (v 1) (v 2) := (termShift_defined L).df.iff v - -instance termShift_definable : 𝚺₁-Function₂ L.termShift := - (termShift_defined L).to_definable +instance Language.termShift_definable : 𝚺₁-Function₁ L.termShift := L.termShift_defined.to_definable -@[definability, simp] instance termShift_definable' (Γ i) : Γ-[i + 1]-Function₂ L.termShift := .of_sigmaOne (termShift_definable L) _ _ +instance Language.termShift_definable' : Γ-[i + 1]-Function₁ L.termShift := L.termShift_definable.of_sigmaOne -lemma termShiftVec_defined : 𝚺₁-Function₃ L.termShiftVec via pL.termShiftVecDef := by +lemma Language.termShiftVec_defined : 𝚺₁-Function₂ L.termShiftVec via pL.termShiftVecDef := by intro v; simpa [LDef.termShiftVecDef, Language.termShiftVec] using (construction L).resultVec_defined v -@[simp] lemma eval_termShiftVecDef (v : Fin 4 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termShiftVecDef ↔ v 0 = L.termShiftVec (v 1) (v 2) (v 3) := (termShiftVec_defined L).df.iff v +instance Language.termShiftVec_definable : 𝚺₁-Function₂ L.termShiftVec := L.termShiftVec_defined.to_definable -instance termShiftVec_definable : 𝚺₁-Function₃ L.termShiftVec := - (termShiftVec_defined L).to_definable - -@[simp, definability] instance termShiftVec_definable' (Γ i) : Γ-[i + 1]-Function₃ L.termShiftVec := - .of_sigmaOne (termShiftVec_definable L) _ _ +instance Language.termShiftVec_definable' : Γ-[i + 1]-Function₂ L.termShiftVec := L.termShiftVec_definable.of_sigmaOne end -@[simp] lemma len_termShiftVec {k n ts : V} (hts : L.SemitermVec k n ts) : - len (L.termShiftVec k n ts) = k := (construction L).resultVec_lh _ hts +@[simp] lemma len_termShiftVec {k ts : V} (hts : L.IsUTermVec k ts) : + len (L.termShiftVec k ts) = k := (construction L).resultVec_lh _ hts -@[simp] lemma nth_termShiftVec {k n ts i : V} (hts : L.SemitermVec k n ts) (hi : i < k) : - (L.termShiftVec k n ts).[i] = L.termShift n ts.[i] := - (construction L).nth_resultVec _ hts hi +@[simp] lemma nth_termShiftVec {k ts i : V} (hts : L.IsUTermVec k ts) (hi : i < k) : + (L.termShiftVec k ts).[i] = L.termShift ts.[i] := (construction L).nth_resultVec _ hts hi -@[simp] lemma termShiftVec_nil (n : V) : L.termShiftVec 0 n 0 = 0 := - (construction L).resultVec_nil ![] _ +@[simp] lemma termShiftVec_nil : L.termShiftVec 0 0 = 0 := (construction L).resultVec_nil ![] -lemma termShiftVec_cons {k n t ts : V} (ht : L.Semiterm n t) (hts : L.SemitermVec k n ts) : - L.termShiftVec (k + 1) n (t ∷ ts) = L.termShift n t ∷ L.termShiftVec k n ts := +lemma termShiftVec_cons {k t ts : V} (ht : L.IsUTerm t) (hts : L.IsUTermVec k ts) : + L.termShiftVec (k + 1) (t ∷ ts) = L.termShift t ∷ L.termShiftVec k ts := (construction L).resultVec_cons ![] hts ht -@[simp] lemma termShiftVec_cons₁ {n t₁ : V} (ht₁ : L.Semiterm n t₁) : - L.termShiftVec 1 n ?[t₁] = ?[L.termShift n t₁] := by +@[simp] lemma termShiftVec_cons₁ {t₁ : V} (ht₁ : L.IsUTerm t₁) : + L.termShiftVec 1 ?[t₁] = ?[L.termShift t₁] := by rw [show (1 : V) = 0 + 1 by simp, termShiftVec_cons] <;> simp [*] -@[simp] lemma termShiftVec_cons₂ {n t₁ t₂ : V} (ht₁ : L.Semiterm n t₁) (ht₂ : L.Semiterm n t₂) : - L.termShiftVec 2 n ?[t₁, t₂] = ?[L.termShift n t₁, L.termShift n t₂] := by - rw [show (2 : V) = 0 + 1 + 1 by simp [one_add_one_eq_two], termShiftVec_cons] <;> simp [*] +@[simp] lemma termShiftVec_cons₂ {t₁ t₂ : V} (ht₁ : L.IsUTerm t₁) (ht₂ : L.IsUTerm t₂) : + L.termShiftVec 2 ?[t₁, t₂] = ?[L.termShift t₁, L.termShift t₂] := by + rw [show (2 : V) = 0 + 1 + 1 by simp [one_add_one_eq_two], termShiftVec_cons] <;> simp [ht₁, ht₂] + +@[simp] lemma Language.IsUTerm.termShift {t} (ht : L.IsUTerm t) : L.IsUTerm (L.termShift t) := by + apply Language.IsUTerm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + · definability + · intro z; simp + · intro x; simp + · intro k f v hkf hv ih; + simp only [hkf, hv, termShift_func, func_iff, true_and] + exact ⟨by simp [hv], by intro i hi; rw [nth_termShiftVec hv hi]; exact ih i hi⟩ -@[simp] lemma Language.Semiterm.termShift {t} (ht : L.Semiterm n t) : L.Semiterm n (L.termShift n t) := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +@[simp] lemma Language.IsSemiterm.termShift {t} (ht : L.IsSemiterm n t) : L.IsSemiterm n (L.termShift t) := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z hz; simp [hz] · intro x; simp · intro k f v hkf hv ih; - simp only [hkf, hv, termShift_func, Language.Semiterm.func_iff, true_and] - exact ⟨by simp [Language.termShiftVec, hv], fun i hi ↦ by - rw [nth_termShiftVec hv hi] - exact ih i hi⟩ + simp only [hkf, hv.isUTerm, termShift_func, func, true_and] + refine Language.IsSemitermVec.iff.mpr ⟨?_, ?_⟩ + · simp [termShiftVec, hv.isUTerm] + · intro i hi + rw [nth_termShiftVec hv.isUTerm hi] + exact ih i hi + +@[simp] lemma Language.IsUTermVec.termShiftVec {k v} (hv : L.IsUTermVec k v) : L.IsUTermVec k (L.termShiftVec k v) := + ⟨by simp [termShiftVec, hv], fun i hi ↦ by rw [nth_termShiftVec hv hi]; exact (hv.nth hi).termShift⟩ + +@[simp] lemma Language.IsSemitermVec.termShiftVec {k n v} (hv : L.IsSemitermVec k n v) : L.IsSemitermVec k n (L.termShiftVec k v) := + Language.IsSemitermVec.iff.mpr + ⟨by simp [termShiftVec, hv.isUTerm], fun i hi ↦ by + rw [nth_termShiftVec hv.isUTerm hi]; exact (hv.nth hi).termShift⟩ + +@[simp] lemma Language.IsUTerm.termBVtermShift {t} (ht : L.IsUTerm t) : L.termBV (L.termShift t) = L.termBV t := by + apply Language.IsUTerm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + · definability + · simp + · simp + · intro k f v hf hv ih + rw [termShift_func hf hv, + termBV_func hf hv.termShiftVec, + termBV_func hf hv] + congr 1 + apply nth_ext' k (by simp [*]) (by simp [*]) + intro i hi + simp [*] -@[simp] lemma Language.SemitermVec.termShiftVec {k n v} (hv : L.SemitermVec k n v) : L.SemitermVec k n (L.termShiftVec k n v) := - ⟨by simp [Language.termShiftVec, hv], fun i hi ↦ by - rw [nth_termShiftVec hv hi] - exact (hv.prop hi).termShift⟩ +@[simp] lemma Language.IsUTermVec.termBVVectermShiftVec {v} (hv : L.IsUTermVec k v) : + L.termBVVec k (L.termShiftVec k v) = L.termBVVec k v := by + apply nth_ext' k (by simp [*]) (by simp [*]) + intro i hi + simp [*, Language.IsUTerm.termBVtermShift (hv.nth hi)] end termShift namespace TermBShift def blueprint (pL : LDef) : Language.TermRec.Blueprint pL 0 where - bvar := .mkSigma “y n z m w | !qqBvarDef y (z + 1)” (by simp) - fvar := .mkSigma “y n x m w | !qqFvarDef y x” (by simp) - func := .mkSigma “y n k f v v' m w | !qqFuncDef y k f v'” (by simp) + bvar := .mkSigma “y z | !qqBvarDef y (z + 1)” (by simp) + fvar := .mkSigma “y x | !qqFvarDef y x” (by simp) + func := .mkSigma “y k f v v' | !qqFuncDef y k f v'” (by simp) variable (L) def construction : Language.TermRec.Construction V L (blueprint pL) where - bvar (_ _ z) := ^#(z + 1) - fvar (_ _ x) := ^&x - func (_ _ k f _ v') := ^func k f v' + bvar (_ z) := ^#(z + 1) + fvar (_ x) := ^&x + func (_ k f _ v') := ^func k f v' bvar_defined := by intro v; simp [blueprint] fvar_defined := by intro v; simp [blueprint] func_defined := by intro v; simp [blueprint]; rfl @@ -313,213 +325,236 @@ open TermBShift variable (L) -def Language.termBShift (n t : V) : V := (construction L).result ![] n t +def Language.termBShift (t : V) : V := (construction L).result ![] t -def Language.termBShiftVec (k n v : V) : V := (construction L).resultVec ![] k n v +def Language.termBShiftVec (k v : V) : V := (construction L).resultVec ![] k v variable {L} -variable {n : V} - -@[simp] lemma termBShift_bvar {z} (hz : z < n) : - L.termBShift n ^#z = ^#(z + 1) := by simp [Language.termBShift, hz, construction] +@[simp] lemma Language.termBShift_bvar (z) : + L.termBShift ^#z = ^#(z + 1) := by simp [Language.termBShift, construction] -@[simp] lemma termBShift_fvar (x) : - L.termBShift n ^&x = ^&x := by simp [Language.termBShift, construction] +@[simp] lemma Language.termBShift_fvar (x) : + L.termBShift ^&x = ^&x := by simp [Language.termBShift, construction] -@[simp] lemma termBShift_func {k f v} (hkf : L.Func k f) (hv : L.SemitermVec k n v) : - L.termBShift n (^func k f v) = ^func k f (L.termBShiftVec k n v) := by +@[simp] lemma Language.termBShift_func {k f v} (hkf : L.Func k f) (hv : L.IsUTermVec k v) : + L.termBShift (^func k f v) = ^func k f (L.termBShiftVec k v) := by simp [Language.termBShift, construction, hkf, hv]; rfl section -def _root_.LO.FirstOrder.Arith.LDef.termBShiftDef (pL : LDef) : 𝚺₁.Semisentence 3 := - (blueprint pL).result +def _root_.LO.FirstOrder.Arith.LDef.termBShiftDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result -def _root_.LO.FirstOrder.Arith.LDef.termBShiftVecDef (pL : LDef) : 𝚺₁.Semisentence 4 := (blueprint pL).resultVec +def _root_.LO.FirstOrder.Arith.LDef.termBShiftVecDef (pL : LDef) : 𝚺₁.Semisentence 3 := (blueprint pL).resultVec variable (L) -lemma termBShift_defined : 𝚺₁-Function₂ L.termBShift via pL.termBShiftDef := by +lemma Language.termBShift_defined : 𝚺₁-Function₁ L.termBShift via pL.termBShiftDef := by intro v; simpa using (construction L).result_defined v -@[simp] lemma eval_termBShiftDef (v : Fin 3 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termBShiftDef ↔ v 0 = L.termBShift (v 1) (v 2) := (termBShift_defined L).df.iff v - -instance termBShift_definable : 𝚺₁-Function₂ L.termBShift := - (termBShift_defined L).to_definable +instance Language.termBShift_definable : 𝚺₁-Function₁ L.termBShift := L.termBShift_defined.to_definable -@[definability, simp] instance termBShift_definable' (Γ i) : Γ-[i + 1]-Function₂ L.termBShift := .of_sigmaOne (termBShift_definable L) _ _ +instance Language.termBShift_definable' : Γ-[i + 1]-Function₁ L.termBShift := L.termBShift_definable.of_sigmaOne -lemma termBShiftVec_defined : 𝚺₁-Function₃ L.termBShiftVec via pL.termBShiftVecDef := by +lemma Language.termBShiftVec_defined : 𝚺₁-Function₂ L.termBShiftVec via pL.termBShiftVecDef := by intro v; simpa using (construction L).resultVec_defined v -@[simp] lemma eval_termBShiftVecDef (v : Fin 4 → V) : - Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termBShiftVecDef ↔ v 0 = L.termBShiftVec (v 1) (v 2) (v 3) := (termBShiftVec_defined L).df.iff v +instance Language.termBShiftVec_definable : 𝚺₁-Function₂ L.termBShiftVec := L.termBShiftVec_defined.to_definable -instance termBShiftVec_definable : 𝚺₁-Function₃ L.termBShiftVec := - (termBShiftVec_defined L).to_definable - -@[simp, definability] instance termBShiftVec_definable' (Γ i) : Γ-[i + 1]-Function₃ L.termBShiftVec := - .of_sigmaOne (termBShiftVec_definable L) _ _ +instance Language.termBShiftVec_definable' : Γ-[i + 1]-Function₂ L.termBShiftVec := L.termBShiftVec_definable.of_sigmaOne end -@[simp] lemma len_termBShiftVec {k n ts : V} (hts : L.SemitermVec k n ts) : - len (L.termBShiftVec k n ts) = k := (construction L).resultVec_lh _ hts +@[simp] lemma len_termBShiftVec {k ts : V} (hts : L.IsUTermVec k ts) : + len (L.termBShiftVec k ts) = k := (construction L).resultVec_lh _ hts -@[simp] lemma nth_termBShiftVec {k n ts i : V} (hts : L.SemitermVec k n ts) (hi : i < k) : - (L.termBShiftVec k n ts).[i] = L.termBShift n ts.[i] := +@[simp] lemma nth_termBShiftVec {k ts i : V} (hts : L.IsUTermVec k ts) (hi : i < k) : + (L.termBShiftVec k ts).[i] = L.termBShift ts.[i] := (construction L).nth_resultVec _ hts hi -@[simp] lemma termBShiftVec_nil (n : V) : L.termBShiftVec 0 n 0 = 0 := - (construction L).resultVec_nil ![] _ +@[simp] lemma termBShiftVec_nil : L.termBShiftVec 0 0 = 0 := + (construction L).resultVec_nil ![] -lemma termBShiftVec_cons {k n t ts : V} (ht : L.Semiterm n t) (hts : L.SemitermVec k n ts) : - L.termBShiftVec (k + 1) n (t ∷ ts) = L.termBShift n t ∷ L.termBShiftVec k n ts := +lemma termBShiftVec_cons {k t ts : V} (ht : L.IsUTerm t) (hts : L.IsUTermVec k ts) : + L.termBShiftVec (k + 1) (t ∷ ts) = L.termBShift t ∷ L.termBShiftVec k ts := (construction L).resultVec_cons ![] hts ht -@[simp] lemma termBShiftVec_cons₁ {n t₁ : V} (ht₁ : L.Semiterm n t₁) : - L.termBShiftVec 1 n ?[t₁] = ?[L.termBShift n t₁] := by +@[simp] lemma termBShiftVec_cons₁ {t₁ : V} (ht₁ : L.IsUTerm t₁) : + L.termBShiftVec 1 ?[t₁] = ?[L.termBShift t₁] := by rw [show (1 : V) = 0 + 1 by simp, termBShiftVec_cons] <;> simp [*] -@[simp] lemma termBShiftVec_cons₂ {n t₁ t₂ : V} (ht₁ : L.Semiterm n t₁) (ht₂ : L.Semiterm n t₂) : - L.termBShiftVec 2 n ?[t₁, t₂] = ?[L.termBShift n t₁, L.termBShift n t₂] := by +@[simp] lemma termBShiftVec_cons₂ {t₁ t₂ : V} (ht₁ : L.IsUTerm t₁) (ht₂ : L.IsUTerm t₂) : + L.termBShiftVec 2 ?[t₁, t₂] = ?[L.termBShift t₁, L.termBShift t₂] := by rw [show (2 : V) = 0 + 1 + 1 by simp [one_add_one_eq_two], termBShiftVec_cons] <;> simp [*] -@[simp] lemma Language.Semiterm.termBShift {t} (ht : L.Semiterm n t) : L.Semiterm (n + 1) (L.termBShift n t) := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +@[simp] lemma Language.IsSemiterm.termBShift {t} (ht : L.IsSemiterm n t) : L.IsSemiterm (n + 1) (L.termBShift t) := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z hz; simp [hz] · intro x; simp · intro k f v hkf hv ih; - simp only [hkf, hv, termBShift_func, Language.Semiterm.func_iff, true_and] - exact ⟨by simp [Language.termBShiftVec, hv], fun i hi ↦ by - rw [nth_termBShiftVec hv hi] - exact ih i hi⟩ - -@[simp] lemma Language.SemitermVec.termBShiftVec {k n v} (hv : L.SemitermVec k n v) : L.SemitermVec k (n + 1) (L.termBShiftVec k n v) := - ⟨by simp [Language.termBShiftVec, hv], fun i hi ↦ by - rw [nth_termBShiftVec hv hi] - exact (hv.prop hi).termBShift⟩ - -lemma termBShift_termShift {t} (ht : L.Semiterm n t) : L.termBShift n (L.termShift n t) = L.termShift (n + 1) (L.termBShift n t) := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + simp only [hkf, hv.isUTerm, termBShift_func, func, true_and] + refine Language.IsSemitermVec.iff.mpr ⟨?_, ?_⟩ + · simp [hv.isUTerm] + · intro i hi + rw [nth_termBShiftVec hv.isUTerm hi] + exact ih i hi + +@[simp] lemma Language.IsSemitermVec.termBShiftVec {k n v} (hv : L.IsSemitermVec k n v) : L.IsSemitermVec k (n + 1) (L.termBShiftVec k v) := + Language.IsSemitermVec.iff.mpr + ⟨by simp [hv.isUTerm], fun i hi ↦ by + rw [nth_termBShiftVec hv.isUTerm hi]; exact (hv.nth hi).termBShift⟩ + +lemma termBShift_termShift {t} (ht : L.IsSemiterm n t) : L.termBShift (L.termShift t) = L.termShift (L.termBShift t) := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z hz; simp [hz] · intro x; simp · intro k f v hkf hv ih - simp only [termShift_func, Language.SemitermVec.termShiftVec, termBShift_func, - Language.SemitermVec.termBShiftVec, qqFunc_inj, true_and, hkf, hv] - apply nth_ext' k (by simp [hv]) (by simp [hv]) + rw [L.termShift_func hkf hv.isUTerm, + L.termBShift_func hkf hv.termShiftVec.isUTerm, + L.termBShift_func hkf hv.isUTerm, + L.termShift_func hkf hv.termBShiftVec.isUTerm] + simp only [qqFunc_inj, true_and] + apply nth_ext' k + (by rw [len_termBShiftVec hv.termShiftVec.isUTerm]) (by rw [len_termShiftVec hv.termBShiftVec.isUTerm]) intro i hi - rw [nth_termBShiftVec hv.termShiftVec hi, nth_termShiftVec hv hi, - nth_termShiftVec hv.termBShiftVec hi, nth_termBShiftVec hv hi, ih i hi] + rw [nth_termBShiftVec hv.termShiftVec.isUTerm hi, nth_termShiftVec hv.isUTerm hi, + nth_termShiftVec hv.termBShiftVec.isUTerm hi, nth_termBShiftVec hv.isUTerm hi, ih i hi] end termBShift variable (L) -def Language.qVec (k n w : V) : V := ^#0 ∷ L.termBShiftVec k n w +def Language.qVec (w : V) : V := ^#0 ∷ L.termBShiftVec (len w) w variable {L} -@[simp] lemma len_qVec {k n w : V} (h : L.SemitermVec k n w) : len (L.qVec k n w) = k + 1 := by simp [Language.qVec, h] - -lemma Language.SemitermVec.qVec {k n w : V} (h : L.SemitermVec k n w) : L.SemitermVec (k + 1) (n + 1) (L.qVec k n w) := - ⟨by simp [h], by - intro i hi - rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) - · simp [Language.qVec] - · simpa [Language.qVec, nth_termBShiftVec h (by simpa using hi)] using - h.prop (by simpa using hi) |>.termBShift⟩ - -lemma substs_cons_bShift {n m u t w} (ht : L.Semiterm n t) (hw : L.SemitermVec n m w) : - L.termSubst (n + 1) m (u ∷ w) (L.termBShift n t) = L.termSubst n m w t := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht +@[simp] lemma len_qVec {k w : V} (h : L.IsUTermVec k w) : len (L.qVec w) = k + 1 := by + rcases h.lh; simp [Language.qVec, h, h] + +lemma Language.IsSemitermVec.qVec {k n w : V} (h : L.IsSemitermVec k n w) : L.IsSemitermVec (k + 1) (n + 1) (L.qVec w) := by + rcases h.lh + refine Language.IsSemitermVec.iff.mpr ⟨?_, ?_⟩ + · simp [h.isUTerm, Language.qVec] + · intro i hi + rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) + · simp [Language.qVec] + · simpa [Language.qVec, nth_termBShiftVec h.isUTerm (by simpa using hi)] using + h.nth (by simpa using hi) |>.termBShift + +lemma substs_cons_bShift {u t w} (ht : L.IsSemiterm n t) : + L.termSubst (u ∷ w) (L.termBShift t) = L.termSubst w t := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability · intro z hz; simp [hz] · intro x; simp · intro k f v hf hv ih - simp [hf, hv] - apply nth_ext' k (by simp [hv, hw]) (by simp [hv, hw]) + rw [Language.termBShift_func hf hv.isUTerm, + Language.termSubst_func hf hv.termBShiftVec.isUTerm, + Language.termSubst_func hf hv.isUTerm] + simp [hf, hv.isUTerm] + apply nth_ext' k + (by rw [len_termSubstVec hv.termBShiftVec.isUTerm]) + (by rw [len_termSubstVec hv.isUTerm]) intro i hi - simp [nth_termSubstVec hv.termBShiftVec hi, nth_termSubstVec hv hi, nth_termBShiftVec hv hi, ih i hi] - -lemma termShift_termSubsts {n m w t} (ht : L.Semiterm n t) (hw : L.SemitermVec n m w) : - L.termShift m (L.termSubst n m w t) = L.termSubst n m (L.termShiftVec n m w) (L.termShift n t) := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + rw [nth_termSubstVec hv.termBShiftVec.isUTerm hi, + nth_termSubstVec hv.isUTerm hi, + nth_termBShiftVec hv.isUTerm hi, + ih i hi] + +lemma termShift_termSubsts {n m w t} (ht : L.IsSemiterm n t) (hw : L.IsSemitermVec n m w) : + L.termShift (L.termSubst w t) = L.termSubst (L.termShiftVec n w) (L.termShift t) := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability - · intro z hz; simp [hz, nth_termShiftVec hw hz] + · intro z hz; simp [nth_termShiftVec hw.isUTerm hz] · intro x; simp · intro k f v hf hv ih - simp only [termSubst_func, Language.SemitermVec.termSubstVec, termShift_func, - Language.SemitermVec.termShiftVec, qqFunc_inj, true_and, hf, hv, hw] - apply nth_ext' k (by simp [hw, hv]) (by simp [hv]) + rw [L.termSubst_func hf hv.isUTerm, + L.termShift_func hf (hw.termSubstVec hv).isUTerm, + L.termShift_func hf hv.isUTerm, + L.termSubst_func hf hv.termShiftVec.isUTerm] + simp only [qqFunc_inj, true_and] + apply nth_ext' k + (by rw [len_termShiftVec (hw.termSubstVec hv).isUTerm]) + (by rw [len_termSubstVec hv.termShiftVec.isUTerm]) intro i hi - rw [nth_termShiftVec (hw.termSubstVec hv) hi, - nth_termSubstVec hv hi, - nth_termSubstVec hv.termShiftVec hi, - nth_termShiftVec hv hi, ih i hi] - -lemma bShift_substs {n m w t} (ht : L.Semiterm n t) (hw : L.SemitermVec n m w) : - L.termBShift m (L.termSubst n m w t) = L.termSubst n (m + 1) (L.termBShiftVec n m w) t := by - apply Language.Semiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht + rw [nth_termShiftVec (hw.termSubstVec hv).isUTerm hi, + nth_termSubstVec hv.isUTerm hi, + nth_termSubstVec hv.termShiftVec.isUTerm hi, + nth_termShiftVec hv.isUTerm hi, ih i hi] + +lemma bShift_substs {n m w t} (ht : L.IsSemiterm n t) (hw : L.IsSemitermVec n m w) : + L.termBShift (L.termSubst w t) = L.termSubst (L.termBShiftVec n w) t := by + apply Language.IsSemiterm.induction 𝚺 ?_ ?_ ?_ ?_ t ht · definability - · intro z hz; simp [hz, nth_termBShiftVec hw hz] + · intro z hz; simp [hz, nth_termBShiftVec hw.isUTerm hz] · intro x; simp · intro k f v hf hv ih - simp only [hf, hv, termSubst_func, hw.termSubstVec hv, termBShift_func, qqFunc_inj, true_and] - apply nth_ext' k (by simp [hw, hv]) (by simp [hv]) + rw [L.termSubst_func hf hv.isUTerm, + L.termBShift_func hf (hw.termSubstVec hv).isUTerm, + L.termSubst_func hf hv.isUTerm] + simp only [qqFunc_inj, true_and] + apply nth_ext' k + (by rw [len_termBShiftVec (hw.termSubstVec hv).isUTerm]) + (by rw [len_termSubstVec hv.isUTerm]) intro i hi - simp [nth_termBShiftVec (hw.termSubstVec hv) hi, nth_termSubstVec hv hi, ih i hi] - -lemma substs_qVec_bShift {n t m w} (ht : L.Semiterm n t) (hw : L.SemitermVec n m w) : - L.termSubst (n + 1) (m + 1) (L.qVec n m w) (L.termBShift n t) = L.termBShift m (L.termSubst n m w t) := by - simp [Language.qVec, substs_cons_bShift ht hw.termBShiftVec, bShift_substs ht hw] - -lemma termSubstVec_qVec_qVec {l n m : V} (hv : L.SemitermVec l n v) (hw : L.SemitermVec n m w) : - L.termSubstVec (l + 1) (n + 1) (m + 1) (L.qVec n m w) (L.qVec l n v) = L.qVec l m (L.termSubstVec l n m w v) := by - apply nth_ext' (l + 1) (by rw[len_termSubstVec hv.qVec]) (by simp [hw, hv]) + simp [nth_termBShiftVec (hw.termSubstVec hv).isUTerm hi, nth_termSubstVec hv.isUTerm hi, ih i hi] + +lemma substs_qVec_bShift {n t m w} (ht : L.IsSemiterm n t) (hw : L.IsSemitermVec n m w) : + L.termSubst (L.qVec w) (L.termBShift t) = L.termBShift (L.termSubst w t) := by + rcases hw.lh + simp [Language.qVec] + rw [substs_cons_bShift ht, bShift_substs ht hw] + +lemma termSubstVec_qVec_qVec {l n m : V} (hv : L.IsSemitermVec l n v) (hw : L.IsSemitermVec n m w) : + L.termSubstVec (l + 1) (L.qVec w) (L.qVec v) = L.qVec (L.termSubstVec l w v) := by + apply nth_ext' (len v + 1) + (by rw [len_termSubstVec hv.qVec.isUTerm, hv.lh]) + (by rw [len_qVec (hw.termSubstVec hv).isUTerm, hv.lh]) intro i hi unfold Language.qVec - rw [termSubstVec_cons (by simp) hv.termBShiftVec] + rcases hv.lh; rcases hw.lh + rw [(hw.termSubstVec hv).lh] + rw [termSubstVec_cons (by simp) (by rcases hv.lh; exact hv.termBShiftVec.isUTerm)] rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp · simp - have hi : i < l := by simpa using hi - rw [nth_termSubstVec hv.termBShiftVec hi, - nth_termBShiftVec hv hi, - nth_termBShiftVec (hw.termSubstVec hv) hi, - nth_termSubstVec hv hi, - substs_cons_bShift (hv.2 i hi) hw.termBShiftVec, - bShift_substs (hv.2 i hi) hw] - -lemma termShift_qVec {n m w : V} (hw : L.SemitermVec n m w) : - L.termShiftVec (n + 1) (m + 1) (L.qVec n m w) = L.qVec n m (L.termShiftVec n m w) := by - apply nth_ext' (n + 1) (by rw [len_termShiftVec hw.qVec]) (by simp [hw]) + have hi : i < len v := by simpa using hi + rw [nth_termSubstVec hv.termBShiftVec.isUTerm hi, + nth_termBShiftVec hv.isUTerm hi, + nth_termBShiftVec (hw.termSubstVec hv).isUTerm hi, + nth_termSubstVec hv.isUTerm hi, + substs_cons_bShift (hv.nth hi), + bShift_substs (hv.nth hi) hw] + +lemma termShift_qVec {n m w : V} (hw : L.IsSemitermVec n m w) : + L.termShiftVec (n + 1) (L.qVec w) = L.qVec (L.termShiftVec n w) := by + apply nth_ext' (n + 1) + (by rw [len_termShiftVec hw.qVec.isUTerm]) + (by rw [len_qVec hw.termShiftVec.isUTerm]) intro i hi - rw [nth_termShiftVec hw.qVec hi] + rw [nth_termShiftVec hw.qVec.isUTerm hi] unfold Language.qVec rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp - · rw [nth_cons_succ, nth_cons_succ, - nth_termBShiftVec hw (by simpa using hi), - nth_termBShiftVec hw.termShiftVec (by simpa using hi), - nth_termShiftVec hw (by simpa using hi), - termBShift_termShift (hw.2 i (by simpa using hi))] + · rcases hw.lh + rw [nth_cons_succ, nth_cons_succ, + nth_termBShiftVec hw.isUTerm (by simpa using hi), + nth_termBShiftVec (by simpa [hw.isUTerm] using hw.termShiftVec.isUTerm) (by simpa [hw.isUTerm] using hi), + nth_termShiftVec hw.isUTerm (by simpa using hi), + termBShift_termShift (hw.nth (by simpa using hi))] section fvfree variable (L) -def Language.IsTermFVFree (n t : V) : Prop := L.Semiterm n t ∧ L.termShift n t = t +def Language.IsTermFVFree (n t : V) : Prop := L.IsSemiterm n t ∧ L.termShift t = t variable {L} @[simp] lemma Language.IsTermFVFree.bvar (x : V) : L.IsTermFVFree n ^#x ↔ x < n := by simp [Language.IsTermFVFree] - intro h; simp [h] @[simp] lemma Language.IsTermFVFree.fvar (x : V) : ¬L.IsTermFVFree n ^&x := by simp [Language.IsTermFVFree] @@ -560,27 +595,25 @@ lemma qqAdd_defined : 𝚺₁-Function₂ (qqAdd : V → V → V) via qqAddDef : lemma qqMul_defined : 𝚺₁-Function₂ (qqMul : V → V → V) via qqMulDef := by intro v; simp [qqMulDef, numeral_eq_natCast, qqMul] -instance (Γ m) : Γ-[m + 1]-Function₂ (qqAdd : V → V → V) := .of_sigmaOne qqAdd_defined.to_definable _ _ +instance : Γ-[m + 1]-Function₂ (qqAdd : V → V → V) := .of_sigmaOne qqAdd_defined.to_definable -instance (Γ m) : Γ-[m + 1]-Function₂ (qqMul : V → V → V) := .of_sigmaOne qqMul_defined.to_definable _ _ +instance : Γ-[m + 1]-Function₂ (qqMul : V → V → V) := .of_sigmaOne qqMul_defined.to_definable end lemma qqFunc_absolute (k f v : ℕ) : ((^func k f v : ℕ) : V) = ^func (k : V) (f : V) (v : V) := by simp [qqFunc, nat_cast_pair] -@[simp] lemma zero_semiterm : ⌜ℒₒᵣ⌝.Semiterm n (𝟎 : V) := by +@[simp] lemma zero_semiterm : ⌜ℒₒᵣ⌝.IsSemiterm n (𝟎 : V) := by simp [Formalized.zero, qqFunc_absolute] -@[simp] lemma one_semiterm : ⌜ℒₒᵣ⌝.Semiterm n (𝟏 : V) := by +@[simp] lemma one_semiterm : ⌜ℒₒᵣ⌝.IsSemiterm n (𝟏 : V) := by simp [Formalized.one, qqFunc_absolute] namespace Numeral def blueprint : PR.Blueprint 0 where - zero := .mkSigma “y | y = !!(Semiterm.Operator.numeral ℒₒᵣ Formalized.one)” (by simp) - succ := .mkSigma “y t n | ∃ p, - !mkVec₂Def p t !!(Semiterm.Operator.numeral ℒₒᵣ Formalized.one) ∧ - !qqFuncDef y 2 !!(Semiterm.Operator.numeral ℒₒᵣ addIndex) p” (by simp) + zero := .mkSigma “y | y = ↑Formalized.one” (by simp) + succ := .mkSigma “y t n | ∃ p, !mkVec₂Def p t ↑Formalized.one ∧ !qqFuncDef y 2 ↑addIndex p” (by simp) def construction : PR.Construction V blueprint where zero := fun _ ↦ 𝟏 @@ -608,7 +641,7 @@ instance seqExp_definable : 𝚺-[0 + 1]-Function₁ (numeralAux : V → V) := n end -@[simp] lemma numeralAux_semiterm (n x : V) : ⌜ℒₒᵣ⌝.Semiterm n (numeralAux x) := by +@[simp] lemma numeralAux_semiterm (n x : V) : ⌜ℒₒᵣ⌝.IsSemiterm n (numeralAux x) := by induction x using induction_sigma1 · definability case zero => simp @@ -633,14 +666,16 @@ lemma numeral_succ_pos (pos : 0 < n) : numeral (n + 1 : V) = numeral n ^+ 𝟏 : · simp at pos simp [numeral, ←one_add_one_eq_two, ←add_assoc] -@[simp] lemma numeral_semiterm (n x : V) : ⌜ℒₒᵣ⌝.Semiterm n (numeral x) := by +@[simp] lemma numeral_semiterm (n x : V) : ⌜ℒₒᵣ⌝.IsSemiterm n (numeral x) := by by_cases hx : x = 0 <;> simp [hx, numeral] +@[simp] lemma numeral_uterm (x : V) : ⌜ℒₒᵣ⌝.IsUTerm (numeral x) := (numeral_semiterm 0 x).isUTerm + section def _root_.LO.FirstOrder.Arith.numeralDef : 𝚺₁.Semisentence 2 := .mkSigma “t x | - (x = 0 → t = !!(Semiterm.Operator.numeral ℒₒᵣ Formalized.zero)) ∧ + (x = 0 → t = ↑Formalized.zero) ∧ (x ≠ 0 → ∃ x', !subDef x' x 1 ∧ !numeralAuxDef t x')” (by simp) @@ -651,34 +686,38 @@ lemma numeral_defined : 𝚺₁-Function₁ (numeral : V → V) via numeralDef : @[simp] lemma eval_numeralDef (v) : Semiformula.Evalbm V v numeralDef.val ↔ v 0 = numeral (v 1) := numeral_defined.df.iff v -@[simp] instance numeral_definable : 𝚺₁-Function₁ (numeral : V → V) := numeral_defined.to_definable +instance numeral_definable : 𝚺₁-Function₁ (numeral : V → V) := numeral_defined.to_definable -@[simp] instance numeral_definable' (Γ m) : Γ-[m + 1]-Function₁ (numeral : V → V) := .of_sigmaOne numeral_definable _ _ +instance numeral_definable' : Γ-[m + 1]-Function₁ (numeral : V → V) := .of_sigmaOne numeral_definable end -@[simp] lemma numeral_substs {w : V} (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) (x : V) : - ⌜ℒₒᵣ⌝.termSubst n m w (numeral x) = numeral x := by +@[simp] lemma numeral_substs {w : V} (hw : ⌜ℒₒᵣ⌝.IsSemitermVec n m w) (x : V) : + ⌜ℒₒᵣ⌝.termSubst w (numeral x) = numeral x := by induction x using induction_sigma1 · definability case zero => simp [hw, Formalized.zero, qqFunc_absolute] case succ x ih => rcases zero_or_succ x with (rfl | ⟨x, rfl⟩) · simp [hw, Formalized.one, qqFunc_absolute] - · simp [qqAdd, hw, ih, Formalized.one, qqFunc_absolute] + · simp only [numeral_add_two, qqAdd, LOR_func_addIndex] + rw [Language.termSubst_func (by simp) (by simp [Formalized.one, qqFunc_absolute])] + simp [ih, Formalized.one, qqFunc_absolute] @[simp] lemma numeral_shift (x : V) : - ⌜ℒₒᵣ⌝.termShift n (numeral x) = numeral x := by + ⌜ℒₒᵣ⌝.termShift (numeral x) = numeral x := by induction x using induction_sigma1 · definability case zero => simp [Formalized.zero, qqFunc_absolute] case succ x ih => rcases zero_or_succ x with (rfl | ⟨x, rfl⟩) · simp [Formalized.one, qqFunc_absolute] - · simp [qqAdd, ih, Formalized.one, qqFunc_absolute] + · simp only [numeral_add_two, qqAdd, LOR_func_addIndex] + rw [Language.termShift_func (by simp) (by simp [Formalized.one, qqFunc_absolute])] + simp [ih, Formalized.one, qqFunc_absolute] @[simp] lemma numeral_bShift (x : V) : - ⌜ℒₒᵣ⌝.termBShift n (numeral x) = numeral x := by + ⌜ℒₒᵣ⌝.termBShift (numeral x) = numeral x := by induction x using induction_sigma1 · definability case zero => simp [Formalized.zero, qqFunc_absolute] diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean index fc11a45..383e73f 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean @@ -2,7 +2,7 @@ import Arithmetization.ISigmaOne.Metamath.Term.Functions /-! -# Typed Formalized Semiterm/Term +# Typed Formalized IsSemiterm/Term -/ @@ -42,312 +42,317 @@ section typed_term variable (L) -structure Language.TSemiterm (n : V) where +structure Language.Semiterm (n : V) where val : V - prop : L.Semiterm n val + prop : L.IsSemiterm n val -structure Language.TSemitermVec (m n : V) where +structure Language.SemitermVec (m n : V) where val : V - prop : L.SemitermVec m n val + prop : L.IsSemitermVec m n val -attribute [simp] Language.TSemiterm.prop Language.TSemitermVec.prop +attribute [simp] Language.Semiterm.prop Language.SemitermVec.prop -abbrev Language.TTerm := L.TSemiterm 0 +abbrev Language.Term := L.Semiterm 0 @[ext] -lemma Language.TSemiterm.ext {t u : L.TSemiterm n} +lemma Language.Semiterm.ext {t u : L.Semiterm n} (h : t.val = u.val) : t = u := by rcases t; rcases u; simpa using h -lemma Language.TSemiterm.ext_iff {t u : L.TSemiterm n} : t = u ↔ t.val = u.val := by rcases t; rcases u; simp +lemma Language.Semiterm.ext_iff {t u : L.Semiterm n} : t = u ↔ t.val = u.val := by rcases t; rcases u; simp + +@[simp] lemma Language.Semiterm.isUTerm (t : L.Semiterm n) : L.IsUTerm t.val := t.prop.isUTerm + +@[simp] lemma Language.SemitermVec.isUTerm (v : L.SemitermVec k n) : L.IsUTermVec k v.val := v.prop.isUTerm @[ext] -lemma Language.TSemitermVec.ext {v w : L.TSemitermVec k n} +lemma Language.SemitermVec.ext {v w : L.SemitermVec k n} (h : v.val = w.val) : v = w := by rcases v; rcases w; simpa using h -def Language.bvar {n : V} (z : V) (hz : z < n := by simp) : L.TSemiterm n := ⟨^#z, by simp [hz]⟩ +def Language.bvar {n : V} (z : V) (hz : z < n := by simp) : L.Semiterm n := ⟨^#z, by simp [hz]⟩ -def Language.fvar {n : V} (x : V) : L.TSemiterm n := ⟨^&x, by simp⟩ +def Language.fvar {n : V} (x : V) : L.Semiterm n := ⟨^&x, by simp⟩ -def Language.func {n k f : V} (hf : L.Func k f) (v : L.TSemitermVec k n) : - L.TSemiterm n := ⟨^func k f v.val , by simp [hf]⟩ +def Language.func {n k f : V} (hf : L.Func k f) (v : L.SemitermVec k n) : + L.Semiterm n := ⟨^func k f v.val , by simp [hf]⟩ variable {L} -abbrev bv {n : V} (x : V) (h : x < n := by simp) : L.TSemiterm n := L.bvar x h -abbrev fv {n : V} (x : V) : L.TSemiterm n := L.fvar x +abbrev bv {n : V} (x : V) (h : x < n := by simp) : L.Semiterm n := L.bvar x h +abbrev fv {n : V} (x : V) : L.Semiterm n := L.fvar x scoped prefix:max "#'" => bv scoped prefix:max "&'" => fv @[simp] lemma Language.val_bvar {n : V} (z : V) (hz : z < n) : (L.bvar z hz).val = ^#z := rfl -@[simp] lemma Language.val_fvar {n : V} (x : V) : (L.fvar x : L.TSemiterm n).val = ^&x := rfl +@[simp] lemma Language.val_fvar {n : V} (x : V) : (L.fvar x : L.Semiterm n).val = ^&x := rfl -def Language.TSemiterm.cons {m n} (t : L.TSemiterm n) (v : L.TSemitermVec m n) : - L.TSemitermVec (m + 1) n := ⟨t.val ∷ v.val, v.prop.cons t.prop⟩ +def Language.Semiterm.cons {m n} (t : L.Semiterm n) (v : L.SemitermVec m n) : + L.SemitermVec (m + 1) n := ⟨t.val ∷ v.val, by simp⟩ -scoped infixr:67 " ∷ᵗ " => Language.TSemiterm.cons +scoped infixr:67 " ∷ᵗ " => Language.Semiterm.cons -@[simp] lemma Language.TSemitermvec.val_cons {m n : V} (t : L.TSemiterm n) (v : L.TSemitermVec m n) : - (t ∷ᵗ v).val = t.val ∷ v.val := by simp [Language.TSemiterm.cons] +@[simp] lemma Language.Semitermvec.val_cons {m n : V} (t : L.Semiterm n) (v : L.SemitermVec m n) : + (t ∷ᵗ v).val = t.val ∷ v.val := by simp [Language.Semiterm.cons] variable (L) -def Language.TSemitermVec.nil (n) : L.TSemitermVec 0 n := ⟨0, by simp⟩ +def Language.SemitermVec.nil (n) : L.SemitermVec 0 n := ⟨0, by simp⟩ variable {L} -@[simp] lemma Language.TSemitermvec.val_nil (n : V) : - (Language.TSemitermVec.nil L n).val = 0 := rfl +@[simp] lemma Language.Semitermvec.val_nil (n : V) : + (Language.SemitermVec.nil L n).val = 0 := rfl -abbrev Language.TSemiterm.sing {n} (t : L.TSemiterm n) : L.TSemitermVec (0 + 1) n := t ∷ᵗ .nil L n +abbrev Language.Semiterm.sing {n} (t : L.Semiterm n) : L.SemitermVec (0 + 1) n := t ∷ᵗ .nil L n -namespace Language.TSemiterm +namespace Language.Semiterm -def shift (t : L.TSemiterm n) : L.TSemiterm n := - ⟨L.termShift n t.val, Language.Semiterm.termShift t.prop⟩ +def shift (t : L.Semiterm n) : L.Semiterm n := + ⟨L.termShift t.val, Language.IsSemiterm.termShift t.prop⟩ -def bShift (t : L.TSemiterm n) : L.TSemiterm (n + 1) := - ⟨L.termBShift n t.val, Language.Semiterm.termBShift t.prop⟩ +def bShift (t : L.Semiterm n) : L.Semiterm (n + 1) := + ⟨L.termBShift t.val, Language.IsSemiterm.termBShift t.prop⟩ -def substs (t : L.TSemiterm n) (w : L.TSemitermVec n m) : L.TSemiterm m := - ⟨L.termSubst n m w.val t.val, termSubst_rng_semiterm w.prop t.prop⟩ +def substs (t : L.Semiterm n) (w : L.SemitermVec n m) : L.Semiterm m := + ⟨L.termSubst w.val t.val, w.prop.termSubst t.prop⟩ -@[simp] lemma val_shift (t : L.TSemiterm n) : t.shift.val = L.termShift n t.val := rfl -@[simp] lemma val_bShift (t : L.TSemiterm n) : t.bShift.val = L.termBShift n t.val := rfl -@[simp] lemma val_substs (w : L.TSemitermVec n m) (t : L.TSemiterm n) : (t.substs w).val = L.termSubst n m w.val t.val := rfl +@[simp] lemma val_shift (t : L.Semiterm n) : t.shift.val = L.termShift t.val := rfl +@[simp] lemma val_bShift (t : L.Semiterm n) : t.bShift.val = L.termBShift t.val := rfl +@[simp] lemma val_substs (w : L.SemitermVec n m) (t : L.Semiterm n) : (t.substs w).val = L.termSubst w.val t.val := rfl -end Language.TSemiterm +end Language.Semiterm -notation t:max "^ᵗ/[" w "]" => Language.TSemiterm.substs t w +notation t:max "^ᵗ/[" w "]" => Language.Semiterm.substs t w -namespace Language.TSemitermVec +namespace Language.SemitermVec -def shift (v : L.TSemitermVec k n) : L.TSemitermVec k n := - ⟨L.termShiftVec k n v.val, Language.SemitermVec.termShiftVec v.prop⟩ +def shift (v : L.SemitermVec k n) : L.SemitermVec k n := + ⟨L.termShiftVec k v.val, Language.IsSemitermVec.termShiftVec v.prop⟩ -def bShift (v : L.TSemitermVec k n) : L.TSemitermVec k (n + 1) := - ⟨L.termBShiftVec k n v.val, Language.SemitermVec.termBShiftVec v.prop⟩ +def bShift (v : L.SemitermVec k n) : L.SemitermVec k (n + 1) := + ⟨L.termBShiftVec k v.val, Language.IsSemitermVec.termBShiftVec v.prop⟩ -def substs (v : L.TSemitermVec k n) (w : L.TSemitermVec n m) : L.TSemitermVec k m := - ⟨L.termSubstVec k n m w.val v.val, Language.SemitermVec.termSubstVec w.prop v.prop⟩ +def substs (v : L.SemitermVec k n) (w : L.SemitermVec n m) : L.SemitermVec k m := + ⟨L.termSubstVec k w.val v.val, Language.IsSemitermVec.termSubstVec w.prop v.prop⟩ -@[simp] lemma val_shift (v : L.TSemitermVec k n) : v.shift.val = L.termShiftVec k n v.val := rfl -@[simp] lemma val_bShift (v : L.TSemitermVec k n) : v.bShift.val = L.termBShiftVec k n v.val := rfl -@[simp] lemma val_substs (v : L.TSemitermVec k n) (w : L.TSemitermVec n m) : (v.substs w).val = L.termSubstVec k n m w.val v.val := rfl +@[simp] lemma val_shift (v : L.SemitermVec k n) : v.shift.val = L.termShiftVec k v.val := rfl +@[simp] lemma val_bShift (v : L.SemitermVec k n) : v.bShift.val = L.termBShiftVec k v.val := rfl +@[simp] lemma val_substs (v : L.SemitermVec k n) (w : L.SemitermVec n m) : (v.substs w).val = L.termSubstVec k w.val v.val := rfl @[simp] lemma bShift_nil (n : V) : (nil L n).bShift = nil L (n + 1) := by ext; simp [bShift] -@[simp] lemma bShift_cons (t : L.TSemiterm n) (v : L.TSemitermVec k n) : +@[simp] lemma bShift_cons (t : L.Semiterm n) (v : L.SemitermVec k n) : (t ∷ᵗ v).bShift = t.bShift ∷ᵗ v.bShift := by - ext; simp [bShift, Language.TSemiterm.bShift, termBShiftVec_cons t.prop v.prop] + ext; simp [bShift, Language.Semiterm.bShift, termBShiftVec_cons t.prop.isUTerm v.prop.isUTerm] @[simp] lemma shift_nil (n : V) : (nil L n).shift = nil L n := by ext; simp [shift] -@[simp] lemma shift_cons (t : L.TSemiterm n) (v : L.TSemitermVec k n) : +@[simp] lemma shift_cons (t : L.Semiterm n) (v : L.SemitermVec k n) : (t ∷ᵗ v).shift = t.shift ∷ᵗ v.shift := by - ext; simp [shift, Language.TSemiterm.shift, termShiftVec_cons t.prop v.prop] + ext; simp [shift, Language.Semiterm.shift, termShiftVec_cons t.prop.isUTerm v.prop.isUTerm] -@[simp] lemma substs_nil (w : L.TSemitermVec n m) : +@[simp] lemma substs_nil (w : L.SemitermVec n m) : (nil L n).substs w = nil L m := by ext; simp [substs] -@[simp] lemma substs_cons (w : L.TSemitermVec n m) (t : L.TSemiterm n) (v : L.TSemitermVec k n) : +@[simp] lemma substs_cons (w : L.SemitermVec n m) (t : L.Semiterm n) (v : L.SemitermVec k n) : (t ∷ᵗ v).substs w = t.substs w ∷ᵗ v.substs w := by - ext; simp [substs, Language.TSemiterm.substs, termSubstVec_cons t.prop v.prop] + ext; simp [substs, Language.Semiterm.substs, termSubstVec_cons t.prop.isUTerm v.prop.isUTerm] -def nth (t : L.TSemitermVec k n) (i : V) (hi : i < k := by simp) : L.TSemiterm n := - ⟨t.val.[i], t.prop.prop hi⟩ +def nth (t : L.SemitermVec k n) (i : V) (hi : i < k := by simp) : L.Semiterm n := + ⟨t.val.[i], t.prop.nth hi⟩ -@[simp] lemma nth_val (v : L.TSemitermVec k n) (i : V) (hi : i < k) : (v.nth i hi).val = v.val.[i] := by simp [nth] +@[simp] lemma nth_val (v : L.SemitermVec k n) (i : V) (hi : i < k) : (v.nth i hi).val = v.val.[i] := by simp [nth] -@[simp] lemma nth_zero (t : L.TSemiterm n) (v : L.TSemitermVec k n) : (t ∷ᵗ v).nth 0 = t := by ext; simp [nth] +@[simp] lemma nth_zero (t : L.Semiterm n) (v : L.SemitermVec k n) : (t ∷ᵗ v).nth 0 = t := by ext; simp [nth] -@[simp] lemma nth_succ (t : L.TSemiterm n) (v : L.TSemitermVec k n) (i : V) (hi : i < k) : +@[simp] lemma nth_succ (t : L.Semiterm n) (v : L.SemitermVec k n) (i : V) (hi : i < k) : (t ∷ᵗ v).nth (i + 1) (by simp [hi]) = v.nth i hi := by ext; simp [nth] -@[simp] lemma nth_one (t : L.TSemiterm n) (v : L.TSemitermVec (k + 1) n) : +@[simp] lemma nth_one (t : L.Semiterm n) (v : L.SemitermVec (k + 1) n) : (t ∷ᵗ v).nth 1 (by simp) = v.nth 0 (by simp) := by ext; simp [nth] -lemma nth_of_pos (t : L.TSemiterm n) (v : L.TSemitermVec k n) (i : V) (ipos : 0 < i) (hi : i < k + 1) : +lemma nth_of_pos (t : L.Semiterm n) (v : L.SemitermVec k n) (i : V) (ipos : 0 < i) (hi : i < k + 1) : (t ∷ᵗ v).nth i (by simp [hi]) = v.nth (i - 1) (tsub_lt_iff_left (one_le_of_zero_lt i ipos) |>.mpr hi) := by - ext; simp only [nth, TSemitermvec.val_cons] + ext; simp only [nth, Semitermvec.val_cons] rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp at ipos · simp -def q (w : L.TSemitermVec k n) : L.TSemitermVec (k + 1) (n + 1) := L.bvar (0 : V) ∷ᵗ w.bShift +def q (w : L.SemitermVec k n) : L.SemitermVec (k + 1) (n + 1) := L.bvar (0 : V) ∷ᵗ w.bShift -@[simp] lemma q_zero (w : L.TSemitermVec k n) : w.q.nth 0 = L.bvar 0 := by simp [q] +@[simp] lemma q_zero (w : L.SemitermVec k n) : w.q.nth 0 = L.bvar 0 := by simp [q] -@[simp] lemma q_succ (w : L.TSemitermVec k n) {i} (hi : i < k) : +@[simp] lemma q_succ (w : L.SemitermVec k n) {i} (hi : i < k) : w.q.nth (i + 1) (by simp [hi]) = (w.nth i hi).bShift := by simp only [q, gt_iff_lt, hi, nth_succ] - ext; simp [bShift, nth, Language.TSemiterm.bShift, hi] + ext; simp [bShift, nth, Language.Semiterm.bShift, hi] -@[simp] lemma q_one (w : L.TSemitermVec k n) (h : 0 < k) : w.q.nth 1 (by simp [h]) = (w.nth 0 h).bShift := by +@[simp] lemma q_one (w : L.SemitermVec k n) (h : 0 < k) : w.q.nth 1 (by simp [h]) = (w.nth 0 h).bShift := by simpa using q_succ w h -lemma q_of_pos (w : L.TSemitermVec k n) (i) (ipos : 0 < i) (hi : i < k + 1) : +lemma q_of_pos (w : L.SemitermVec k n) (i) (ipos : 0 < i) (hi : i < k + 1) : w.q.nth i (by simp [hi]) = (w.nth (i - 1) (tsub_lt_iff_left (one_le_of_zero_lt i ipos) |>.mpr hi)).bShift := by rcases zero_or_succ i with (rfl | ⟨i, rfl⟩) · simp at ipos · simp [q_succ w (by simpa using hi)] -@[simp] lemma q_val_eq_qVec (w : L.TSemitermVec k n) : w.q.val = L.qVec k n w.val := by simp [q, Language.qVec, Language.bvar, bShift] +@[simp] lemma q_val_eq_qVec (w : L.SemitermVec k n) : w.q.val = L.qVec w.val := by simp [q, Language.qVec, Language.bvar, bShift, w.prop.lh] -end Language.TSemitermVec +end Language.SemitermVec -namespace Language.TSemiterm +namespace Language.Semiterm @[simp] lemma shift_bvar {z n : V} (hz : z < n) : shift (L.bvar z hz) = L.bvar z hz := by ext; simp [Language.bvar, shift, hz] @[simp] lemma shift_fvar (x : V) : - shift (L.fvar x : L.TSemiterm n) = L.fvar (x + 1) := by ext; simp [Language.fvar, shift] + shift (L.fvar x : L.Semiterm n) = L.fvar (x + 1) := by ext; simp [Language.fvar, shift] -@[simp] lemma shift_func {k f} (hf : L.Func k f) (v : L.TSemitermVec k n) : - shift (L.func hf v) = L.func hf v.shift := by ext; simp [Language.func, shift, TSemitermVec.shift, hf] +@[simp] lemma shift_func {k f} (hf : L.Func k f) (v : L.SemitermVec k n) : + shift (L.func hf v) = L.func hf v.shift := by ext; simp [Language.func, shift, SemitermVec.shift, hf] @[simp] lemma bShift_bvar {z n : V} (hz : z < n) : bShift (L.bvar z hz) = L.bvar (z + 1) (by simpa using hz) := by ext; simp [Language.bvar, bShift, hz] @[simp] lemma bShift_fvar (x : V) : - bShift (L.fvar x : L.TSemiterm n) = L.fvar x := by ext; simp [Language.fvar, bShift] + bShift (L.fvar x : L.Semiterm n) = L.fvar x := by ext; simp [Language.fvar, bShift] -@[simp] lemma bShift_func {k f} (hf : L.Func k f) (v : L.TSemitermVec k n) : - bShift (L.func hf v) = L.func hf v.bShift := by ext; simp [Language.func, bShift, TSemitermVec.bShift, hf] +@[simp] lemma bShift_func {k f} (hf : L.Func k f) (v : L.SemitermVec k n) : + bShift (L.func hf v) = L.func hf v.bShift := by ext; simp [Language.func, bShift, SemitermVec.bShift, hf] -@[simp] lemma substs_bvar {z m : V} (w : L.TSemitermVec n m) (hz : z < n) : - (L.bvar z hz).substs w = w.nth z hz := by ext; simp [Language.bvar, substs, hz, Language.TSemitermVec.nth] +@[simp] lemma substs_bvar {z m : V} (w : L.SemitermVec n m) (hz : z < n) : + (L.bvar z hz).substs w = w.nth z hz := by ext; simp [Language.bvar, substs, hz, Language.SemitermVec.nth] -@[simp] lemma substs_fvar (w : L.TSemitermVec n m) (x : V) : - (L.fvar x : L.TSemiterm n).substs w = L.fvar x := by ext; simp [Language.fvar, substs] +@[simp] lemma substs_fvar (w : L.SemitermVec n m) (x : V) : + (L.fvar x : L.Semiterm n).substs w = L.fvar x := by ext; simp [Language.fvar, substs] -@[simp] lemma substs_func {k f} (w : L.TSemitermVec n m) (hf : L.Func k f) (v : L.TSemitermVec k n) : +@[simp] lemma substs_func {k f} (w : L.SemitermVec n m) (hf : L.Func k f) (v : L.SemitermVec k n) : (L.func hf v).substs w = L.func hf (v.substs w) := by - ext; simp [Language.func, substs, TSemitermVec.substs, hf] + ext; simp [Language.func, substs, SemitermVec.substs, hf] -@[simp] lemma bShift_substs_q (t : L.TSemiterm n) (w : L.TSemitermVec n m) : +@[simp] lemma bShift_substs_q (t : L.Semiterm n) (w : L.SemitermVec n m) : t.bShift.substs w.q = (t.substs w).bShift := by - ext; simp only [substs, TSemitermVec.q_val_eq_qVec, bShift, prop, TSemitermVec.prop, substs_qVec_bShift] + ext; simp only [substs, SemitermVec.q_val_eq_qVec, bShift, prop, SemitermVec.prop, substs_qVec_bShift t.prop w.prop] -@[simp] lemma bShift_substs_sing (t u : L.TTerm) : +@[simp] lemma bShift_substs_sing (t u : L.Term) : t.bShift.substs u.sing = t := by ext; simp [substs, bShift] - rw [show (1 : V) = 0 + 1 by simp, substs_cons_bShift] <;> simp + rw [substs_cons_bShift t.prop]; simp -lemma bShift_shift_comm (t : L.TSemiterm n) : +lemma bShift_shift_comm (t : L.Semiterm n) : t.shift.bShift = t.bShift.shift := by - ext; simp [termBShift_termShift] + ext; simp [termBShift_termShift t.prop] -end Language.TSemiterm +end Language.Semiterm end typed_term section typed_isfvfree -namespace Language.TSemiterm +namespace Language.Semiterm -def FVFree (t : L.TSemiterm n) : Prop := L.IsTermFVFree n t.val +def FVFree (t : L.Semiterm n) : Prop := L.IsTermFVFree n t.val -lemma FVFree.iff {t : L.TSemiterm n} : t.FVFree ↔ t.shift = t := by +lemma FVFree.iff {t : L.Semiterm n} : t.FVFree ↔ t.shift = t := by simp [FVFree, Language.IsTermFVFree, ext_iff] @[simp] lemma FVFree.bvar (z : V) (h : z < n) : (L.bvar z h).FVFree := by simp [FVFree, h] -@[simp] lemma FVFree.bShift (t : L.TSemiterm n) (ht : t.FVFree) : +@[simp] lemma FVFree.bShift (t : L.Semiterm n) (ht : t.FVFree) : t.bShift.FVFree := by simp [FVFree.iff, ←bShift_shift_comm, FVFree.iff.mp ht] -end Language.TSemiterm +end Language.Semiterm end typed_isfvfree namespace Formalized -def typedNumeral (n m : V) : ⌜ℒₒᵣ⌝.TSemiterm n := ⟨numeral m, by simp⟩ +def typedNumeral (n m : V) : ⌜ℒₒᵣ⌝.Semiterm n := ⟨numeral m, by simp⟩ -def add {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiterm n := ⟨t.val ^+ u.val, by simp [qqAdd]⟩ +def add {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiterm n := ⟨t.val ^+ u.val, by simp [qqAdd]⟩ -def mul {n : V} (t u : ⌜ℒₒᵣ⌝.TSemiterm n) : ⌜ℒₒᵣ⌝.TSemiterm n := ⟨t.val ^* u.val, by simp [qqMul]⟩ +def mul {n : V} (t u : ⌜ℒₒᵣ⌝.Semiterm n) : ⌜ℒₒᵣ⌝.Semiterm n := ⟨t.val ^* u.val, by simp [qqMul]⟩ -instance (n : V) : Add (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨add⟩ +instance (n : V) : Add (⌜ℒₒᵣ⌝.Semiterm n) := ⟨add⟩ -instance (n : V) : Mul (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨mul⟩ +instance (n : V) : Mul (⌜ℒₒᵣ⌝.Semiterm n) := ⟨mul⟩ -instance coeNumeral (n : V) : Coe V (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨typedNumeral n⟩ +instance coeNumeral (n : V) : Coe V (⌜ℒₒᵣ⌝.Semiterm n) := ⟨typedNumeral n⟩ variable {n : V} -@[simp] lemma val_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.TSemiterm n).val = numeral x := rfl +@[simp] lemma val_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.Semiterm n).val = numeral x := rfl -@[simp] lemma val_add (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ + t₂).val = t₁.val ^+ t₂.val := rfl +@[simp] lemma val_add (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ + t₂).val = t₁.val ^+ t₂.val := rfl -@[simp] lemma val_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ * t₂).val = t₁.val ^* t₂.val := rfl +@[simp] lemma val_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ * t₂).val = t₁.val ^* t₂.val := rfl -@[simp] lemma add_inj_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma add_inj_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : t₁ + t₂ = u₁ + u₂ ↔ t₁ = u₁ ∧ t₂ = u₂ := by - simp [Language.TSemiterm.ext_iff, qqAdd] + simp [Language.Semiterm.ext_iff, qqAdd] -@[simp] lemma mul_inj_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TSemiterm n} : +@[simp] lemma mul_inj_iff {t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Semiterm n} : t₁ * t₂ = u₁ * u₂ ↔ t₁ = u₁ ∧ t₂ = u₂ := by - simp [Language.TSemiterm.ext_iff, qqMul] + simp [Language.Semiterm.ext_iff, qqMul] -@[simp] lemma subst_numeral {m n : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (x : V) : - (↑x : ⌜ℒₒᵣ⌝.TSemiterm n).substs w = ↑x := by - ext; simp [Language.TSemiterm.substs] +@[simp] lemma subst_numeral {m n : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (x : V) : + (↑x : ⌜ℒₒᵣ⌝.Semiterm n).substs w = ↑x := by + ext; simp [Language.Semiterm.substs, numeral_substs w.prop] -@[simp] lemma subst_add {m n : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma subst_add {m n : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ + t₂).substs w = t₁.substs w + t₂.substs w := by - ext; simp [qqAdd, Language.TSemiterm.substs] + ext; simp [qqAdd, Language.Semiterm.substs] -@[simp] lemma subst_mul {m n : V} (w : ⌜ℒₒᵣ⌝.TSemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : +@[simp] lemma subst_mul {m n : V} (w : ⌜ℒₒᵣ⌝.SemitermVec n m) (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ * t₂).substs w = t₁.substs w * t₂.substs w := by - ext; simp [qqMul, Language.TSemiterm.substs] + ext; simp [qqMul, Language.Semiterm.substs] -@[simp] lemma shift_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.TSemiterm n).shift = ↑x := by - ext; simp [Language.TSemiterm.shift] +@[simp] lemma shift_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.Semiterm n).shift = ↑x := by + ext; simp [Language.Semiterm.shift] -@[simp] lemma shift_add (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ + t₂).shift = t₁.shift + t₂.shift := by - ext; simp [qqAdd, Language.TSemiterm.shift] +@[simp] lemma shift_add (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ + t₂).shift = t₁.shift + t₂.shift := by + ext; simp [qqAdd, Language.Semiterm.shift] -@[simp] lemma shift_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ * t₂).shift = t₁.shift * t₂.shift := by - ext; simp [qqMul, Language.TSemiterm.shift] +@[simp] lemma shift_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ * t₂).shift = t₁.shift * t₂.shift := by + ext; simp [qqMul, Language.Semiterm.shift] -@[simp] lemma bShift_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.TSemiterm n).bShift = ↑x := by - ext; simp [Language.TSemiterm.bShift] +@[simp] lemma bShift_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.Semiterm n).bShift = ↑x := by + ext; simp [Language.Semiterm.bShift] -@[simp] lemma bShift_add (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ + t₂).bShift = t₁.bShift + t₂.bShift := by - ext; simp [qqAdd, Language.TSemiterm.bShift] +@[simp] lemma bShift_add (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ + t₂).bShift = t₁.bShift + t₂.bShift := by + ext; simp [qqAdd, Language.Semiterm.bShift] -@[simp] lemma bShift_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : (t₁ * t₂).bShift = t₁.bShift * t₂.bShift := by - ext; simp [qqMul, Language.TSemiterm.bShift] +@[simp] lemma bShift_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : (t₁ * t₂).bShift = t₁.bShift * t₂.bShift := by + ext; simp [qqMul, Language.Semiterm.bShift] -@[simp] lemma fvFree_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.TSemiterm n).FVFree := by simp [Language.TSemiterm.FVFree.iff] +@[simp] lemma fvFree_numeral (x : V) : (↑x : ⌜ℒₒᵣ⌝.Semiterm n).FVFree := by simp [Language.Semiterm.FVFree.iff] -@[simp] lemma fvFree_add (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : - (t₁ + t₂).FVFree ↔ t₁.FVFree ∧ t₂.FVFree := by simp [Language.TSemiterm.FVFree.iff] +@[simp] lemma fvFree_add (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : + (t₁ + t₂).FVFree ↔ t₁.FVFree ∧ t₂.FVFree := by simp [Language.Semiterm.FVFree.iff] -@[simp] lemma fvFree_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.TSemiterm n) : - (t₁ * t₂).FVFree ↔ t₁.FVFree ∧ t₂.FVFree := by simp [Language.TSemiterm.FVFree.iff] +@[simp] lemma fvFree_mul (t₁ t₂ : ⌜ℒₒᵣ⌝.Semiterm n) : + (t₁ * t₂).FVFree ↔ t₁.FVFree ∧ t₂.FVFree := by simp [Language.Semiterm.FVFree.iff] +/- lemma replace {P : α → Prop} {x y} (hx : P x) (h : x = y) : P y := h ▸ hx -lemma semiterm_induction (Γ) {n : V} {P : ⌜ℒₒᵣ⌝.TSemiterm n → Prop} - (hP : Γ-[1]-Predicate (fun x ↦ (h : ⌜ℒₒᵣ⌝.Semiterm n x) → P ⟨x, h⟩)) +lemma semiterm_induction (Γ) {n : V} {P : ⌜ℒₒᵣ⌝.Semiterm n → Prop} + (hP : Γ-[1]-Predicate (fun x ↦ (h : ⌜ℒₒᵣ⌝.IsSemiterm n x) → P ⟨x, h⟩)) (hBvar : ∀ (z : V) (h : z < n), P (⌜ℒₒᵣ⌝.bvar z h)) (hFvar : ∀ x, P (⌜ℒₒᵣ⌝.fvar x)) - (hZero : P ((0 : V) : ⌜ℒₒᵣ⌝.TSemiterm n)) - (hOne : P ((1 : V) : ⌜ℒₒᵣ⌝.TSemiterm n)) + (hZero : P ((0 : V) : ⌜ℒₒᵣ⌝.Semiterm n)) + (hOne : P ((1 : V) : ⌜ℒₒᵣ⌝.Semiterm n)) (hAdd : ∀ t₁ t₂, P t₁ → P t₂ → P (t₁ + t₂)) (hMul : ∀ t₁ t₂, P t₁ → P t₂ → P (t₁ * t₂)) : - ∀ (t : ⌜ℒₒᵣ⌝[V].TSemiterm n), P t := by - let Q := fun x ↦ (h : ⌜ℒₒᵣ⌝.Semiterm n x) → P ⟨x, h⟩ - suffices ∀ t, ⌜ℒₒᵣ⌝.Semiterm n t → Q t by intro t; exact this t.val t.prop t.prop - apply Language.Semiterm.induction Γ hP + ∀ (t : ⌜ℒₒᵣ⌝[V].Semiterm n), P t := by + let Q := fun x ↦ (h : ⌜ℒₒᵣ⌝.IsSemiterm n x) → P ⟨x, h⟩ + suffices ∀ t, ⌜ℒₒᵣ⌝.IsSemiterm n t → Q t by intro t; exact this t.val t.prop t.prop + apply Language.IsSemiterm.induction Γ hP case hbvar => intro z hz _; exact hBvar z hz case hfvar => intro x _; exact hFvar x case hfunc => @@ -357,14 +362,15 @@ lemma semiterm_induction (Γ) {n : V} {P : ⌜ℒₒᵣ⌝.TSemiterm n → Prop} exact replace hZero (by ext; simp [Formalized.zero, qqFunc_absolute]) · rcases (by simpa using hv) exact replace hOne (by ext; simp [Formalized.one, qqFunc_absolute]) - · rcases Language.SemitermVec.two_iff.mp hv with ⟨t₁, t₂, ht₁, ht₂, rfl⟩ + · rcases Language.IsSemitermVec.two_iff.mp hv with ⟨t₁, t₂, ht₁, ht₂, rfl⟩ exact hAdd ⟨t₁, ht₁⟩ ⟨t₂, ht₂⟩ (by simpa using ih 0 (by simp) (by simp [ht₁])) (by simpa using ih 1 (by simp) (by simp [ht₂])) - · rcases Language.SemitermVec.two_iff.mp hv with ⟨t₁, t₂, ht₁, ht₂, rfl⟩ + · rcases Language.IsSemitermVec.two_iff.mp hv with ⟨t₁, t₂, ht₁, ht₂, rfl⟩ exact hMul ⟨t₁, ht₁⟩ ⟨t₂, ht₂⟩ (by simpa using ih 0 (by simp) (by simp [ht₁])) (by simpa using ih 1 (by simp) (by simp [ht₂])) +-/ end Formalized diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean index bca9f29..9ee82d4 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean @@ -35,15 +35,15 @@ variable {V} class EQTheory (T : LOR.TTheory (V := V)) where refl : T ⊢ (#'0 =' #'0).all - replace (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) : T ⊢ (#'1 =' #'0 ⟶ p^/[(#'1).sing] ⟶ p^/[(#'0).sing]).all.all + replace (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) : T ⊢ (#'1 =' #'0 ⟶ p^/[(#'1).sing] ⟶ p^/[(#'0).sing]).all.all -abbrev oneAbbrev {n} : ⌜ℒₒᵣ⌝[V].TSemiterm n := (1 : V) +abbrev oneAbbrev {n} : ⌜ℒₒᵣ⌝[V].Semiterm n := (1 : V) scoped notation "^1" => oneAbbrev class R₀Theory (T : LOR.TTheory (V := V)) extends EQTheory T where - add (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) - mul (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) + add (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n + m) + mul (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n * m) ne {n m : V} : n ≠ m → T ⊢ ↑n ≠' ↑m ltNumeral (n : V) : T ⊢ (#'0 <' ↑n ⟷ (tSubstItr (#'0).sing (#'1 =' #'0) n).disj).all /- @@ -65,22 +65,22 @@ section EQTheory variable [EQTheory T] -def eqRefl (t : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t =' t := by +def eqRefl (t : ⌜ℒₒᵣ⌝.Term) : T ⊢ t =' t := by have : T ⊢ (#'0 =' #'0).all := EQTheory.refl - simpa [Language.TSemiformula.substs₁] using specialize this t + simpa [Language.Semiformula.substs₁] using specialize this t -lemma eq_refl! (t : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t =' t := ⟨eqRefl T t⟩ +lemma eq_refl! (t : ⌜ℒₒᵣ⌝.Term) : T ⊢! t =' t := ⟨eqRefl T t⟩ -noncomputable def replace (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : +noncomputable def replace (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢ t =' u ⟶ p^/[t.sing] ⟶ p^/[u.sing] := by have : T ⊢ (#'1 =' #'0 ⟶ p^/[(#'1).sing] ⟶ p^/[(#'0).sing]).all.all := EQTheory.replace p have := by simpa using specialize this t - simpa [Language.TSemitermVec.q_of_pos, Language.TSemiformula.substs₁, + simpa [Language.SemitermVec.q_of_pos, Language.Semiformula.substs₁, Language.TSemifromula.substs_substs] using specialize this u -lemma replace! (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t =' u ⟶ p^/[t.sing] ⟶ p^/[u.sing] := ⟨replace T p t u⟩ +lemma replace! (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢! t =' u ⟶ p^/[t.sing] ⟶ p^/[u.sing] := ⟨replace T p t u⟩ -def eqSymm (t₁ t₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ t₂ =' t₁ := by +def eqSymm (t₁ t₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ t₂ =' t₁ := by apply deduct' let Γ := [t₁ =' t₂] have e₁ : Γ ⊢[T] t₁ =' t₂ := FiniteContext.byAxm (by simp [Γ]) @@ -89,11 +89,11 @@ def eqSymm (t₁ t₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ t₂ =' simpa using replace T (#'0 =' t₁.bShift) t₁ t₂ exact this ⨀ e₁ ⨀ e₂ -lemma eq_symm! (t₁ t₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ t₂ =' t₁ := ⟨eqSymm T t₁ t₂⟩ +lemma eq_symm! (t₁ t₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ t₂ =' t₁ := ⟨eqSymm T t₁ t₂⟩ -lemma eq_symm'! {t₁ t₂ : ⌜ℒₒᵣ⌝.TTerm} (h : T ⊢! t₁ =' t₂) : T ⊢! t₂ =' t₁ := eq_symm! T t₁ t₂ ⨀ h +lemma eq_symm'! {t₁ t₂ : ⌜ℒₒᵣ⌝.Term} (h : T ⊢! t₁ =' t₂) : T ⊢! t₂ =' t₁ := eq_symm! T t₁ t₂ ⨀ h -def eqTrans (t₁ t₂ t₃ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ t₂ =' t₃ ⟶ t₁ =' t₃ := by +def eqTrans (t₁ t₂ t₃ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ t₂ =' t₃ ⟶ t₁ =' t₃ := by apply deduct' apply deduct let Γ := [t₂ =' t₃, t₁ =' t₂] @@ -103,9 +103,9 @@ def eqTrans (t₁ t₂ t₃ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ t simpa using replace T (t₁.bShift =' #'0) t₂ t₃ exact this ⨀ e₂ ⨀ e₁ -lemma eq_trans! (t₁ t₂ t₃ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ t₂ =' t₃ ⟶ t₁ =' t₃ := ⟨eqTrans T t₁ t₂ t₃⟩ +lemma eq_trans! (t₁ t₂ t₃ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ t₂ =' t₃ ⟶ t₁ =' t₃ := ⟨eqTrans T t₁ t₂ t₃⟩ -noncomputable def addExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ + u₁) =' (t₂ + u₂) := by +noncomputable def addExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ + u₁) =' (t₂ + u₂) := by apply deduct' apply deduct let Γ := [u₁ =' u₂, t₁ =' t₂] @@ -120,9 +120,9 @@ noncomputable def addExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t simpa using this exact of (Γ := Γ) this ⨀ bu ⨀ b -lemma add_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ + u₁) =' (t₂ + u₂) := ⟨addExt T t₁ t₂ u₁ u₂⟩ +lemma add_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ + u₁) =' (t₂ + u₂) := ⟨addExt T t₁ t₂ u₁ u₂⟩ -noncomputable def mulExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ * u₁) =' (t₂ * u₂) := by +noncomputable def mulExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ * u₁) =' (t₂ * u₂) := by apply deduct' apply deduct let Γ := [u₁ =' u₂, t₁ =' t₂] @@ -137,9 +137,9 @@ noncomputable def mulExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t simpa using this exact of (Γ := Γ) this ⨀ bu ⨀ b -lemma mul_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ * u₁) =' (t₂ * u₂) := ⟨mulExt T t₁ t₂ u₁ u₂⟩ +lemma mul_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ (t₁ * u₁) =' (t₂ * u₂) := ⟨mulExt T t₁ t₂ u₁ u₂⟩ -noncomputable def eqExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ =' u₁ ⟶ t₂ =' u₂ := by +noncomputable def eqExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ =' u₁ ⟶ t₂ =' u₂ := by apply deduct' apply deduct apply deduct @@ -150,10 +150,10 @@ noncomputable def eqExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t have e3 : Γ ⊢[T] u₁ =' u₂ := FiniteContext.byAxm (by simp [Γ]) exact (of <| eqTrans T t₂ u₁ u₂) ⨀ ((of <| eqTrans T t₂ t₁ u₁) ⨀ e1 ⨀ e2) ⨀ e3 -lemma eq_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ =' u₁ ⟶ t₂ =' u₂ := +lemma eq_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ =' u₁ ⟶ t₂ =' u₂ := ⟨eqExt T t₁ t₂ u₁ u₂⟩ -noncomputable def neExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≠' u₁ ⟶ t₂ ≠' u₂ := by +noncomputable def neExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≠' u₁ ⟶ t₂ ≠' u₂ := by apply deduct' apply deduct apply deduct @@ -169,10 +169,10 @@ noncomputable def neExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t simpa using replace T (t₂.bShift ≠' #'0) u₁ u₂ exact of (Γ := Γ) this ⨀ bu ⨀ b -lemma ne_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≠' u₁ ⟶ t₂ ≠' u₂ := +lemma ne_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≠' u₁ ⟶ t₂ ≠' u₂ := ⟨neExt T t₁ t₂ u₁ u₂⟩ -noncomputable def ltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ <' u₁ ⟶ t₂ <' u₂ := by +noncomputable def ltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ <' u₁ ⟶ t₂ <' u₂ := by apply deduct' apply deduct apply deduct @@ -189,9 +189,9 @@ noncomputable def ltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t simpa using this exact of (Γ := Γ) this ⨀ bu ⨀ b -lemma lt_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ <' u₁ ⟶ t₂ <' u₂ := ⟨ltExt T t₁ t₂ u₁ u₂⟩ +lemma lt_ext! (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ <' u₁ ⟶ t₂ <' u₂ := ⟨ltExt T t₁ t₂ u₁ u₂⟩ -noncomputable def nltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≮' u₁ ⟶ t₂ ≮' u₂ := by +noncomputable def nltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢ t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≮' u₁ ⟶ t₂ ≮' u₂ := by apply deduct' apply deduct apply deduct @@ -208,20 +208,20 @@ noncomputable def nltExt (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢ t simpa using this exact of (Γ := Γ) this ⨀ bu ⨀ b -lemma nlt_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≮' u₁ ⟶ t₂ ≮' u₂ := ⟨nltExt T t₁ t₂ u₁ u₂⟩ +lemma nlt_ext (t₁ t₂ u₁ u₂ : ⌜ℒₒᵣ⌝.Term) : T ⊢! t₁ =' t₂ ⟶ u₁ =' u₂ ⟶ t₁ ≮' u₁ ⟶ t₂ ≮' u₂ := ⟨nltExt T t₁ t₂ u₁ u₂⟩ -noncomputable def ballReplace (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : +noncomputable def ballReplace (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢ t =' u ⟶ p.ball t ⟶ p.ball u := by simpa [Language.TSemifromula.substs_substs] using replace T ((p^/[(#'0).sing]).ball #'0) t u -lemma ball_replace! (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : +lemma ball_replace! (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢! t =' u ⟶ p.ball t ⟶ p.ball u := ⟨ballReplace T p t u⟩ -noncomputable def bexReplace (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : +noncomputable def bexReplace (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢ t =' u ⟶ p.bex t ⟶ p.bex u := by simpa [Language.TSemifromula.substs_substs] using replace T ((p^/[(#'0).sing]).bex #'0) t u -lemma bex_replace! (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.TTerm) : +lemma bex_replace! (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (t u : ⌜ℒₒᵣ⌝.Term) : T ⊢! t =' u ⟶ p.bex t ⟶ p.bex u := ⟨bexReplace T p t u⟩ def eqComplete {n m : V} (h : n = m) : T ⊢ ↑n =' ↑m := by @@ -236,20 +236,20 @@ end EQTheory section Q₀ class Q₀Theory (T : LOR.TTheory (V := V)) extends EQTheory T where - add_zero : T ⊢ ((#'0 + ((0 : V) : ⌜ℒₒᵣ⌝[V].TSemiterm (0 + 1)) : ⌜ℒₒᵣ⌝[V].TSemiterm (0 + 1)) =' #'0).all - add_succ : T ⊢ ((#'1 + (#'0 + ((1 : V) : ⌜ℒₒᵣ⌝[V].TSemiterm 2))) =' ((#'1 + #'0) + ((1 : V) : ⌜ℒₒᵣ⌝[V].TSemiterm 2))).all.all - mul : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) + add_zero : T ⊢ ((#'0 + ((0 : V) : ⌜ℒₒᵣ⌝[V].Semiterm (0 + 1)) : ⌜ℒₒᵣ⌝[V].Semiterm (0 + 1)) =' #'0).all + add_succ : T ⊢ ((#'1 + (#'0 + ((1 : V) : ⌜ℒₒᵣ⌝[V].Semiterm 2))) =' ((#'1 + #'0) + ((1 : V) : ⌜ℒₒᵣ⌝[V].Semiterm 2))).all.all + mul : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n * m) ne {n m : V} : n ≠ m → T ⊢ ↑n ≠' ↑m - ltNumeral (n : V) : T ⊢ (#'1 <' (#'0 + ((1 : V) : ⌜ℒₒᵣ⌝[V].TSemiterm 2)) ⟷ #'1 =' #'0 ⋎ #'1 <' #'0).all.all + ltNumeral (n : V) : T ⊢ (#'1 <' (#'0 + ((1 : V) : ⌜ℒₒᵣ⌝[V].Semiterm 2)) ⟷ #'1 =' #'0 ⋎ #'1 <' #'0).all.all variable [Q₀Theory T] -lemma add_zero (t : ⌜ℒₒᵣ⌝.TTerm) : T ⊢! (t + ((0 : V) : ⌜ℒₒᵣ⌝[V].TSemiterm 0) : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' t := ⟨by +lemma add_zero (t : ⌜ℒₒᵣ⌝.Term) : T ⊢! (t + ((0 : V) : ⌜ℒₒᵣ⌝[V].Semiterm 0) : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' t := ⟨by simpa using specialize (Q₀Theory.add_zero (T := T)) t⟩ -lemma addssss (n m : V) : T ⊢! (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) := by { +lemma addssss (n m : V) : T ⊢! (n + m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n + m) := by { revert n m - let d : ∀ n m : V, m ≤ n → T ⊢! ((n - m : V) + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' n := by { + let d : ∀ n m : V, m ≤ n → T ⊢! ((n - m : V) + m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' n := by { intro n m hmn induction m using induction_sigma1 · sorry -- simp [Language.Theory.TProvable.iff_provable]; definability @@ -271,23 +271,23 @@ section R₀ variable [R₀Theory T] -def addComplete (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) := R₀Theory.add n m +def addComplete (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n + m) := R₀Theory.add n m -lemma add_complete! (n m : V) : T ⊢! (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) := ⟨addComplete T n m⟩ +lemma add_complete! (n m : V) : T ⊢! (n + m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n + m) := ⟨addComplete T n m⟩ -def mulComplete (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) := R₀Theory.mul n m +def mulComplete (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n * m) := R₀Theory.mul n m -lemma mul_complete! (n m : V) : T ⊢! (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) := ⟨mulComplete T n m⟩ +lemma mul_complete! (n m : V) : T ⊢! (n * m : ⌜ℒₒᵣ⌝[V].Semiterm 0) =' ↑(n * m) := ⟨mulComplete T n m⟩ def neComplete {n m : V} (h : n ≠ m) : T ⊢ ↑n ≠' ↑m := R₀Theory.ne h lemma ne_complete! {n m : V} (h : n ≠ m) : T ⊢! ↑n ≠' ↑m := ⟨neComplete T h⟩ -def ltNumeral (t : ⌜ℒₒᵣ⌝.TTerm) (n : V) : T ⊢ t <' ↑n ⟷ (tSubstItr t.sing (#'1 =' #'0) n).disj := by +def ltNumeral (t : ⌜ℒₒᵣ⌝.Term) (n : V) : T ⊢ t <' ↑n ⟷ (tSubstItr t.sing (#'1 =' #'0) n).disj := by have : T ⊢ (#'0 <' ↑n ⟷ (tSubstItr (#'0).sing (#'1 =' #'0) n).disj).all := R₀Theory.ltNumeral n - simpa [Language.TSemitermVec.q_of_pos, Language.TSemiformula.substs₁] using specialize this t + simpa [Language.SemitermVec.q_of_pos, Language.Semiformula.substs₁] using specialize this t -noncomputable def nltNumeral (t : ⌜ℒₒᵣ⌝.TTerm) (n : V) : T ⊢ t ≮' ↑n ⟷ (tSubstItr t.sing (#'1 ≠' #'0) n).conj := by +noncomputable def nltNumeral (t : ⌜ℒₒᵣ⌝.Term) (n : V) : T ⊢ t ≮' ↑n ⟷ (tSubstItr t.sing (#'1 ≠' #'0) n).conj := by simpa using negReplaceIff' <| ltNumeral T t n def ltComplete {n m : V} (h : n < m) : T ⊢ ↑n <' ↑m := by @@ -299,7 +299,7 @@ def ltComplete {n m : V} (h : n < m) : T ⊢ ↑n <' ↑m := by lemma lt_complete! {n m : V} (h : n < m) : T ⊢! ↑n <' ↑m := ⟨ltComplete T h⟩ noncomputable def nltComplete {n m : V} (h : m ≤ n) : T ⊢ ↑n ≮' ↑m := by - have : T ⊢ ↑n ≮' ↑m ⟷ (tSubstItr (↑n : ⌜ℒₒᵣ⌝.TTerm).sing (#'1 ≠' #'0) m).conj := by + have : T ⊢ ↑n ≮' ↑m ⟷ (tSubstItr (↑n : ⌜ℒₒᵣ⌝.Term).sing (#'1 ≠' #'0) m).conj := by simpa using negReplaceIff' <| ltNumeral T n m refine andRight this ⨀ ?_ apply conj' @@ -310,38 +310,38 @@ noncomputable def nltComplete {n m : V} (h : m ≤ n) : T ⊢ ↑n ≮' ↑m := lemma nlt_complete {n m : V} (h : m ≤ n) : T ⊢! ↑n ≮' ↑m := ⟨nltComplete T h⟩ -noncomputable def ballIntro (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (n : V) - (bs : ∀ i < n, T ⊢ p ^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing]) : +noncomputable def ballIntro (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (n : V) + (bs : ∀ i < n, T ⊢ p ^/[(i : ⌜ℒₒᵣ⌝.Term).sing]) : T ⊢ p.ball ↑n := by apply all suffices T ⊢ &'0 ≮' ↑n ⋎ p.shift^/[(&'0).sing] by - simpa [Language.TSemiformula.free, Language.TSemiformula.substs₁] + simpa [Language.Semiformula.free, Language.Semiformula.substs₁] have : T ⊢ (tSubstItr (&'0).sing (#'1 ≠' #'0) n).conj ⋎ p.shift^/[(&'0).sing] := by apply conjOr' intro i hi have hi : i < n := by simpa using hi let Γ := [&'0 =' typedNumeral 0 i] suffices Γ ⊢[T] p.shift^/[(&'0).sing] by - simpa [nth_tSubstItr', hi, Language.TSemiformula.imp_def] using deduct' this + simpa [nth_tSubstItr', hi, Language.Semiformula.imp_def] using deduct' this have e : Γ ⊢[T] ↑i =' &'0 := of (eqSymm T &'0 ↑i) ⨀ (FiniteContext.byAxm <| by simp [Γ]) - have : T ⊢ p.shift^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing] := by + have : T ⊢ p.shift^/[(i : ⌜ℒₒᵣ⌝.Term).sing] := by simpa [Language.TSemifromula.shift_substs] using shift (bs i hi) exact of (replace T p.shift ↑i &'0) ⨀ e ⨀ of this exact orReplaceLeft' this (andRight (nltNumeral T (&'0) n)) -lemma ball_intro! (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (n : V) - (bs : ∀ i < n, T ⊢! p ^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing]) : +lemma ball_intro! (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (n : V) + (bs : ∀ i < n, T ⊢! p ^/[(i : ⌜ℒₒᵣ⌝.Term).sing]) : T ⊢! p.ball ↑n := ⟨ballIntro T p n fun i hi ↦ (bs i hi).get⟩ -noncomputable def bexIntro (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (n : V) {i} - (hi : i < n) (b : T ⊢ p ^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing]) : +noncomputable def bexIntro (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (n : V) {i} + (hi : i < n) (b : T ⊢ p ^/[(i : ⌜ℒₒᵣ⌝.Term).sing]) : T ⊢ p.bex ↑n := by apply ex i - suffices T ⊢ i <' n ⋏ p^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing] by simpa + suffices T ⊢ i <' n ⋏ p^/[(i : ⌜ℒₒᵣ⌝.Term).sing] by simpa exact System.andIntro (ltComplete T hi) b -lemma bex_intro! (p : ⌜ℒₒᵣ⌝.TSemiformula (0 + 1)) (n : V) {i} - (hi : i < n) (b : T ⊢! p ^/[(i : ⌜ℒₒᵣ⌝.TTerm).sing]) : +lemma bex_intro! (p : ⌜ℒₒᵣ⌝.Semiformula (0 + 1)) (n : V) {i} + (hi : i < n) (b : T ⊢! p ^/[(i : ⌜ℒₒᵣ⌝.Term).sing]) : T ⊢! p.bex ↑n := ⟨bexIntro T p n hi b.get⟩ end R₀ diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean index 30d1277..7d46d79 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean @@ -17,7 +17,7 @@ def _root_.LO.FirstOrder.SyntacticTheory.tDef (T : SyntacticTheory L) [d : T.Δ abbrev _root_.LO.FirstOrder.Theory.tDef (T : Theory L) [d : T.Δ₁Definable] : Arith.LDef.TDef L.lDef := d.toTDef -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {T : SyntacticTheory L} [T.Δ₁Definable] @@ -34,7 +34,7 @@ lemma Language.SyntacticTheory.codeIn_iff : x ∈ T.codeIn V ↔ V ⊧/![x] T.tD lemma mem_coded_theory {σ} (h : σ ∈ T) : ⌜σ⌝ ∈ T.codeIn V := Language.SyntacticTheory.codeIn_iff.mpr <| by have := consequence_iff_add_eq.mp (sound! <| SyntacticTheory.Δ₁Definable.mem_iff.mp h) V inferInstance - simpa [models_iff, Semiformula.syntacticformula_goedelNumber_def, numeral_eq_natCast] using this + simpa [models_iff, coe_quote, Semiformula.syntacticformula_goedelNumber_def, numeral_eq_natCast] using this instance tDef_defined : (T.codeIn V).Defined T.tDef where defined := ⟨by diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean b/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean index e91f32e..92d90eb 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean @@ -99,7 +99,7 @@ def Language.Theory.singletonDef (p : FirstOrder.SyntacticFormula L) : L.lDef.TD ch := .ofZero (.mkSigma “x | x = ↑⌜p⌝” (by simp)) _ instance const_defined_const (p : FirstOrder.SyntacticFormula L) : (Language.Theory.singleton V L p).Defined (Language.Theory.singletonDef p) where - defined := .of_zero (by intro v; simp [numeral_eq_natCast]) + defined := .of_zero (by intro v; simp [numeral_eq_natCast, coe_quote]) end @@ -142,8 +142,8 @@ instance scheme_defined_scheme (φ : L.Scheme) {ps : pL.SchemeDef} [φ.Defined p (Language.Scheme.Defined.defined (V := V) (φ := φ)).df.iff, eq_comm]⟩ def Language.Craig.toScheme {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] (c : L.Craig) : L.Scheme where - scheme (x) := c.core x ^⋏ qqVerums 0 x - increasing (x) := le_trans (le_qqVerums 0 x) (le_of_lt <| by simp) + scheme (x) := c.core x ^⋏ qqVerums x + increasing (x) := le_trans (le_qqVerums x) (le_of_lt <| by simp) structure _root_.LO.FirstOrder.Arith.LDef.CraigDef (pL : LDef) where core : 𝚺₁.Semisentence 2 @@ -152,7 +152,7 @@ class Language.Craig.Defined (φ : L.Craig) (ps : outParam pL.CraigDef) : Prop w defined : 𝚺₁-Function₁ φ.core via ps.core def _root_.LO.FirstOrder.Arith.LDef.CraigDef.toSchemeDef {pL : LDef} (c : pL.CraigDef) : pL.SchemeDef where - core := .mkSigma “p x | ∃ p', !c.core p' x ∧ ∃ vs, !qqVerumsDef vs 0 x ∧ !qqAndDef p 0 p' vs” (by simp) + core := .mkSigma “p x | ∃ p', !c.core p' x ∧ ∃ vs, !qqVerumsDef vs x ∧ !qqAndDef p p' vs” (by simp) instance (φ : L.Craig) (c : pL.CraigDef) [φ.Defined c] : φ.toScheme.Defined c.toSchemeDef where defined := by intro v; simp [Language.Craig.toScheme, Arith.LDef.CraigDef.toSchemeDef, (Language.Craig.Defined.defined (φ := φ)).df.iff] @@ -191,22 +191,22 @@ section thEQ def eqRefl : ⌜ℒₒᵣ⌝[V].Theory := Language.Theory.singleton V ℒₒᵣ “∀ x, x = x” def eqReplaceC : ⌜ℒₒᵣ⌝[V].Craig where - core := fun p ↦ if ⌜ℒₒᵣ⌝.Semiformula 1 p then ^∀[0] ^∀[1] (^#1 ^=[2] ^#0 ^→[⌜ℒₒᵣ⌝; 2] ⌜ℒₒᵣ⌝.substs 2 ?[^#1] p ^→[⌜ℒₒᵣ⌝; 2] ⌜ℒₒᵣ⌝.substs 2 ?[^#0] p) else 0 + core := fun p ↦ if ⌜ℒₒᵣ⌝.IsSemiformula 1 p then ^∀ ^∀ (^#1 ^= ^#0 ^→[⌜ℒₒᵣ⌝] ⌜ℒₒᵣ⌝.substs ?[^#1] p ^→[⌜ℒₒᵣ⌝] ⌜ℒₒᵣ⌝.substs ?[^#0] p) else 0 def eqReplaceCDef : p⌜ℒₒᵣ⌝.CraigDef where core := .mkSigma “σ p | ( !p⌜ℒₒᵣ⌝.isSemiformulaDef.pi 1 p → let x0 := qqBvarDef 0; let x1 := qqBvarDef 1; - let eq := qqEQDef 2 x1 x0; + let eq := qqEQDef x1 x0; let v0 := mkVec₁Def x0; let v1 := mkVec₁Def x1; - let p0 := p⌜ℒₒᵣ⌝.substsDef 2 v1 p; - let p1 := p⌜ℒₒᵣ⌝.substsDef 2 v0 p; - let imp0 := p⌜ℒₒᵣ⌝.impDef 2 p0 p1; - let imp1 := p⌜ℒₒᵣ⌝.impDef 2 eq imp0; - let all0 := qqAllDef 1 imp1; - !qqAllDef σ 0all0 ) ∧ + let p0 := p⌜ℒₒᵣ⌝.substsDef v1 p; + let p1 := p⌜ℒₒᵣ⌝.substsDef v0 p; + let imp0 := p⌜ℒₒᵣ⌝.impDef p0 p1; + let imp1 := p⌜ℒₒᵣ⌝.impDef eq imp0; + let all0 := qqAllDef imp1; + !qqAllDef σ all0 ) ∧ ( ¬!p⌜ℒₒᵣ⌝.isSemiformulaDef.sigma 1 p → σ = 0)” (by simp) instance : (eqReplaceC (V := V)).Defined eqReplaceCDef where @@ -214,9 +214,9 @@ instance : (eqReplaceC (V := V)).Defined eqReplaceCDef where intro v simp [eqReplaceC, eqReplaceCDef, HierarchySymbol.Semiformula.val_sigma, - (semiformula_defined (LOR (V := V))).df.iff, (semiformula_defined (LOR (V := V))).proper.iff', - (substs_defined (LOR (V := V))).df.iff, (imp_defined (LOR (V := V))).df.iff] - by_cases h : ⌜ℒₒᵣ⌝.Semiformula 1 (v 1) <;> simp [h] + (Language.isSemiformula_defined (LOR (V := V))).df.iff, (Language.isSemiformula_defined (LOR (V := V))).proper.iff', + (Language.substs_defined (LOR (V := V))).df.iff, (Language.imp_defined (LOR (V := V))).df.iff] + by_cases h : ⌜ℒₒᵣ⌝.IsSemiformula 1 (v 1) <;> simp [h] variable (V) @@ -239,8 +239,6 @@ def TTheory.thEQ.eqRefl : ⌜𝐄𝐐'⌝[V] ⊢ (#'0 =' #'0).all := Language.Th simp [Language.Theory.tmem, TTheory.thEQ, Theory.EQ, FirstOrder.Semiformula.quote_all, FirstOrder.Semiformula.quote_eq, Semiformula.Operator.eq_def, Semiterm.quote_bvar] - - end thEQ /- diff --git a/lake-manifest.json b/lake-manifest.json index 597868a..5089206 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -75,7 +75,7 @@ "type": "git", "subDir": null, "scope": "", - "rev": "352b7bd19ddb940b29fb5abb91e94a53317fc5ba", + "rev": "3e8de075c396116e8dc88ae9d919e7d011a11efb", "name": "logic", "manifestFile": "lake-manifest.json", "inputRev": "master",