From 3faa3879ed876260bde653c1180662705405b0ec Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sat, 10 Aug 2024 02:07:07 +0900 Subject: [PATCH 01/10] refactor --- Arithmetization/Definability/Defined.lean | 864 ++++++++++ .../Definability/HSemiformula.lean | 415 +++++ Arithmetization/Definability/Hierarchy.lean | 1390 ----------------- .../{Definability.lean => PolyBounded.lean} | 14 +- Arithmetization/ISigmaOne/HFS/Fixpoint.lean | 10 +- Arithmetization/ISigmaOne/HFS/Seq.lean | 34 +- Arithmetization/ISigmaOne/Ind.lean | 6 +- .../ISigmaOne/Metamath/Formula/Basic.lean | 22 +- .../ISigmaOne/Metamath/Formula/Functions.lean | 16 +- 9 files changed, 1330 insertions(+), 1441 deletions(-) create mode 100644 Arithmetization/Definability/Defined.lean create mode 100644 Arithmetization/Definability/HSemiformula.lean delete mode 100644 Arithmetization/Definability/Hierarchy.lean rename Arithmetization/Definability/{Definability.lean => PolyBounded.lean} (99%) diff --git a/Arithmetization/Definability/Defined.lean b/Arithmetization/Definability/Defined.lean new file mode 100644 index 0000000..5f2805e --- /dev/null +++ b/Arithmetization/Definability/Defined.lean @@ -0,0 +1,864 @@ +import Arithmetization.Definability.HSemiformula +import Arithmetization.Vorspiel.Graph + +namespace LO.FirstOrder.Arith + +end Arith + +def Defined {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semisentence L k) : Prop := + ∀ v, R v ↔ Semiformula.Evalbm M v p + +def DefinedWithParam {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semiformula L M k) : Prop := + ∀ v, R v ↔ Semiformula.Evalm M v id p + +lemma Defined.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : + Semiformula.Evalbm M v p ↔ R v := (h v).symm + +lemma DefinedWithParam.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semiformula L M k} (h : DefinedWithParam R p) (v) : + Semiformula.Evalm M v id p ↔ R v := (h v).symm + +namespace Arith.HierarchySymbol + +variable (ξ : Type*) (n : ℕ) + +open LO.Arith + +variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] + +variable {Γ : HierarchySymbol} + +def Defined (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → Γ.Semisentence k → Prop + | 𝚺-[_], p => FirstOrder.Defined R p.val + | 𝚷-[_], p => FirstOrder.Defined R p.val + | 𝚫-[_], p => p.ProperOn M ∧ FirstOrder.Defined R p.val + +def DefinedWithParam (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → Γ.Semiformula M k → Prop + | 𝚺-[_], p => FirstOrder.DefinedWithParam R p.val + | 𝚷-[_], p => FirstOrder.DefinedWithParam R p.val + | 𝚫-[_], p => p.ProperWithParamOn M ∧ FirstOrder.DefinedWithParam R p.val + +variable (L Γ) + +class Definable {k} (P : (Fin k → M) → Prop) : Prop where + definable : ∃ p : Γ.Semisentence k, Defined P p + +class DefinableWithParam {k} (P : (Fin k → M) → Prop) : Prop where + definable : ∃ p : Γ.Semiformula M k, DefinedWithParam P p + +abbrev DefinedPred (P : M → Prop) (p : Γ.Semisentence 1) : Prop := + Defined (λ v ↦ P (v 0)) p + +abbrev DefinedRel (R : M → M → Prop) (p : Γ.Semisentence 2) : Prop := + Defined (λ v ↦ R (v 0) (v 1)) p + +abbrev DefinedRel₃ (R : M → M → M → Prop) (p : Γ.Semisentence 3) : Prop := + Defined (λ v ↦ R (v 0) (v 1) (v 2)) p + +abbrev DefinedRel₄ (R : M → M → M → M → Prop) (p : Γ.Semisentence 4) : Prop := + Defined (λ v ↦ R (v 0) (v 1) (v 2) (v 3)) p + +variable {L Γ} + +abbrev DefinedFunction {k} (f : (Fin k → M) → M) (p : Γ.Semisentence (k + 1)) : Prop := + Defined (fun v => v 0 = f (v ·.succ)) p + +variable (L Γ) + +abbrev DefinedFunction₁ (f : M → M) (p : Γ.Semisentence 2) : Prop := + DefinedFunction (fun v => f (v 0)) p + +abbrev DefinedFunction₂ (f : M → M → M) (p : Γ.Semisentence 3) : Prop := + DefinedFunction (fun v => f (v 0) (v 1)) p + +abbrev DefinedFunction₃ (f : M → M → M → M) (p : Γ.Semisentence 4) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2)) p + +abbrev DefinedFunction₄ (f : M → M → M → M → M) (p : Γ.Semisentence 5) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3)) p + +abbrev DefinedFunction₅ (f : M → M → M → M → M → M) (p : Γ.Semisentence 6) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3) (v 4)) p + +abbrev DefinableWithParamPred (P : M → Prop) : Prop := Γ.DefinableWithParam (k := 1) (fun v ↦ P (v 0)) + +abbrev DefinableWithParamRel (P : M → M → Prop) : Prop := Γ.DefinableWithParam (k := 2) (fun v ↦ P (v 0) (v 1)) + +abbrev DefinableWithParamRel₃ (P : M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) + +abbrev DefinableWithParamRel₄ (P : M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) + +abbrev DefinableWithParamRel₅ (P : M → M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) + +abbrev DefinableWithParamRel₆ (P : M → M → M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 6) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) + +abbrev DefinableWithParamFunction (f : (Fin k → M) → M) : Prop := Γ.DefinableWithParam (k := k + 1) (fun v ↦ v 0 = f (v ·.succ)) + +abbrev DefinableWithParamFunction₁ (f : M → M) : Prop := Γ.DefinableWithParamFunction (k := 1) (fun v ↦ f (v 0)) + +abbrev DefinableWithParamFunction₂ (f : M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 2) (fun v ↦ f (v 0) (v 1)) + +abbrev DefinableWithParamFunction₃ (f : M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) + +abbrev DefinableWithParamFunction₄ (f : M → M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) + +abbrev DefinableWithParamFunction₅ (f : M → M → M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 5) (fun v ↦ f (v 0) (v 1) (v 2) (v 3) (v 4)) + +variable {L Γ} + +section + +variable {k} {P Q : (Fin k → M) → Prop} + +namespace Defined + +lemma df {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semisentence k} (h : Defined R p) : FirstOrder.Defined R p.val := + match Γ with + | 𝚺-[_] => h + | 𝚷-[_] => h + | 𝚫-[_] => h.2 + +lemma proper {R : (Fin k → M) → Prop} {m} {p : 𝚫-[m].Semisentence k} (h : Defined R p) : p.ProperOn M := h.1 + +lemma of_zero {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : 𝚺₀.Semisentence k} (h : Defined R p) : Defined R (p.ofZero Γ) := + match Γ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simp, by intro _; simp [h.iff]⟩ + +lemma emb {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semisentence k} (h : Defined R p) : Defined R p.emb := + match Γ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ + +lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) + {p : Γ.Semisentence k} (H : Defined Q p) : Defined P p := by + rwa [show P = Q from by funext v; simp [h]] + +lemma to_definable {Γ : HierarchySymbol} (p : Γ.Semisentence k) (hP : Defined P p) : Γ.DefinableWithParam P := ⟨p.rew Rew.emb, by + match Γ with + | 𝚺-[_] => intro; simp [hP.iff] + | 𝚷-[_] => intro; simp [hP.iff] + | 𝚫-[_] => exact ⟨ + fun v ↦ by rcases p; simpa [HierarchySymbol.Semiformula.rew] using hP.proper.rew Rew.emb v, + by intro; simp [hP.df.iff]⟩⟩ + +lemma to_definable₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : + Γ.DefinableWithParam P := Defined.to_definable (p.ofZero Γ) hP.of_zero + +lemma to_definable_oRing (p : Γ.Semisentence k) (hP : Defined P p) : + Γ.DefinableWithParam P := Defined.to_definable p.emb hP.emb + +lemma to_definable_oRing₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : + Γ.DefinableWithParam P := Defined.to_definable₀ p.emb hP.emb + +end Defined + +namespace DefinedFunction + +lemma of_eq {f g : (Fin k → M) → M} (h : ∀ x, f x = g x) + {p : Γ.Semisentence (k + 1)} (H : DefinedFunction f p) : DefinedFunction g p := + Defined.of_iff (by intro; simp [h]) H + +lemma graph_delta {f : (Fin k → M) → M} {p : 𝚺-[m].Semisentence (k + 1)} + (h : DefinedFunction f p) : DefinedFunction f p.graphDelta := + ⟨by cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] + intro e; simp [Empty.eq_elim, h.df.iff] + rw [eq_comm], + by intro v; simp [h.df.iff]⟩ + +end DefinedFunction + +namespace DefinedWithParam + +lemma df {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semiformula M k} (h : DefinedWithParam R p) : FirstOrder.DefinedWithParam R p.val := + match Γ with + | 𝚺-[_] => h + | 𝚷-[_] => h + | 𝚫-[_] => h.2 + +lemma proper {R : (Fin k → M) → Prop} {m} {p : 𝚫-[m].Semiformula M k} (h : DefinedWithParam R p) : p.ProperWithParamOn M := h.1 + +lemma of_zero {R : (Fin k → M) → Prop} {Γ'} {p : Γ'-[0].Semiformula M k} + (h : DefinedWithParam R p) {Γ} : DefinedWithParam R (p.ofZero Γ) := + match Γ with + | 𝚺-[m] => by intro _; simp [h.df.iff] + | 𝚷-[m] => by intro _; simp [h.df.iff] + | 𝚫-[m] => ⟨by simp , by intro _; simp [h.df.iff]⟩ + +lemma of_deltaOne {R : (Fin k → M) → Prop} {Γ m} {p : 𝚫₁.Semiformula M k} + (h : DefinedWithParam R p) : DefinedWithParam R (p.ofDeltaOne Γ m) := + match Γ with + | 𝚺 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma] + | 𝚷 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, h.proper.iff'] + | 𝚫 => ⟨by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma, h.proper.iff'], + by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + +lemma emb {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semiformula M k} + (h : DefinedWithParam R p) : DefinedWithParam R p.emb := + match Γ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ + +lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) + {p : Γ.Semiformula M k} (H : DefinedWithParam Q p) : DefinedWithParam P p := by + rwa [show P = Q from by funext v; simp [h]] + +lemma to_definable {p : Γ.Semiformula M k} (h : DefinedWithParam P p) : Γ.DefinableWithParam P := ⟨p, h⟩ + +lemma to_definable₀ {p : Γ'-[0].Semiformula M k} + (h : DefinedWithParam P p) : Γ.DefinableWithParam P := ⟨p.ofZero Γ, h.of_zero⟩ + +lemma to_definable_deltaOne {p : 𝚫₁.Semiformula M k} {Γ m} + (h : DefinedWithParam P p) : Γ-[m + 1].DefinableWithParam P := ⟨p.ofDeltaOne Γ m, h.of_deltaOne⟩ + +variable {Γ : HierarchySymbol} + +lemma retraction {p : Γ.Semiformula M k} (hp : DefinedWithParam P p) (f : Fin k → Fin l) : + DefinedWithParam (fun v ↦ P fun i ↦ v (f i)) (p.rew <| Rew.substs fun x ↦ #(f x)) := + match Γ with + | 𝚺-[_] => by intro; simp [hp.df.iff] + | 𝚷-[_] => by intro; simp [hp.df.iff] + | 𝚫-[_] => ⟨hp.proper.rew _, by intro; simp [hp.df.iff]⟩ + +@[simp] lemma verum : + DefinedWithParam (fun _ ↦ True) (⊤ : Γ.Semiformula M k) := + match Γ with + | 𝚺-[m] => by intro v; simp + | 𝚷-[m] => by intro v; simp + | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ + +@[simp] lemma falsum : + DefinedWithParam (fun _ ↦ False) (⊥ : Γ.Semiformula M k) := + match Γ with +| 𝚺-[m] => by intro v; simp + | 𝚷-[m] => by intro v; simp + | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ + +lemma and {p q : Γ.Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ∧ Q x) (p ⋏ q) := + match Γ with + | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚫-[m] => ⟨hp.proper.and hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ + +lemma or {p q : Γ.Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ∨ Q x) (p ⋎ q) := + match Γ with + | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚫-[m] => ⟨hp.proper.or hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ + +lemma negSigma {p : 𝚺-[m].Semiformula M k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) p.negSigma := by intro v; simp [hp.iff] + +lemma negPi {p : 𝚷-[m].Semiformula M k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) p.negPi := by intro v; simp [hp.iff] + +lemma not {p : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) (~p) := ⟨hp.proper.neg, by intro v; simp [hp.proper.eval_neg, hp.df.iff]⟩ + +lemma imp {p q : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x → Q x) (p ⟶ q) := (hp.not.or hq).of_iff (by intro x; simp [imp_iff_not_or]) + +lemma iff {p q : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ↔ Q x) (p ⟷ q) := ((hp.imp hq).and (hq.imp hp)).of_iff <| by intro v; simp [iff_iff_implies_and_implies] + +lemma ball {P : (Fin (k + 1) → M) → Prop} {p : Γ.Semiformula M (k + 1)} + (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ M k) : + DefinedWithParam (fun v ↦ ∀ x < t.valm M v id, P (x :> v)) (HierarchySymbol.Semiformula.ball t p) := + match Γ with + | 𝚺-[m] => by intro v; simp [hp.df.iff] + | 𝚷-[m] => by intro v; simp [hp.df.iff] + | 𝚫-[m] => ⟨hp.proper.ball, by intro v; simp [hp.df.iff]⟩ + +lemma bex {P : (Fin (k + 1) → M) → Prop} {p : Γ.Semiformula M (k + 1)} + (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ M k) : + DefinedWithParam (fun v ↦ ∃ x < t.valm M v id, P (x :> v)) (HierarchySymbol.Semiformula.bex t p) := + match Γ with + | 𝚺-[m] => by intro v; simp [hp.df.iff] + | 𝚷-[m] => by intro v; simp [hp.df.iff] + | 𝚫-[m] => ⟨hp.proper.bex, by intro v; simp [hp.df.iff]⟩ + +lemma ex {P : (Fin (k + 1) → M) → Prop} {p : 𝚺-[m + 1].Semiformula M (k + 1)} + (hp : DefinedWithParam P p) : + DefinedWithParam (fun v ↦ ∃ x, P (x :> v)) p.ex := by intro _; simp [hp.df.iff] + +lemma all {P : (Fin (k + 1) → M) → Prop} {p : 𝚷-[m + 1].Semiformula M (k + 1)} + (hp : DefinedWithParam P p) : + DefinedWithParam (fun v ↦ ∀ x, P (x :> v)) p.all := by intro _; simp [hp.df.iff] + +end DefinedWithParam + +namespace DefinableWithParam + +lemma mkPolarity {P : (Fin k → M) → Prop} {Γ : Polarity} + (p : Semiformula ℒₒᵣ M k) (hp : Hierarchy Γ m p) (hP : ∀ v, P v ↔ Semiformula.Evalm M v id p) : Γ-[m].DefinableWithParam P := + match Γ with + | 𝚺 => ⟨.mkSigma p hp, by intro v; simp [hP]⟩ + | 𝚷 => ⟨.mkPi p hp, by intro v; simp [hP]⟩ + +lemma of_iff (Q : (Fin k → M) → Prop) (h : ∀ x, P x ↔ Q x) (H : Γ.DefinableWithParam Q) : Γ.DefinableWithParam P := by + rwa [show P = Q from by funext v; simp [h]] + +lemma of_oRing (h : Γ.DefinableWithParam P) : Γ.DefinableWithParam P := by + rcases h with ⟨p, hP⟩; exact ⟨p.emb, hP.emb⟩ + +lemma of_delta (h : 𝚫-[m].DefinableWithParam P) {Γ} : Γ-[m].DefinableWithParam P := by + rcases h with ⟨p, h⟩ + match Γ with + | 𝚺 => exact ⟨p.sigma, by intro v; simp [HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ + | 𝚷 => exact ⟨p.pi, by intro v; simp [←h.proper v, HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ + | 𝚫 => exact ⟨p, h⟩ + +instance [𝚫-[m].DefinableWithParam P] (Γ) : Γ-[m].DefinableWithParam P := of_delta inferInstance + +lemma of_sigma_of_pi (hσ : 𝚺-[m].DefinableWithParam P) (hπ : 𝚷-[m].DefinableWithParam P) : 𝚫-[m].DefinableWithParam P := by + rcases hσ with ⟨p, hp⟩; rcases hπ with ⟨q, hq⟩ + exact ⟨.mkDelta p q, by intro v; simp [hp.df.iff, hq.df.iff], by intro v; simp [hp.df.iff]⟩ + +lemma of_zero (h : Γ'-[0].DefinableWithParam P) : Γ.DefinableWithParam P := by + rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable₀ + +lemma of_deltaOne (h : 𝚫₁.DefinableWithParam P) (Γ m) : Γ-[m + 1].DefinableWithParam P := by + rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable_deltaOne + +instance [𝚺₀.DefinableWithParam P] (Γ : HierarchySymbol) : Γ.DefinableWithParam P := DefinableWithParam.of_zero (Γ' := 𝚺) (Γ := Γ) inferInstance + +variable {Γ : HierarchySymbol} + +lemma retraction (h : Γ.DefinableWithParam P) (f : Fin k → Fin n) : + Γ.DefinableWithParam fun v ↦ P (fun i ↦ v (f i)) := by + rcases h with ⟨p, h⟩ + exact ⟨p.rew (Rew.substs (fun i ↦ #(f i))), + match Γ with + | 𝚺-[_] => by intro; simp [h.df.iff] + | 𝚷-[_] => by intro; simp [h.df.iff] + | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ + +lemma retractiont (h : Γ.DefinableWithParam P) (f : Fin k → Semiterm ℒₒᵣ M n) : + Γ.DefinableWithParam fun v ↦ P (fun i ↦ Semiterm.valm M v id (f i)) := by + rcases h with ⟨p, h⟩ + exact ⟨p.rew (Rew.substs f), + match Γ with + | 𝚺-[_] => by intro; simp [h.df.iff] + | 𝚷-[_] => by intro; simp [h.df.iff] + | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ + +lemma const {P : Prop} : Γ.DefinableWithParam (fun _ : Fin k → M ↦ P) := of_zero (by + by_cases hP : P + · exact ⟨.mkSigma ⊤ (by simp), by intro; simp[hP]⟩ + · exact ⟨.mkSigma ⊥ (by simp), by intro; simp[hP]⟩) + +lemma and (h₁ : Γ.DefinableWithParam P) (h₂ : Γ.DefinableWithParam Q) : + Γ.DefinableWithParam (fun v ↦ P v ∧ Q v) := by + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁ ⋏ p₂, h₁.and h₂⟩ + +lemma or (h₁ : Γ.DefinableWithParam P) (h₂ : Γ.DefinableWithParam Q) : + Γ.DefinableWithParam (fun v ↦ P v ∨ Q v) := by + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁ ⋎ p₂, h₁.or h₂⟩ + +lemma not {Γ : SigmaPiDelta} (h : Γ.alt-[m].DefinableWithParam P) : + Γ-[m].DefinableWithParam (fun v ↦ ¬P v) := by + match Γ with + | 𝚺 => rcases h with ⟨p, h⟩; exact ⟨p.negPi, h.negPi⟩ + | 𝚷 => rcases h with ⟨p, h⟩; exact ⟨p.negSigma, h.negSigma⟩ + | 𝚫 => rcases h with ⟨p, h⟩; exact ⟨p.negDelta, h.not⟩ + +lemma imp {Γ : SigmaPiDelta} (h₁ : Γ.alt-[m].DefinableWithParam P) (h₂ : Γ-[m].DefinableWithParam Q) : + Γ-[m].DefinableWithParam (fun v ↦ P v → Q v) := by + match Γ with + | 𝚺 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁.negPi.or p₂, (h₁.negPi.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ + | 𝚷 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁.negSigma.or p₂, (h₁.negSigma.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ + | 𝚫 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩; exact ⟨p₁ ⟶ p₂, h₁.imp h₂⟩ + +lemma iff (h₁ : 𝚫-[m].DefinableWithParam P) (h₂ : 𝚫-[m].DefinableWithParam Q) {Γ} : + Γ-[m].DefinableWithParam (fun v ↦ P v ↔ Q v) := + .of_delta (by rcases h₁ with ⟨p, hp⟩; rcases h₂ with ⟨q, hq⟩; exact ⟨p ⟷ q, hp.iff hq⟩) + +lemma all {P : (Fin k → M) → M → Prop} (h : 𝚷-[s + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + 𝚷-[s + 1].DefinableWithParam (fun v ↦ ∀ x, P v x) := by + rcases h with ⟨p, hp⟩ + exact ⟨.mkPi (∀' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ + +lemma ex {P : (Fin k → M) → M → Prop} (h : 𝚺-[s + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + 𝚺-[s + 1].DefinableWithParam (fun v ↦ ∃ x, P v x) := by + rcases h with ⟨p, hp⟩ + exact ⟨.mkSigma (∃' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ + +lemma comp₁ {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : 𝚺-[m + 1].DefinableWithParamFunction f) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamPred P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f v)) := by + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ + exact ⟨(pf ⋏ (p.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ + exact ⟨(pf.negSigma ⋎ (p.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ + exact of_sigma_of_pi + ⟨(pf ⋏ (p.sigma.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ + ⟨(pf.negSigma ⋎ (p.pi.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₁_infer {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : 𝚺-[m + 1].DefinableWithParamFunction f) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamPred P] : Γ-[m + 1].DefinableWithParam fun v ↦ P (f v) := + comp₁ hf inferInstance + +lemma comp₂ {k} {P : M → M → Prop} {f g : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel P) : Γ-[m + 1].DefinableWithParam fun v ↦ P (f v) (g v) := by + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ + exact ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.rew (Rew.substs ![#0, #1]))).ex.ex, by + intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ + exact ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.rew (Rew.substs ![#0, #1]))).all.all, by + intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ + exact of_sigma_of_pi + ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.sigma.rew (Rew.substs ![#0, #1]))).ex.ex, by + intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma + ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.pi.rew (Rew.substs ![#0, #1]))).all.all, by + intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₂_infer {k} {P : M → M → Prop} {f g : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f v) (g v)) := + comp₂ hf hg inferInstance + +lemma comp₃ {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₃ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := by + rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) + ⋏ (p.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma + ⋎ (p.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩ + exact of_sigma_of_pi + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) + ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma + ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₃_infer {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₃ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := + comp₃ hf₁ hf₂ hf₃ inferInstance + +lemma comp₄ {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₄ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := by + rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩; rcases hf₄ with ⟨pf₄, hpf₄⟩ + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) + ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩ + exact of_sigma_of_pi + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) + ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma + ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₄_infer {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₄ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := + comp₄ hf₁ hf₂ hf₃ hf₄ inferInstance + +lemma comp₅ {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₅ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := by + rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ + rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩ + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) + ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩ + exact of_sigma_of_pi + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) + ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) + ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma + ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₅_infer {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₅ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := + comp₅ hf₁ hf₂ hf₃ hf₄ hf₅ inferInstance + +lemma comp₆ {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) (hf₆ : 𝚺-[m + 1].DefinableWithParamFunction f₆) + {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₆ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := by + rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ + rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩; rcases hf₆ with ⟨pf₆, hpf₆⟩ + match Γ with + | 𝚺 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff]⟩ + | 𝚷 => + rcases hP with ⟨p, hp⟩ + exact + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or]⟩ + | 𝚫 => + rcases hP with ⟨p, hp⟩ + exact of_sigma_of_pi + ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) + ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma + ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by + intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ + +lemma comp₆_infer {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} + (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) + (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) + (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) (hf₆ : 𝚺-[m + 1].DefinableWithParamFunction f₆) + {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₆ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := + comp₆ hf₁ hf₂ hf₃ hf₄ hf₅ hf₆ inferInstance + +end DefinableWithParam + +lemma DefinableWithParamPred.of_iff {P : M → Prop} (Q : M → Prop) (h : ∀ x, P x ↔ Q x) (H : Γ.DefinableWithParamPred Q) : Γ.DefinableWithParamPred P := by + rwa [show P = Q from by funext v; simp [h]] + +instance DefinableWithParamFunction₁.graph {f : M → M} [h : Γ.DefinableWithParamFunction₁ f] : + Γ.DefinableWithParamRel (Function.Graph f) := h + +instance DefinableWithParamFunction₂.graph {f : M → M → M} [h : Γ.DefinableWithParamFunction₂ f] : + Γ.DefinableWithParamRel₃ (Function.Graph₂ f) := h + +instance DefinableWithParamFunction₃.graph {f : M → M → M → M} [h : Γ.DefinableWithParamFunction₃ f] : + Γ.DefinableWithParamRel₄ (Function.Graph₃ f) := h + +namespace DefinableWithParamFunction + +lemma graph_delta {k} {f : (Fin k → M) → M} + (h : 𝚺-[m].DefinableWithParamFunction f) : 𝚫-[m].DefinableWithParamFunction f := by + rcases h with ⟨p, h⟩ + exact ⟨p.graphDelta, by + cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] + intro e; simp [Empty.eq_elim, h.df.iff] + exact eq_comm, by + intro v; simp [h.df.iff]⟩ + +instance {k} {f : (Fin k → M) → M} [h : 𝚺-[m].DefinableWithParamFunction f] : 𝚫-[m].DefinableWithParamFunction f := + DefinableWithParamFunction.graph_delta h + +instance {k} {f : (Fin k → M) → M} [𝚺₀.DefinableWithParamFunction f] : Γ.DefinableWithParamFunction f := inferInstance + +lemma of_sigmaOne {k} {f : (Fin k → M) → M} + (h : 𝚺₁.DefinableWithParamFunction f) (Γ m) : Γ-[m + 1].DefinableWithParamFunction f := DefinableWithParam.of_deltaOne (graph_delta h) Γ m + +@[simp] lemma var {k} (i : Fin k) : Γ.DefinableWithParamFunction (fun v : Fin k → M ↦ v i) := + .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!#i.succ” (by simp), by intro _; simp⟩ + +@[simp] lemma const {k} (c : M) : Γ.DefinableWithParamFunction (fun _ : Fin k → M ↦ c) := + .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | #0 = &c” (by simp), by intro v; simp⟩ + +@[simp] lemma term_retraction (t : Semiterm ℒₒᵣ M n) (e : Fin n → Fin k) : + Γ.DefinableWithParamFunction fun v : Fin k → M ↦ Semiterm.valm M (fun x ↦ v (e x)) id t := + .of_zero (Γ' := 𝚺) + ⟨.mkSigma “x | x = !!(Rew.substs (fun x ↦ #(e x).succ) t)” (by simp), by intro v; simp [Semiterm.val_substs]⟩ + +@[simp] lemma term (t : Semiterm ℒₒᵣ M k) : + Γ.DefinableWithParamFunction fun v : Fin k → M ↦ Semiterm.valm M v id t := + .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!(Rew.bShift t)” (by simp), by intro v; simp [Semiterm.val_bShift']⟩ + +lemma of_eq {f : (Fin k → M) → M} (g) (h : ∀ v, f v = g v) (H : Γ.DefinableWithParamFunction f) : Γ.DefinableWithParamFunction g := by + rwa [show g = f from by funext v; simp [h]] + +lemma retraction {n k} {f : (Fin k → M) → M} (hf : Γ.DefinableWithParamFunction f) (e : Fin k → Fin n) : + Γ.DefinableWithParamFunction fun v ↦ f (fun i ↦ v (e i)) := by + have := DefinableWithParam.retraction (n := n + 1) hf (0 :> fun i ↦ (e i).succ); simp at this + exact this.of_iff _ (by intro x; simp) + +lemma retractiont {f : (Fin k → M) → M} (hf : Γ.DefinableWithParamFunction f) (t : Fin k → Semiterm ℒₒᵣ M n) : + Γ.DefinableWithParamFunction fun v ↦ f (fun i ↦ Semiterm.valm M v id (t i)) := by + have := DefinableWithParam.retractiont (n := n + 1) hf (#0 :> fun i ↦ Rew.bShift (t i)); simp at this + exact this.of_iff _ (by intro x; simp [Semiterm.val_bShift']) + +lemma rel {f : (Fin k → M) → M} (h : Γ.DefinableWithParamFunction f) : + Γ.DefinableWithParam (fun v ↦ v 0 = f (v ·.succ)) := h + +end DefinableWithParamFunction + +lemma DefinableWithParamFunction₁.comp {Γ} {k} {f : M → M} {g : (Fin k → M) → M} + (hf : Γ-[m + 1].DefinableWithParamFunction₁ f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g v)) := by + simpa using DefinableWithParam.comp₂ (P := Function.Graph f) (DefinableWithParamFunction.var 0) (hg.retraction Fin.succ) hf + +lemma DefinableWithParamFunction₂.comp {Γ} {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} + (hf : Γ-[m + 1].DefinableWithParamFunction₂ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v)) := by + simpa using DefinableWithParam.comp₃ (P := Function.Graph₂ f) (DefinableWithParamFunction.var 0) (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) hf + +lemma DefinableWithParamFunction₃.comp {Γ} {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} + (hf : Γ-[m + 1].DefinableWithParamFunction₃ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) + (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := by + simpa using DefinableWithParam.comp₄ (P := Function.Graph₃ f) (DefinableWithParamFunction.var 0) + (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) hf + +lemma DefinableWithParamFunction₄.comp {Γ} {k} {f : M → M → M → M → M} {g₁ g₂ g₃ g₄ : (Fin k → M) → M} + (hf : Γ-[m + 1].DefinableWithParamFunction₄ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) + (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) + (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := by + simpa using DefinableWithParam.comp₅ (P := Function.Graph₄ f) (DefinableWithParamFunction.var 0) + (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) (hg₄.retraction Fin.succ) hf + +lemma DefinableWithParamFunction₅.comp {Γ} {k} {f : M → M → M → M → M → M} {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} + (hf : Γ-[m + 1].DefinableWithParamFunction₅ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) + (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) + (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) (hg₅ : 𝚺-[m + 1].DefinableWithParamFunction g₅) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := by + simpa using DefinableWithParam.comp₆ (P := Function.Graph₅ f) (DefinableWithParamFunction.var 0) + (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) + (hg₄.retraction Fin.succ) (hg₅.retraction Fin.succ) hf + +lemma DefinableWithParamFunction.comp₁_infer {Γ} {k} {f : M → M} [Γ-[m + 1].DefinableWithParamFunction₁ f] + {g : (Fin k → M) → M} (hg : 𝚺-[m + 1].DefinableWithParamFunction g) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g v)) := + DefinableWithParamFunction₁.comp inferInstance hg + +lemma DefinableWithParamFunction.comp₂_infer {Γ} {k} {f : M → M → M} [Γ-[m + 1].DefinableWithParamFunction₂ f] + {g₁ g₂ : (Fin k → M) → M} (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v)) := + DefinableWithParamFunction₂.comp inferInstance hg₁ hg₂ + +lemma DefinableWithParamFunction.comp₃_infer {Γ} {k} {f : M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₃ f] + {g₁ g₂ g₃ : (Fin k → M) → M} + (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := + DefinableWithParamFunction₃.comp inferInstance hg₁ hg₂ hg₃ + +lemma DefinableWithParamFunction.comp₄_infer {Γ} {k} {f : M → M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₄ f] + {g₁ g₂ g₃ g₄ : (Fin k → M) → M} + (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) + (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := + DefinableWithParamFunction₄.comp inferInstance hg₁ hg₂ hg₃ hg₄ + +lemma DefinableWithParamFunction.comp₅_infer {Γ} {k} {f : M → M → M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₅ f] + {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} + (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) + (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) + (hg₅ : 𝚺-[m + 1].DefinableWithParamFunction g₅) : + Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := + DefinableWithParamFunction₅.comp inferInstance hg₁ hg₂ hg₃ hg₄ hg₅ + +namespace DefinableWithParamRel + +@[simp] instance eq : Γ.DefinableWithParamRel (Eq : M → M → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1” (by simp)) (by intro _; simp) + +@[simp] instance lt : Γ.DefinableWithParamRel (LT.lt : M → M → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 < #1” (by simp)) (by intro _; simp) + +@[simp] instance le : Γ.DefinableWithParamRel (LE.le : M → M → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 ≤ #1” (by simp)) (by intro _; simp) + +end DefinableWithParamRel + +namespace DefinableWithParamFunction₂ + +@[simp] instance add : Γ.DefinableWithParamFunction₂ ((· + ·) : M → M → M) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) + +@[simp] instance mul : Γ.DefinableWithParamFunction₂ ((· * ·) : M → M → M) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) + +@[simp] instance hAdd : Γ.DefinableWithParamFunction₂ (HAdd.hAdd : M → M → M) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) + +@[simp] instance hMul : Γ.DefinableWithParamFunction₂ (HMul.hMul : M → M → M) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) + +end DefinableWithParamFunction₂ + +namespace DefinableWithParam + +lemma ball_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x < f v, P v x) := by + rcases hf with ⟨bf, hbf⟩ + rcases h with ⟨p, hp⟩ + match Γ with + | 𝚺 => exact + ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚷 => exact + ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚫 => + exact .of_sigma_of_pi + ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ + ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ + +lemma bex_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x < f v, P v x) := by + rcases hf with ⟨bf, hbf⟩ + rcases h with ⟨p, hp⟩ + match Γ with + | 𝚺 => exact + ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚷 => exact + ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚫 => + exact .of_sigma_of_pi + ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ + ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ + +lemma ball_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x ≤ f v, P v x) := by + have : Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x < f v + 1, P v x) := ball_lt (DefinableWithParamFunction₂.comp (by simp) hf (by simp)) h + exact this.of_iff <| by intro v; simp [lt_succ_iff_le] + +lemma bex_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x ≤ f v, P v x) := by + have : Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x < f v + 1, P v x) := bex_lt (DefinableWithParamFunction₂.comp (by simp) hf (by simp)) h + exact this.of_iff <| by intro v; simp [lt_succ_iff_le] + +lemma ball_lt' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ {x}, x < f v → P v x) := ball_lt hf h + +lemma ball_le' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} + (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ {x}, x ≤ f v → P v x) := ball_le hf h + +end DefinableWithParam + +end + +end Arith.HierarchySymbol + +end LO.FirstOrder diff --git a/Arithmetization/Definability/HSemiformula.lean b/Arithmetization/Definability/HSemiformula.lean new file mode 100644 index 0000000..d801de8 --- /dev/null +++ b/Arithmetization/Definability/HSemiformula.lean @@ -0,0 +1,415 @@ +import Arithmetization.Vorspiel.Lemmata +import Logic.FirstOrder.Arith.StrictHierarchy + +/-! + +# Arithmetical Formula Sorted by Arithmetical Hierarchy + +This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of first-order logic. + +- `𝚺-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚺-[m]`. +- `𝚷-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚷-[m]`. +- `𝚫-[m].Semiformula ξ n` is a pair of `𝚺-[m].Semiformula ξ n` and `𝚷-[m].Semiformula ξ n`. +- `ProperOn` : `p.ProperOn M` iff `p`'s two element `p.sigma` and `p.pi` are equivalent on model `M`. + +-/ + +namespace LO.FirstOrder.Arith + +structure HierarchySymbol where + Γ : SigmaPiDelta + rank : ℕ + +scoped notation:max Γ:max "-[" n "]" => HierarchySymbol.mk Γ n + +abbrev HierarchySymbol.sigmaZero : HierarchySymbol := 𝚺-[0] + +abbrev HierarchySymbol.piZero : HierarchySymbol := 𝚷-[0] + +abbrev HierarchySymbol.deltaZero : HierarchySymbol := 𝚫-[0] + +abbrev HierarchySymbol.sigmaOne : HierarchySymbol := 𝚺-[1] + +abbrev HierarchySymbol.piOne : HierarchySymbol := 𝚷-[1] + +abbrev HierarchySymbol.deltaOne : HierarchySymbol := 𝚫-[1] + +notation "𝚺₀" => HierarchySymbol.sigmaZero + +notation "𝚷₀" => HierarchySymbol.piZero + +notation "𝚫₀" => HierarchySymbol.deltaZero + +notation "𝚺₁" => HierarchySymbol.sigmaOne + +notation "𝚷₁" => HierarchySymbol.piOne + +notation "𝚫₁" => HierarchySymbol.deltaOne + +namespace HierarchySymbol + +variable (ξ : Type*) (n : ℕ) + +protected inductive Semiformula : HierarchySymbol → Type _ where + | mkSigma {m} : (p : Semiformula ℒₒᵣ ξ n) → Hierarchy 𝚺 m p → 𝚺-[m].Semiformula + | mkPi {m} : (p : Semiformula ℒₒᵣ ξ n) → Hierarchy 𝚷 m p → 𝚷-[m].Semiformula + | mkDelta {m} : 𝚺-[m].Semiformula → 𝚷-[m].Semiformula → 𝚫-[m].Semiformula + +protected abbrev Semisentence (Γ : HierarchySymbol) (n : ℕ) := Γ.Semiformula Empty n + +variable {Γ : HierarchySymbol} + +variable {ξ n} + +namespace Semiformula + +def val {Γ : HierarchySymbol} : Γ.Semiformula ξ n → Semiformula ℒₒᵣ ξ n + | mkSigma p _ => p + | mkPi p _ => p + | mkDelta p _ => p.val + +@[simp] lemma val_mkSigma (p : Semiformula ℒₒᵣ ξ n) (hp : Hierarchy 𝚺 m p) : (mkSigma p hp).val = p := rfl + +@[simp] lemma val_mkPi (p : Semiformula ℒₒᵣ ξ n) (hp : Hierarchy 𝚷 m p) : (mkPi p hp).val = p := rfl + +@[simp] lemma val_mkDelta (p : 𝚺-[m].Semiformula ξ n) (q : 𝚷-[m].Semiformula ξ n) : (mkDelta p q).val = p.val := rfl + +instance : Coe (𝚺₀.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ +instance : Coe (𝚷₀.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ +instance : Coe (𝚫₀.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ + +instance : Coe (𝚺₁.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ +instance : Coe (𝚷₁.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ +instance : Coe (𝚫₁.Semisentence n) (Semisentence ℒₒᵣ n) := ⟨Semiformula.val⟩ + +@[simp] lemma sigma_prop : (p : 𝚺-[m].Semiformula ξ n) → Hierarchy 𝚺 m p.val + | mkSigma _ h => h + +@[simp] lemma pi_prop : (p : 𝚷-[m].Semiformula ξ n) → Hierarchy 𝚷 m p.val + | mkPi _ h => h + +@[simp] lemma polarity_prop : {Γ : Polarity} → (p : Γ-[m].Semiformula ξ n) → Hierarchy Γ m p.val + | 𝚺, p => p.sigma_prop + | 𝚷, p => p.pi_prop + +def sigma : 𝚫-[m].Semiformula ξ n → 𝚺-[m].Semiformula ξ n + | mkDelta p _ => p + +@[simp] lemma sigma_mkDelta (p : 𝚺-[m].Semiformula ξ n) (q : 𝚷-[m].Semiformula ξ n) : (mkDelta p q).sigma = p := rfl + +def pi : 𝚫-[m].Semiformula ξ n → 𝚷-[m].Semiformula ξ n + | mkDelta _ p => p + +@[simp] lemma pi_mkDelta (p : 𝚺-[m].Semiformula ξ n) (q : 𝚷-[m].Semiformula ξ n) : (mkDelta p q).pi = q := rfl + +lemma val_sigma (p : 𝚫-[m].Semiformula ξ n) : p.sigma.val = p.val := by rcases p; simp + +def mkPolarity (p : Semiformula ℒₒᵣ ξ n) : (Γ : Polarity) → Hierarchy Γ m p → Γ-[m].Semiformula ξ n + | 𝚺, h => mkSigma p h + | 𝚷, h => mkPi p h + +@[simp] lemma val_mkPolarity (p : Semiformula ℒₒᵣ ξ n) {Γ} (h : Hierarchy Γ m p) : (mkPolarity p Γ h).val = p := by cases Γ <;> rfl + +@[simp] lemma hierarchy_sigma (p : 𝚺-[m].Semiformula ξ n) : Hierarchy 𝚺 m p.val := p.sigma_prop + +@[simp] lemma hierarchy_pi (p : 𝚷-[m].Semiformula ξ n) : Hierarchy 𝚷 m p.val := p.pi_prop + +@[simp] lemma hierarchy_zero {Γ Γ' m} (p : Γ-[0].Semiformula ξ n) : Hierarchy Γ' m p.val := by + cases Γ + · exact Hierarchy.of_zero p.sigma_prop + · exact Hierarchy.of_zero p.pi_prop + · cases p + simp; exact Hierarchy.of_zero (sigma_prop _) + +variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] + +variable (M) + +def ProperOn (p : 𝚫-[m].Semisentence n) : Prop := + ∀ (e : Fin n → M), Semiformula.Evalbm M e p.sigma.val ↔ Semiformula.Evalbm M e p.pi.val + +def ProperWithParamOn (p : 𝚫-[m].Semiformula M n) : Prop := + ∀ (e : Fin n → M), Semiformula.Evalm M e id p.sigma.val ↔ Semiformula.Evalm M e id p.pi.val + +variable {M} + +lemma ProperOn.iff {p : 𝚫-[m].Semisentence n} + (h : p.ProperOn M) (e : Fin n → M) : + Semiformula.Evalbm M e p.sigma.val ↔ Semiformula.Evalbm M e p.pi.val := h e + +lemma ProperWithParamOn.iff {p : 𝚫-[m].Semiformula M n} + (h : p.ProperWithParamOn M) (e : Fin n → M) : + Semiformula.Evalm M e id p.sigma.val ↔ Semiformula.Evalm (L := ℒₒᵣ) M e id p.pi.val := h e + +lemma ProperOn.iff' {p : 𝚫-[m].Semisentence n} + (h : p.ProperOn M) (e : Fin n → M) : + Semiformula.Evalbm M e p.pi.val ↔ Semiformula.Evalbm M e p.val := by simp [←h.iff, val_sigma] + +lemma ProperWithParamOn.iff' {p : 𝚫-[m].Semiformula M n} + (h : p.ProperWithParamOn M) (e : Fin n → M) : + Semiformula.Evalm M e id p.pi.val ↔ Semiformula.Evalm (L := ℒₒᵣ) M e id p.val := by simp [←h.iff, val_sigma] + +def rew (ω : Rew ℒₒᵣ ξ₁ n₁ ξ₂ n₂) : {Γ : HierarchySymbol} → Γ.Semiformula ξ₁ n₁ → Γ.Semiformula ξ₂ n₂ + | 𝚺-[_], mkSigma p hp => mkSigma (ω.hom p) (by simpa using hp) + | 𝚷-[_], mkPi p hp => mkPi (ω.hom p) (by simpa using hp) + | 𝚫-[_], mkDelta p q => mkDelta (p.rew ω) (q.rew ω) + +@[simp] lemma val_rew (ω : Rew ℒₒᵣ ξ₁ n₁ ξ₂ n₂) {Γ : HierarchySymbol} (p : Γ.Semiformula ξ₁ n₁) : (p.rew ω).val = ω.hom p.val := by + rcases Γ with ⟨Γ, m⟩; rcases p with (_ | _ | ⟨⟨p, _⟩, ⟨q, _⟩⟩) <;> simp [rew] + +@[simp] lemma ProperOn.rew {p : 𝚫-[m].Semisentence n₁} (h : p.ProperOn M) (ω : Rew ℒₒᵣ Empty n₁ Empty n₂) : (p.rew ω).ProperOn M := by + rcases p; simp [ProperOn, Semiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] + intro e; exact h.iff _ + +@[simp] lemma ProperOn.rew' {p : 𝚫-[m].Semisentence n₁} (h : p.ProperOn M) (ω : Rew ℒₒᵣ Empty n₁ M n₂) : (p.rew ω).ProperWithParamOn M := by + rcases p; intro e; simp [ProperOn, Semiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] + simpa using h.iff _ + +@[simp] lemma ProperWithParamOn.rew {p : 𝚫-[m].Semiformula M n₁} + (h : p.ProperWithParamOn M) (f : Fin n₁ → Semiterm ℒₒᵣ M n₂) : (p.rew (Rew.substs f)).ProperWithParamOn M := by + rcases p; intro e; + simp [ProperOn, Semiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] + exact h.iff _ + +def emb : {Γ : HierarchySymbol} → Γ.Semiformula ξ n → Γ.Semiformula ξ n + | 𝚺-[_], mkSigma p hp => mkSigma (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) + | 𝚷-[_], mkPi p hp => mkPi (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) + | 𝚫-[_], mkDelta p q => mkDelta p.emb q.emb + +@[simp] lemma val_emb {Γ : HierarchySymbol} (p : Γ.Semiformula ξ n) : p.emb.val = Semiformula.lMap Language.oringEmb p.val := by + rcases Γ with ⟨Γ, m⟩; rcases p with (_ | _ | ⟨⟨p, _⟩, ⟨q, _⟩⟩) <;> simp [rew, val] + +@[simp] lemma pi_emb (p : 𝚫-[m].Semiformula ξ n) : p.emb.pi = p.pi.emb := by cases p; rfl + +@[simp] lemma sigma_emb (p : 𝚫-[m].Semiformula ξ n) : p.emb.sigma = p.sigma.emb := by cases p; rfl + +@[simp] lemma emb_proper (p : 𝚫-[m].Semisentence n) : p.emb.ProperOn M ↔ p.ProperOn M := by + rcases p; simp [ProperOn, emb] + +@[simp] lemma emb_properWithParam (p : 𝚫-[m].Semiformula M n) : p.emb.ProperWithParamOn M ↔ p.ProperWithParamOn M := by + rcases p; simp [ProperWithParamOn, emb] + +def extd {Γ : HierarchySymbol} : Γ.Semiformula ξ n → Γ.Semiformula ξ n + | mkSigma p hp => mkSigma (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) + | mkPi p hp => mkPi (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) + | mkDelta p q => mkDelta p.extd q.extd + +@[simp] +lemma eval_extd_iff {e ε} {p : Γ.Semiformula ξ n} : + Semiformula.Evalm M e ε p.extd.val ↔ Semiformula.Evalm M e ε p.val := by + induction p <;> simp [extd, *] + +lemma ProperOn.extd {p : 𝚫-[m].Semisentence n} (h : p.ProperOn M) : p.extd.ProperOn M := by + intro e; rcases p; simpa [Semiformula.extd] using h.iff e + +lemma ProperWithParamOn.extd {p : 𝚫-[m].Semisentence n} (h : p.ProperOn M) : p.extd.ProperOn M := by + intro e; rcases p; simpa [Semiformula.extd] using h.iff e + +lemma sigma_extd_val (p : 𝚺-[m].Semiformula ξ n) : + p.extd.val = Semiformula.lMap Language.oringEmb p.val := by + rcases p; simp [extd] + +lemma pi_extd_val (p : 𝚷-[m].Semiformula ξ n) : + p.extd.val = Semiformula.lMap Language.oringEmb p.val := by + rcases p; simp [extd] + +lemma sigmaZero {Γ} (p : Γ-[0].Semiformula ξ k) : Hierarchy 𝚺 0 p.val := + match Γ with + | 𝚺 => p.sigma_prop + | 𝚷 => p.pi_prop.of_zero + | 𝚫 => by simp [val_sigma] + +def ofZero {Γ'} (p : Γ'-[0].Semiformula ξ k) : (Γ : HierarchySymbol) → Γ.Semiformula ξ k + | 𝚺-[_] => mkSigma p.val p.sigmaZero.of_zero + | 𝚷-[_] => mkPi p.val p.sigmaZero.of_zero + | 𝚫-[_] => mkDelta (mkSigma p.val p.sigmaZero.of_zero) (mkPi p.val p.sigmaZero.of_zero) + +def ofDeltaOne (p : 𝚫₁.Semiformula ξ k) : (Γ : SigmaPiDelta) → (m : ℕ) → Γ-[m+1].Semiformula ξ k + | 𝚺, m => mkSigma p.sigma.val (p.sigma.sigma_prop.mono (by simp)) + | 𝚷, m => mkPi p.pi.val (p.pi.pi_prop.mono (by simp)) + | 𝚫, m => mkDelta (mkSigma p.sigma.val (p.sigma.sigma_prop.mono (by simp))) (mkPi p.pi.val (p.pi.pi_prop.mono (by simp))) + +@[simp] lemma ofZero_val {Γ'} (p : Γ'-[0].Semiformula ξ n) (Γ) : (ofZero p Γ).val = p.val := by + match Γ with + | 𝚺-[_] => simp [ofZero] + | 𝚷-[_] => simp [ofZero] + | 𝚫-[_] => simp [ofZero] + +@[simp] lemma ProperOn.of_zero (p : Γ'-[0].Semisentence k) (m) : (ofZero p 𝚫-[m]).ProperOn M := by + simp [ProperOn, ofZero] + +@[simp] lemma ProperWithParamOn.of_zero (p : Γ'-[0].Semiformula M k) (m) : (ofZero p 𝚫-[m]).ProperWithParamOn M := by + simp [ProperWithParamOn, ofZero] + +def verum : {Γ : HierarchySymbol} → Γ.Semiformula ξ n + | 𝚺-[m] => mkSigma ⊤ (by simp) + | 𝚷-[m] => mkPi ⊤ (by simp) + | 𝚫-[m] => mkDelta (mkSigma ⊤ (by simp)) (mkPi ⊤ (by simp)) + +def falsum : {Γ : HierarchySymbol} → Γ.Semiformula ξ n + | 𝚺-[m] => mkSigma ⊥ (by simp) + | 𝚷-[m] => mkPi ⊥ (by simp) + | 𝚫-[m] => mkDelta (mkSigma ⊥ (by simp)) (mkPi ⊥ (by simp)) + +def and : {Γ : HierarchySymbol} → Γ.Semiformula ξ n → Γ.Semiformula ξ n → Γ.Semiformula ξ n + | 𝚺-[m], p, q => mkSigma (p.val ⋏ q.val) (by simp) + | 𝚷-[m], p, q => mkPi (p.val ⋏ q.val) (by simp) + | 𝚫-[m], p, q => mkDelta (mkSigma (p.sigma.val ⋏ q.sigma.val) (by simp)) (mkPi (p.pi.val ⋏ q.pi.val) (by simp)) + +def or : {Γ : HierarchySymbol} → Γ.Semiformula ξ n → Γ.Semiformula ξ n → Γ.Semiformula ξ n + | 𝚺-[m], p, q => mkSigma (p.val ⋎ q.val) (by simp) + | 𝚷-[m], p, q => mkPi (p.val ⋎ q.val) (by simp) + | 𝚫-[m], p, q => mkDelta (mkSigma (p.sigma.val ⋎ q.sigma.val) (by simp)) (mkPi (p.pi.val ⋎ q.pi.val) (by simp)) + +def negSigma (p : 𝚺-[m].Semiformula ξ n) : 𝚷-[m].Semiformula ξ n := mkPi (~p.val) (by simp) + +def negPi (p : 𝚷-[m].Semiformula ξ n) : 𝚺-[m].Semiformula ξ n := mkSigma (~p.val) (by simp) + +def negDelta (p : 𝚫-[m].Semiformula ξ n) : 𝚫-[m].Semiformula ξ n := mkDelta (p.pi.negPi) (p.sigma.negSigma) + +def ball (t : Semiterm ℒₒᵣ ξ n) : {Γ : HierarchySymbol} → Γ.Semiformula ξ (n + 1) → Γ.Semiformula ξ n + | 𝚺-[m], p => mkSigma (∀[“#0 < !!(Rew.bShift t)”] p.val) (by simp) + | 𝚷-[m], p => mkPi (∀[“#0 < !!(Rew.bShift t)”] p.val) (by simp) + | 𝚫-[m], p => + mkDelta (mkSigma (∀[“#0 < !!(Rew.bShift t)”] p.sigma.val) (by simp)) (mkPi (∀[“#0 < !!(Rew.bShift t)”] p.pi.val) (by simp)) + +def bex (t : Semiterm ℒₒᵣ ξ n) : {Γ : HierarchySymbol} → Γ.Semiformula ξ (n + 1) → Γ.Semiformula ξ n + | 𝚺-[m], p => mkSigma (∃[“#0 < !!(Rew.bShift t)”] p.val) (by simp) + | 𝚷-[m], p => mkPi (∃[“#0 < !!(Rew.bShift t)”] p.val) (by simp) + | 𝚫-[m], p => + mkDelta (mkSigma (∃[“#0 < !!(Rew.bShift t)”] p.sigma.val) (by simp)) (mkPi (∃[“#0 < !!(Rew.bShift t)”] p.pi.val) (by simp)) + +def all (p : 𝚷-[m + 1].Semiformula ξ (n + 1)) : 𝚷-[m + 1].Semiformula ξ n := mkPi (∀' p.val) p.pi_prop.all + +def ex (p : 𝚺-[m + 1].Semiformula ξ (n + 1)) : 𝚺-[m + 1].Semiformula ξ n := mkSigma (∃' p.val) p.sigma_prop.ex + +instance : Top (Γ.Semiformula ξ n) := ⟨verum⟩ + +instance : Bot (Γ.Semiformula ξ n) := ⟨falsum⟩ + +instance : Wedge (Γ.Semiformula ξ n) := ⟨and⟩ + +instance : Vee (Γ.Semiformula ξ n) := ⟨or⟩ + +instance : Tilde (𝚫-[m].Semiformula ξ n) := ⟨negDelta⟩ + +instance : LogicalConnective (𝚫-[m].Semiformula ξ n) where + arrow p q := ~p ⋎ q + +instance : ExQuantifier (𝚺-[m + 1].Semiformula ξ) := ⟨ex⟩ + +instance : UnivQuantifier (𝚷-[m + 1].Semiformula ξ) := ⟨all⟩ + +def substSigma (p : 𝚺-[m + 1].Semiformula ξ 1) (F : 𝚺-[m + 1].Semiformula ξ (n + 1)) : + 𝚺-[m + 1].Semiformula ξ n := (F ⋏ p.rew (Rew.substs ![#0])).ex + +@[simp] lemma val_verum : (⊤ : Γ.Semiformula ξ n).val = ⊤ := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val] + +@[simp] lemma sigma_verum {m} : (⊤ : 𝚫-[m].Semiformula ξ n).sigma = ⊤ := by simp [Top.top, verum] + +@[simp] lemma pi_verum {m} : (⊤ : 𝚫-[m].Semiformula ξ n).pi = ⊤ := by simp [Top.top, verum] + +@[simp] lemma val_falsum : (⊥ : Γ.Semiformula ξ n).val = ⊥ := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val] + +@[simp] lemma sigma_falsum {m} : (⊥ : 𝚫-[m].Semiformula ξ n).sigma = ⊥ := by simp [Bot.bot, falsum] + +@[simp] lemma pi_falsum {m} : (⊥ : 𝚫-[m].Semiformula ξ n).pi = ⊥ := by simp [Bot.bot, falsum] + +@[simp] lemma val_and (p q : Γ.Semiformula ξ n) : (p ⋏ q).val = p.val ⋏ q.val := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] + +@[simp] lemma sigma_and (p q : 𝚫-[m].Semiformula ξ n) : (p ⋏ q).sigma = p.sigma ⋏ q.sigma := by simp [Wedge.wedge, and] + +@[simp] lemma pi_and (p q : 𝚫-[m].Semiformula ξ n) : (p ⋏ q).pi = p.pi ⋏ q.pi := by simp [Wedge.wedge, and] + +@[simp] lemma val_or (p q : Γ.Semiformula ξ n) : (p ⋎ q).val = p.val ⋎ q.val := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] + +@[simp] lemma sigma_or (p q : 𝚫-[m].Semiformula ξ n) : (p ⋎ q).sigma = p.sigma ⋎ q.sigma := by simp [Vee.vee, or] + +@[simp] lemma pi_or (p q : 𝚫-[m].Semiformula ξ n) : (p ⋎ q).pi = p.pi ⋎ q.pi := by simp [Vee.vee, or] + +@[simp] lemma val_negSigma {m} (p : 𝚺-[m].Semiformula ξ n) : p.negSigma.val = ~p.val := by simp [val, val_sigma] + +@[simp] lemma val_negPi {m} (p : 𝚷-[m].Semiformula ξ n) : p.negPi.val = ~p.val := by simp [val, val_sigma] + +lemma val_negDelta {m} (p : 𝚫-[m].Semiformula ξ n) : (~p).val = ~p.pi.val := by simp [Tilde.tilde, negDelta] + +@[simp] lemma sigma_negDelta {m} (p : 𝚫-[m].Semiformula ξ n) : (~p).sigma = p.pi.negPi := by simp [Tilde.tilde, negDelta] + +@[simp] lemma sigma_negPi {m} (p : 𝚫-[m].Semiformula ξ n) : (~p).pi = p.sigma.negSigma := by simp [Tilde.tilde, negDelta] + +@[simp] lemma val_ball (t : Semiterm ℒₒᵣ ξ n) (p : Γ.Semiformula ξ (n + 1)) : (ball t p).val = ∀[“#0 < !!(Rew.bShift t)”] p.val := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] + +@[simp] lemma val_bex (t : Semiterm ℒₒᵣ ξ n) (p : Γ.Semiformula ξ (n + 1)) : (bex t p).val = ∃[“#0 < !!(Rew.bShift t)”] p.val := by + rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] + +@[simp] lemma val_exSigma {m} (p : 𝚺-[m + 1].Semiformula ξ (n + 1)) : (ex p).val = ∃' p.val := rfl + +@[simp] lemma val_allPi {m} (p : 𝚷-[m + 1].Semiformula ξ (n + 1)) : (all p).val = ∀' p.val := rfl + +@[simp] lemma ProperOn.verum : (⊤ : 𝚫-[m].Semisentence k).ProperOn M := by intro e; simp + +@[simp] lemma ProperOn.falsum : (⊥ : 𝚫-[m].Semisentence k).ProperOn M := by intro e; simp + +lemma ProperOn.and {p q : 𝚫-[m].Semisentence k} (hp : p.ProperOn M) (hq : q.ProperOn M) : (p ⋏ q).ProperOn M := by + intro e; simp [hp.iff, hq.iff] + +lemma ProperOn.or {p q : 𝚫-[m].Semisentence k} (hp : p.ProperOn M) (hq : q.ProperOn M) : (p ⋎ q).ProperOn M := by + intro e; simp [hp.iff, hq.iff] + +lemma ProperOn.neg {p : 𝚫-[m].Semisentence k} (hp : p.ProperOn M) : (~p).ProperOn M := by + intro e; simp [hp.iff] + +lemma ProperOn.eval_neg {p : 𝚫-[m].Semisentence k} (hp : p.ProperOn M) (e) : + Semiformula.Evalbm M e (~p).val ↔ ¬Semiformula.Evalbm M e p.val := by + simp [val, ←val_sigma, hp.iff] + +lemma ProperOn.ball {t} {p : 𝚫-[m + 1].Semisentence (k + 1)} (hp : p.ProperOn M) : (ball t p).ProperOn M := by + intro e; simp [Semiformula.ball, hp.iff] + +lemma ProperOn.bex {t} {p : 𝚫-[m + 1].Semisentence (k + 1)} (hp : p.ProperOn M) : (bex t p).ProperOn M := by + intro e; simp [Semiformula.bex, hp.iff] + +@[simp] lemma ProperWithParamOn.verum : (⊤ : 𝚫-[m].Semiformula M k).ProperWithParamOn M := by intro e; simp + +@[simp] lemma ProperWithParamOn.falsum : (⊥ : 𝚫-[m].Semiformula M k).ProperWithParamOn M := by intro e; simp + +lemma ProperWithParamOn.and {p q : 𝚫-[m].Semiformula M k} + (hp : p.ProperWithParamOn M) (hq : q.ProperWithParamOn M) : (p ⋏ q).ProperWithParamOn M := by + intro e; simp [hp.iff, hq.iff] + +lemma ProperWithParamOn.or {p q : 𝚫-[m].Semiformula M k} + (hp : p.ProperWithParamOn M) (hq : q.ProperWithParamOn M) : (p ⋎ q).ProperWithParamOn M := by + intro e; simp [hp.iff, hq.iff] + +lemma ProperWithParamOn.neg {p : 𝚫-[m].Semiformula M k} (hp : p.ProperWithParamOn M) : (~p).ProperWithParamOn M := by + intro e; simp [hp.iff] + +lemma ProperWithParamOn.eval_neg {p : 𝚫-[m].Semiformula M k} (hp : p.ProperWithParamOn M) (e) : + Semiformula.Evalm M e id (~p).val ↔ ¬Semiformula.Evalm M e id p.val := by + simp [val, ←val_sigma, hp.iff] + +lemma ProperWithParamOn.ball {t} {p : 𝚫-[m].Semiformula M (k + 1)} + (hp : p.ProperWithParamOn M) : (ball t p).ProperWithParamOn M := by + intro e; simp [Semiformula.ball, hp.iff] + +lemma ProperWithParamOn.bex {t} {p : 𝚫-[m].Semiformula M (k + 1)} + (hp : p.ProperWithParamOn M) : (bex t p).ProperWithParamOn M := by + intro e; simp [Semiformula.bex, hp.iff] + +def graphDelta (p : 𝚺-[m].Semiformula ξ (k + 1)) : 𝚫-[m].Semiformula ξ (k + 1) := + match m with + | 0 => p.ofZero _ + | m + 1 => mkDelta p (mkPi “x | ∀ y, !p.val y ⋯ → y = x” (by simp)) + +@[simp] lemma graphDelta_val (p : 𝚺-[m].Semiformula ξ (k + 1)) : p.graphDelta.val = p.val := by cases m <;> simp [graphDelta] + +end Semiformula + +end HierarchySymbol + +end LO.FirstOrder.Arith diff --git a/Arithmetization/Definability/Hierarchy.lean b/Arithmetization/Definability/Hierarchy.lean deleted file mode 100644 index 91f992d..0000000 --- a/Arithmetization/Definability/Hierarchy.lean +++ /dev/null @@ -1,1390 +0,0 @@ -import Arithmetization.Vorspiel.Lemmata -import Arithmetization.Definability.Init -import Arithmetization.Vorspiel.Graph -import Logic.FirstOrder.Arith.StrictHierarchy -import Aesop - -namespace LO.FirstOrder.Arith - -abbrev HierarchySymbol := SigmaPiDelta × ℕ - -abbrev HierarchySymbol.sigmaZero : HierarchySymbol := (𝚺, 0) - -abbrev HierarchySymbol.piZero : HierarchySymbol := (𝚷, 0) - -abbrev HierarchySymbol.deltaZero : HierarchySymbol := (𝚫, 0) - -abbrev HierarchySymbol.sigmaOne : HierarchySymbol := (𝚺, 1) - -abbrev HierarchySymbol.piOne : HierarchySymbol := (𝚷, 1) - -abbrev HierarchySymbol.deltaOne : HierarchySymbol := (𝚫, 1) - -notation "𝚺₀" => HierarchySymbol.sigmaZero - -notation "𝚷₀" => HierarchySymbol.piZero - -notation "𝚫₀" => HierarchySymbol.deltaZero - -notation "𝚺₁" => HierarchySymbol.sigmaOne - -notation "𝚷₁" => HierarchySymbol.piOne - -notation "𝚫₁" => HierarchySymbol.deltaOne - -namespace HierarchySymbol -/- -inductive Rel : HierarchySymbol → HierarchySymbol → Prop where - | delta_le_sigma (m) : Rel (𝚫, m) (𝚺, m) - | delta_le_pi (m) : Rel (𝚫, m) (𝚷, m) - | sigma_le_delta_succ (m) : Rel (𝚺, m) (𝚫, m + 1) - | pi_le_delta_succ (m) : Rel (𝚷, m) (𝚫, m + 1) - | sigma_le_delta_zero : Rel (𝚺, 0) (𝚫, 0) - | pi_le_delta_zero : Rel (𝚷, 0) (𝚫, 0) - -/-- Order structure of arithmetical hierarchy -/ -protected inductive LE : HierarchySymbol → HierarchySymbol → Prop where - | of_rel {Γ₁ Γ₂} : Rel Γ₁ Γ₂ → HierarchySymbol.LE Γ₁ Γ₂ - | refl (Γ) : HierarchySymbol.LE Γ Γ - | trans {Γ₁ Γ₂ Γ₃} : HierarchySymbol.LE Γ₁ Γ₂ → HierarchySymbol.LE Γ₂ Γ₃ → HierarchySymbol.LE Γ₁ Γ₃ - -instance : LE HierarchySymbol := ⟨HierarchySymbol.LE⟩ - -instance : Preorder HierarchySymbol where - le_refl := HierarchySymbol.LE.refl - le_trans := fun _ _ _ ↦ HierarchySymbol.LE.trans - -@[simp] lemma delta_le : (Γ : SigmaPiDelta) → (m : ℕ) → (𝚫, m) ≤ (Γ, m) - | 𝚺, m => HierarchySymbol.LE.of_rel (Rel.delta_le_sigma m) - | 𝚷, m => HierarchySymbol.LE.of_rel (Rel.delta_le_pi m) - | 𝚫, m => by rfl - -@[simp] lemma le_delta_succ : (Γ : SigmaPiDelta) → (m : ℕ) → (Γ, m) ≤ (𝚫, m + 1) - | 𝚺, m => HierarchySymbol.LE.of_rel (Rel.sigma_le_delta_succ m) - | 𝚷, m => HierarchySymbol.LE.of_rel (Rel.pi_le_delta_succ m) - | 𝚫, m => le_trans (delta_le 𝚺 m) (HierarchySymbol.LE.of_rel (Rel.sigma_le_delta_succ m)) - -@[simp] lemma le_succ (Γ₁ Γ₂ : SigmaPiDelta) (m : ℕ) : (Γ₁, m) ≤ (Γ₂, m + 1) := - le_trans (le_delta_succ Γ₁ m) (delta_le Γ₂ (m + 1)) - -lemma le_of_le (Γ : SigmaPiDelta) {m n : ℕ} (h : m ≤ n) : (Γ, m) ≤ (Γ, n) := by - have : n = m + (n - m) := (Nat.add_sub_of_le h).symm - generalize e : n - m = d - rw [e] at this; rcases this; clear e - induction' d with d IH - · rfl - · exact le_trans (IH <| by simp) (by simp [Nat.add_succ]) - -lemma le_of_lt (Γ₁ Γ₂ : SigmaPiDelta) {m n : ℕ} (h : m < n) : (Γ₁, m) ≤ (Γ₂, n) := by - cases' n with n - · simp_all - · exact le_trans (le_of_le Γ₁ (by simpa [Nat.lt_succ] using h)) (le_succ Γ₁ Γ₂ n) - -@[simp] lemma zero_le (Γ₁ Γ₂ : SigmaPiDelta) : (Γ₁, 0) ≤ (Γ₂, 0) := - match Γ₁, Γ₂ with - | 𝚺, 𝚺 => by rfl - | 𝚺, 𝚷 => le_trans (HierarchySymbol.LE.of_rel Rel.sigma_le_delta_zero) (by simp) - | 𝚺, 𝚫 => HierarchySymbol.LE.of_rel Rel.sigma_le_delta_zero - | 𝚷, 𝚺 => le_trans (HierarchySymbol.LE.of_rel Rel.pi_le_delta_zero) (by simp) - | 𝚷, 𝚷 => by rfl - | 𝚷, 𝚫 => HierarchySymbol.LE.of_rel Rel.pi_le_delta_zero - | 𝚫, 𝚺 => by simp - | 𝚫, 𝚷 => by simp - | 𝚫, 𝚫 => by rfl --/ - -end HierarchySymbol - -end Arith - -def Defined {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semisentence L k) : Prop := - ∀ v, R v ↔ Semiformula.Evalbm M v p - -def DefinedWithParam {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semiformula L M k) : Prop := - ∀ v, R v ↔ Semiformula.Evalm M v id p - -lemma Defined.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : - Semiformula.Evalbm M v p ↔ R v := (h v).symm - -lemma DefinedWithParam.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semiformula L M k} (h : DefinedWithParam R p) (v) : - Semiformula.Evalm M v id p ↔ R v := (h v).symm - -namespace Arith - -variable (L : Language.{u}) [L.ORing] (ξ : Type v) (n : ℕ) - -inductive HSemiformula : HierarchySymbol → Type _ where - | mkSigma {m} : (p : Semiformula L ξ n) → Hierarchy 𝚺 m p → HSemiformula (𝚺, m) - | mkPi {m} : (p : Semiformula L ξ n) → Hierarchy 𝚷 m p → HSemiformula (𝚷, m) - | mkDelta {m} : HSemiformula (𝚺, m) → HSemiformula (𝚷, m) → HSemiformula (𝚫, m) - -abbrev HSemisentence (Γ : HierarchySymbol) := HSemiformula L Empty n Γ - -scoped[LO.FirstOrder.Arith] notation Γ "-Semisentence " n => HSemisentence ℒₒᵣ n Γ - -variable {L ξ n} - -namespace HSemiformula - -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [Structure L M] [Structure.ORing L M] - -def val : HSemiformula L ξ n Γ → Semiformula L ξ n - | mkSigma p _ => p - | mkPi p _ => p - | mkDelta p _ => p.val - -@[simp] lemma val_mkSigma (p : Semiformula L ξ n) (hp : Hierarchy 𝚺 m p) : (mkSigma p hp).val = p := rfl - -@[simp] lemma val_mkPi (p : Semiformula L ξ n) (hp : Hierarchy 𝚷 m p) : (mkPi p hp).val = p := rfl - -@[simp] lemma val_mkDelta (p : HSemiformula L ξ n (𝚺, m)) (q : HSemiformula L ξ n (𝚷, m)) : (mkDelta p q).val = p.val := rfl - -instance : Coe (HSemisentence L n 𝚺₀) (Semisentence L n) := ⟨HSemiformula.val⟩ -instance : Coe (HSemisentence L n 𝚷₀) (Semisentence L n) := ⟨HSemiformula.val⟩ -instance : Coe (HSemisentence L n 𝚫₀) (Semisentence L n) := ⟨HSemiformula.val⟩ - -instance : Coe (HSemisentence L n 𝚺₁) (Semisentence L n) := ⟨HSemiformula.val⟩ -instance : Coe (HSemisentence L n 𝚷₁) (Semisentence L n) := ⟨HSemiformula.val⟩ -instance : Coe (HSemisentence L n 𝚫₁) (Semisentence L n) := ⟨HSemiformula.val⟩ - -@[simp] lemma sigma_prop : (p : HSemiformula L ξ n (𝚺, m)) → Hierarchy 𝚺 m p.val - | mkSigma _ h => h - -@[simp] lemma pi_prop : (p : HSemiformula L ξ n (𝚷, m)) → Hierarchy 𝚷 m p.val - | mkPi _ h => h - -@[simp] lemma polarity_prop : {Γ : Polarity} → (p : HSemiformula L ξ n (Γ, m)) → Hierarchy Γ m p.val - | 𝚺, p => p.sigma_prop - | 𝚷, p => p.pi_prop - -def sigma : HSemiformula L ξ n (𝚫, m) → HSemiformula L ξ n (𝚺, m) - | mkDelta p _ => p - -@[simp] lemma sigma_mkDelta (p : HSemiformula L ξ n (𝚺, m)) (q : HSemiformula L ξ n (𝚷, m)) : (mkDelta p q).sigma = p := rfl - -def pi : HSemiformula L ξ n (𝚫, m) → HSemiformula L ξ n (𝚷, m) - | mkDelta _ p => p - -@[simp] lemma pi_mkDelta (p : HSemiformula L ξ n (𝚺, m)) (q : HSemiformula L ξ n (𝚷, m)) : (mkDelta p q).pi = q := rfl - -lemma val_sigma (p : HSemiformula L ξ n (𝚫, m)) : p.sigma.val = p.val := by rcases p; simp - -def mkPolarity (p : Semiformula L ξ n) : (Γ : Polarity) → Hierarchy Γ m p → HSemiformula L ξ n (Γ, m) - | 𝚺, h => mkSigma p h - | 𝚷, h => mkPi p h - -@[simp] lemma val_mkPolarity (p : Semiformula L ξ n) {Γ} (h : Hierarchy Γ m p) : (mkPolarity p Γ h).val = p := by cases Γ <;> rfl - -variable (M) - -def ProperOn (p : HSemisentence L n (𝚫, m)) : Prop := - ∀ (e : Fin n → M), Semiformula.Evalbm M e p.sigma.val ↔ Semiformula.Evalbm M e p.pi.val - -def ProperWithParamOn (p : HSemiformula L M n (𝚫, m)) : Prop := - ∀ (e : Fin n → M), Semiformula.Evalm M e id p.sigma.val ↔ Semiformula.Evalm M e id p.pi.val - -variable {M} - -lemma ProperOn.iff {p : HSemisentence L n (𝚫, m)} - (h : p.ProperOn M) (e : Fin n → M) : - Semiformula.Evalbm M e p.sigma.val ↔ Semiformula.Evalbm M e p.pi.val := h e - -lemma ProperWithParamOn.iff {p : HSemiformula L M n (𝚫, m)} - (h : p.ProperWithParamOn M) (e : Fin n → M) : - Semiformula.Evalm M e id p.sigma.val ↔ Semiformula.Evalm (L := L) M e id p.pi.val := h e - -lemma ProperOn.iff' {p : HSemisentence L n (𝚫, m)} - (h : p.ProperOn M) (e : Fin n → M) : - Semiformula.Evalbm M e p.pi.val ↔ Semiformula.Evalbm M e p.val := by simp [←h.iff, val_sigma] - -lemma ProperWithParamOn.iff' {p : HSemiformula L M n (𝚫, m)} - (h : p.ProperWithParamOn M) (e : Fin n → M) : - Semiformula.Evalm M e id p.pi.val ↔ Semiformula.Evalm (L := L) M e id p.val := by simp [←h.iff, val_sigma] - -def rew (ω : Rew L ξ₁ n₁ ξ₂ n₂) : {Γ : HierarchySymbol} → HSemiformula L ξ₁ n₁ Γ → HSemiformula L ξ₂ n₂ Γ - | (𝚺, _), mkSigma p hp => mkSigma (ω.hom p) (by simpa using hp) - | (𝚷, _), mkPi p hp => mkPi (ω.hom p) (by simpa using hp) - | (𝚫, _), mkDelta p q => mkDelta (p.rew ω) (q.rew ω) - -@[simp] lemma val_rew (ω : Rew L ξ₁ n₁ ξ₂ n₂) {Γ} (p : HSemiformula L ξ₁ n₁ Γ) : (p.rew ω).val = ω.hom p.val := by - rcases Γ with ⟨Γ, m⟩; rcases p with (_ | _ | ⟨⟨p, _⟩, ⟨q, _⟩⟩) <;> simp [rew] - -@[simp] lemma ProperOn.rew {p : HSemisentence L n₁ (𝚫, m)} (h : p.ProperOn M) (ω : Rew L Empty n₁ Empty n₂) : (p.rew ω).ProperOn M := by - rcases p; simp [ProperOn, HSemiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] - intro e; exact h.iff _ - -@[simp] lemma ProperOn.rew' {p : HSemisentence L n₁ (𝚫, m)} (h : p.ProperOn M) (ω : Rew L Empty n₁ M n₂) : (p.rew ω).ProperWithParamOn M := by - rcases p; intro e; simp [ProperOn, HSemiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] - simpa using h.iff _ - -@[simp] lemma ProperWithParamOn.rew {p : HSemiformula L M n₁ (𝚫, m)} - (h : p.ProperWithParamOn M) (f : Fin n₁ → Semiterm L M n₂) : (p.rew (Rew.substs f)).ProperWithParamOn M := by - rcases p; intro e; - simp [ProperOn, HSemiformula.rew, Semiformula.eval_rew, Function.comp, h.iff, Empty.eq_elim] - exact h.iff _ - -variable (L) - -def emb : {Γ : HierarchySymbol} → HSemiformula ℒₒᵣ ξ n Γ → HSemiformula L ξ n Γ - | (𝚺, _), mkSigma p hp => mkSigma (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) - | (𝚷, _), mkPi p hp => mkPi (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) - | (𝚫, _), mkDelta p q => mkDelta p.emb q.emb - -variable {L} - -@[simp] lemma val_emb {Γ} (p : HSemiformula ℒₒᵣ ξ n Γ) : (p.emb L).val = Semiformula.lMap Language.oringEmb p.val := by - rcases Γ with ⟨Γ, m⟩; rcases p with (_ | _ | ⟨⟨p, _⟩, ⟨q, _⟩⟩) <;> simp [rew, val] - -@[simp] lemma pi_emb (p : HSemiformula ℒₒᵣ ξ n (𝚫, m)) : (p.emb L).pi = p.pi.emb L := by cases p; rfl - -@[simp] lemma sigma_emb (p : HSemiformula ℒₒᵣ ξ n (𝚫, m)) : (p.emb L).sigma = p.sigma.emb L := by cases p; rfl - -@[simp] lemma emb_proper (p : HSemisentence ℒₒᵣ n (𝚫, m)) : (p.emb L).ProperOn M ↔ p.ProperOn M := by - rcases p; simp [ProperOn, emb] - -@[simp] lemma emb_properWithParam (p : HSemiformula ℒₒᵣ M n (𝚫, m)) : (p.emb L).ProperWithParamOn M ↔ p.ProperWithParamOn M := by - rcases p; simp [ProperWithParamOn, emb] - -variable (L) - -def extd : HSemiformula ℒₒᵣ ξ n Γ → HSemiformula L ξ n Γ - | mkSigma p hp => mkSigma (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) - | mkPi p hp => mkPi (Semiformula.lMap Language.oringEmb p) (Hierarchy.oringEmb hp) - | mkDelta p q => mkDelta p.extd q.extd - -variable {L} - -@[simp] -lemma eval_extd_iff {e ε} {p : HSemiformula ℒₒᵣ ξ n Γ} : - Semiformula.Evalm M e ε (p.extd L).val ↔ Semiformula.Evalm M e ε p.val := by - induction p <;> simp [extd, *] - -lemma ProperOn.extd {p : HSemisentence ℒₒᵣ n (𝚫, m)} (h : p.ProperOn M) : (p.extd L).ProperOn M := by - intro e; rcases p; simpa [HSemiformula.extd] using h.iff e - -lemma ProperWithParamOn.extd {p : HSemisentence ℒₒᵣ n (𝚫, m)} (h : p.ProperOn M) : (p.extd L).ProperOn M := by - intro e; rcases p; simpa [HSemiformula.extd] using h.iff e - -lemma sigma_extd_val (p : HSemiformula ℒₒᵣ ξ n (𝚺, m)) : - (p.extd L).val = Semiformula.lMap Language.oringEmb p.val := by - rcases p; simp [extd] - -lemma pi_extd_val (p : HSemiformula ℒₒᵣ ξ n (𝚷, m)) : - (p.extd L).val = Semiformula.lMap Language.oringEmb p.val := by - rcases p; simp [extd] - -/- -def ofRel : {Γ₁ Γ₂ : HierarchySymbol} → HierarchySymbol.Rel Γ₁ Γ₂ → HSemiformula L ξ k Γ₁ → HSemiformula L ξ k Γ₂ - | (𝚺, m), (𝚫, n + 1), H, (mkSigma p hp) => - have : n = m := by cases H; rfl - mkDelta (mkSigma p <| Hierarchy.strict_mono hp 𝚺 (by simp [this])) (mkPi p <| Hierarchy.strict_mono hp 𝚷 (by simp [this])) - | (𝚷, m), (𝚫, n + 1), H, (mkPi p hp) => - have : n = m := by cases H; rfl - mkDelta (mkSigma p <| Hierarchy.strict_mono hp 𝚺 (by simp [this])) (mkPi p <| Hierarchy.strict_mono hp 𝚷 (by simp [this])) - | (𝚫, m), (𝚺, n), H, (mkDelta (mkSigma p hp) _) => - have : n = m := by cases H; rfl - mkSigma p (by simpa [this] using hp) - | (𝚫, m), (𝚷, n), H, (mkDelta _ (mkPi p hp)) => - have : n = m := by cases H; rfl - mkPi p (by simpa [this] using hp) - | (𝚷, m), (𝚫, 0), H, (mkPi p hp) => - have : m = 0 := by cases H; rfl - mkDelta (mkSigma p <| Hierarchy.of_zero (by simpa [this] using hp)) (mkPi p <| by simpa [this] using hp) - | (𝚺, m), (𝚫, 0), H, (mkSigma p hp) => - have : m = 0 := by cases H; rfl - mkDelta (mkSigma p <| by simpa [this] using hp) (mkPi p <| Hierarchy.of_zero (by simpa [this] using hp)) --/ - -lemma sigmaZero (p : HSemiformula L ξ k (Γ, 0)) : Hierarchy 𝚺 0 p.val := - match Γ with - | 𝚺 => p.sigma_prop - | 𝚷 => p.pi_prop.of_zero - | 𝚫 => by simpa [val_sigma] using p.sigma.sigma_prop - -def ofZero (p : HSemiformula L ξ k (Γ', 0)) : (Γ : HierarchySymbol) → HSemiformula L ξ k Γ - | (𝚺, _) => mkSigma p.val p.sigmaZero.of_zero - | (𝚷, _) => mkPi p.val p.sigmaZero.of_zero - | (𝚫, _) => mkDelta (mkSigma p.val p.sigmaZero.of_zero) (mkPi p.val p.sigmaZero.of_zero) - -def ofDeltaOne (p : HSemiformula L ξ k 𝚫₁) : (Γ : SigmaPiDelta) → (m : ℕ) → HSemiformula L ξ k (Γ, m + 1) - | 𝚺, m => mkSigma p.sigma.val (p.sigma.sigma_prop.mono (by simp)) - | 𝚷, m => mkPi p.pi.val (p.pi.pi_prop.mono (by simp)) - | 𝚫, m => mkDelta (mkSigma p.sigma.val (p.sigma.sigma_prop.mono (by simp))) (mkPi p.pi.val (p.pi.pi_prop.mono (by simp))) - -@[simp] lemma ofZero_val (p : HSemiformula L ξ k (Γ', 0)) (Γ) : (ofZero p Γ).val = p.val := by - match Γ with - | (𝚺, _) => simp [ofZero] - | (𝚷, _) => simp [ofZero] - | (𝚫, _) => simp [ofZero] - -@[simp] lemma ProperOn.of_zero (p : HSemisentence L k (Γ', 0)) (m) : (ofZero p (𝚫, m)).ProperOn M := by - simp [ProperOn, ofZero] - -@[simp] lemma ProperWithParamOn.of_zero (p : HSemiformula L M k (Γ', 0)) (m) : (ofZero p (𝚫, m)).ProperWithParamOn M := by - simp [ProperWithParamOn, ofZero] - -def verum : {Γ : HierarchySymbol} → HSemiformula L ξ n Γ - | (𝚺, m) => mkSigma ⊤ (by simp) - | (𝚷, m) => mkPi ⊤ (by simp) - | (𝚫, m) => mkDelta (mkSigma ⊤ (by simp)) (mkPi ⊤ (by simp)) - -def falsum : {Γ : HierarchySymbol} → HSemiformula L ξ n Γ - | (𝚺, m) => mkSigma ⊥ (by simp) - | (𝚷, m) => mkPi ⊥ (by simp) - | (𝚫, m) => mkDelta (mkSigma ⊥ (by simp)) (mkPi ⊥ (by simp)) - -def and : {Γ : HierarchySymbol} → HSemiformula L ξ n Γ → HSemiformula L ξ n Γ → HSemiformula L ξ n Γ - | (𝚺, m), p, q => mkSigma (p.val ⋏ q.val) (by simp) - | (𝚷, m), p, q => mkPi (p.val ⋏ q.val) (by simp) - | (𝚫, m), p, q => mkDelta (mkSigma (p.sigma.val ⋏ q.sigma.val) (by simp)) (mkPi (p.pi.val ⋏ q.pi.val) (by simp)) - -def or : {Γ : HierarchySymbol} → HSemiformula L ξ n Γ → HSemiformula L ξ n Γ → HSemiformula L ξ n Γ - | (𝚺, m), p, q => mkSigma (p.val ⋎ q.val) (by simp) - | (𝚷, m), p, q => mkPi (p.val ⋎ q.val) (by simp) - | (𝚫, m), p, q => mkDelta (mkSigma (p.sigma.val ⋎ q.sigma.val) (by simp)) (mkPi (p.pi.val ⋎ q.pi.val) (by simp)) - -def negSigma (p : HSemiformula L ξ n (𝚺, m)) : HSemiformula L ξ n (𝚷, m) := mkPi (~p.val) (by simp) - -def negPi (p : HSemiformula L ξ n (𝚷, m)) : HSemiformula L ξ n (𝚺, m) := mkSigma (~p.val) (by simp) - -def negDelta (p : HSemiformula L ξ n (𝚫, m)) : HSemiformula L ξ n (𝚫, m) := mkDelta (p.pi.negPi) (p.sigma.negSigma) - -def ball (t : Semiterm L ξ n) : {Γ : HierarchySymbol} → HSemiformula L ξ (n + 1) Γ → HSemiformula L ξ n Γ - | (𝚺, m), p => mkSigma (∀[“#0 < !!(Rew.bShift t)”] p.val) (by simp) - | (𝚷, m), p => mkPi (∀[“#0 < !!(Rew.bShift t)”] p.val) (by simp) - | (𝚫, m), p => - mkDelta (mkSigma (∀[“#0 < !!(Rew.bShift t)”] p.sigma.val) (by simp)) (mkPi (∀[“#0 < !!(Rew.bShift t)”] p.pi.val) (by simp)) - -def bex (t : Semiterm L ξ n) : {Γ : HierarchySymbol} → HSemiformula L ξ (n + 1) Γ → HSemiformula L ξ n Γ - | (𝚺, m), p => mkSigma (∃[“#0 < !!(Rew.bShift t)”] p.val) (by simp) - | (𝚷, m), p => mkPi (∃[“#0 < !!(Rew.bShift t)”] p.val) (by simp) - | (𝚫, m), p => - mkDelta (mkSigma (∃[“#0 < !!(Rew.bShift t)”] p.sigma.val) (by simp)) (mkPi (∃[“#0 < !!(Rew.bShift t)”] p.pi.val) (by simp)) - -def all (p : HSemiformula L ξ (n + 1) (𝚷, m + 1)) : HSemiformula L ξ n (𝚷, m + 1) := mkPi (∀' p.val) p.pi_prop.all - -def ex (p : HSemiformula L ξ (n + 1) (𝚺, m + 1)) : HSemiformula L ξ n (𝚺, m + 1) := mkSigma (∃' p.val) p.sigma_prop.ex - -instance : Top (HSemiformula L ξ n Γ) := ⟨verum⟩ - -instance : Bot (HSemiformula L ξ n Γ) := ⟨falsum⟩ - -instance : Wedge (HSemiformula L ξ n Γ) := ⟨and⟩ - -instance : Vee (HSemiformula L ξ n Γ) := ⟨or⟩ - -instance : Tilde (HSemiformula L ξ n (𝚫, m)) := ⟨negDelta⟩ - -instance : LogicalConnective (HSemiformula L ξ n (𝚫, m)) where - arrow p q := ~p ⋎ q - -instance : ExQuantifier (HSemiformula L ξ · (𝚺, m + 1)) := ⟨ex⟩ - -instance : UnivQuantifier (HSemiformula L ξ · (𝚷, m + 1)) := ⟨all⟩ - -def substSigma (p : HSemiformula L ξ 1 (𝚺, m + 1)) (F : HSemiformula L ξ (n + 1) (𝚺, m + 1)) : - HSemiformula L ξ n (𝚺, m + 1) := (F ⋏ p.rew (Rew.substs ![#0])).ex - -@[simp] lemma val_verum {Γ}: (⊤ : HSemiformula L ξ n Γ).val = ⊤ := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val] - -@[simp] lemma sigma_verum {m} : (⊤ : HSemiformula L ξ n (𝚫, m)).sigma = ⊤ := by simp [Top.top, verum] - -@[simp] lemma pi_verum {m} : (⊤ : HSemiformula L ξ n (𝚫, m)).pi = ⊤ := by simp [Top.top, verum] - -@[simp] lemma val_falsum {Γ}: (⊥ : HSemiformula L ξ n Γ).val = ⊥ := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val] - -@[simp] lemma sigma_falsum {m} : (⊥ : HSemiformula L ξ n (𝚫, m)).sigma = ⊥ := by simp [Bot.bot, falsum] - -@[simp] lemma pi_falsum {m} : (⊥ : HSemiformula L ξ n (𝚫, m)).pi = ⊥ := by simp [Bot.bot, falsum] - -@[simp] lemma val_and {Γ} (p q : HSemiformula L ξ n Γ) : (p ⋏ q).val = p.val ⋏ q.val := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] - -@[simp] lemma sigma_and (p q : HSemiformula L ξ n (𝚫, m)) : (p ⋏ q).sigma = p.sigma ⋏ q.sigma := by simp [Wedge.wedge, and] - -@[simp] lemma pi_and (p q : HSemiformula L ξ n (𝚫, m)) : (p ⋏ q).pi = p.pi ⋏ q.pi := by simp [Wedge.wedge, and] - -@[simp] lemma val_or {Γ} (p q : HSemiformula L ξ n Γ) : (p ⋎ q).val = p.val ⋎ q.val := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] - -@[simp] lemma sigma_or (p q : HSemiformula L ξ n (𝚫, m)) : (p ⋎ q).sigma = p.sigma ⋎ q.sigma := by simp [Vee.vee, or] - -@[simp] lemma pi_or (p q : HSemiformula L ξ n (𝚫, m)) : (p ⋎ q).pi = p.pi ⋎ q.pi := by simp [Vee.vee, or] - -@[simp] lemma val_negSigma {m} (p : HSemiformula L ξ n (𝚺, m)) : p.negSigma.val = ~p.val := by simp [val, val_sigma] - -@[simp] lemma val_negPi {m} (p : HSemiformula L ξ n (𝚷, m)) : p.negPi.val = ~p.val := by simp [val, val_sigma] - -lemma val_negDelta {m} (p : HSemiformula L ξ n (𝚫, m)) : (~p).val = ~p.pi.val := by simp [Tilde.tilde, negDelta] - -@[simp] lemma sigma_negDelta {m} (p : HSemiformula L ξ n (𝚫, m)) : (~p).sigma = p.pi.negPi := by simp [Tilde.tilde, negDelta] - -@[simp] lemma sigma_negPi {m} (p : HSemiformula L ξ n (𝚫, m)) : (~p).pi = p.sigma.negSigma := by simp [Tilde.tilde, negDelta] - -@[simp] lemma val_ball {Γ} (t : Semiterm L ξ n) (p : HSemiformula L ξ (n + 1) Γ) : (ball t p).val = ∀[“#0 < !!(Rew.bShift t)”] p.val := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] - -@[simp] lemma val_bex {Γ} (t : Semiterm L ξ n) (p : HSemiformula L ξ (n + 1) Γ) : (bex t p).val = ∃[“#0 < !!(Rew.bShift t)”] p.val := by - rcases Γ with ⟨Γ, m⟩; rcases Γ <;> simp [val, val_sigma] - -@[simp] lemma val_exSigma {m} (p : HSemiformula L ξ (n + 1) (𝚺, (m + 1))) : (ex p).val = ∃' p.val := rfl - -@[simp] lemma val_allPi {m} (p : HSemiformula L ξ (n + 1) (𝚷, (m + 1))) : (all p).val = ∀' p.val := rfl - -@[simp] lemma ProperOn.verum : (⊤ : HSemisentence L k (𝚫, m)).ProperOn M := by intro e; simp - -@[simp] lemma ProperOn.falsum : (⊥ : HSemisentence L k (𝚫, m)).ProperOn M := by intro e; simp - -lemma ProperOn.and {p q : HSemisentence L k (𝚫, m)} (hp : p.ProperOn M) (hq : q.ProperOn M) : (p ⋏ q).ProperOn M := by - intro e; simp [hp.iff, hq.iff] - -lemma ProperOn.or {p q : HSemisentence L k (𝚫, m)} (hp : p.ProperOn M) (hq : q.ProperOn M) : (p ⋎ q).ProperOn M := by - intro e; simp [hp.iff, hq.iff] - -lemma ProperOn.neg {p : HSemisentence L k (𝚫, m)} (hp : p.ProperOn M) : (~p).ProperOn M := by - intro e; simp [hp.iff] - -lemma ProperOn.eval_neg {p : HSemisentence L k (𝚫, m)} (hp : p.ProperOn M) (e) : - Semiformula.Evalbm M e (~p).val ↔ ¬Semiformula.Evalbm M e p.val := by - simp [val, ←val_sigma, hp.iff] - -lemma ProperOn.ball {t} {p : HSemisentence L (k + 1) (𝚫, m)} (hp : p.ProperOn M) : (ball t p).ProperOn M := by - intro e; simp [HSemiformula.ball, hp.iff] - -lemma ProperOn.bex {t} {p : HSemisentence L (k + 1) (𝚫, m)} (hp : p.ProperOn M) : (bex t p).ProperOn M := by - intro e; simp [HSemiformula.bex, hp.iff] - -@[simp] lemma ProperWithParamOn.verum : (⊤ : HSemiformula L M k (𝚫, m)).ProperWithParamOn M := by intro e; simp - -@[simp] lemma ProperWithParamOn.falsum : (⊥ : HSemiformula L M k (𝚫, m)).ProperWithParamOn M := by intro e; simp - -lemma ProperWithParamOn.and {p q : HSemiformula L M k (𝚫, m)} - (hp : p.ProperWithParamOn M) (hq : q.ProperWithParamOn M) : (p ⋏ q).ProperWithParamOn M := by - intro e; simp [hp.iff, hq.iff] - -lemma ProperWithParamOn.or {p q : HSemiformula L M k (𝚫, m)} - (hp : p.ProperWithParamOn M) (hq : q.ProperWithParamOn M) : (p ⋎ q).ProperWithParamOn M := by - intro e; simp [hp.iff, hq.iff] - -lemma ProperWithParamOn.neg {p : HSemiformula L M k (𝚫, m)} (hp : p.ProperWithParamOn M) : (~p).ProperWithParamOn M := by - intro e; simp [hp.iff] - -lemma ProperWithParamOn.eval_neg {p : HSemiformula L M k (𝚫, m)} (hp : p.ProperWithParamOn M) (e) : - Semiformula.Evalm M e id (~p).val ↔ ¬Semiformula.Evalm M e id p.val := by - simp [val, ←val_sigma, hp.iff] - -lemma ProperWithParamOn.ball {t} {p : HSemiformula L M (k + 1) (𝚫, m)} - (hp : p.ProperWithParamOn M) : (ball t p).ProperWithParamOn M := by - intro e; simp [HSemiformula.ball, hp.iff] - -lemma ProperWithParamOn.bex {t} {p : HSemiformula L M (k + 1) (𝚫, m)} - (hp : p.ProperWithParamOn M) : (bex t p).ProperWithParamOn M := by - intro e; simp [HSemiformula.bex, hp.iff] - -def graphDelta (p : HSemiformula L ξ (k + 1) (𝚺, m)) : HSemiformula L ξ (k + 1) (𝚫, m) := - match m with - | 0 => p.ofZero _ - | m + 1 => mkDelta p (mkPi “x | ∀ y, !p.val y ⋯ → y = x” (by simp)) - -@[simp] lemma graphDelta_val (p : HSemiformula L ξ (k + 1) (𝚺, m)) : p.graphDelta.val = p.val := by cases m <;> simp [graphDelta] - -end HSemiformula - -namespace Definability - -namespace HSemiformula - -variable (ξ : Type*) (n) (Γ : SigmaPiDelta) (m : ℕ) - -@[simp] lemma hierarchy_sigma (p : HSemiformula L ξ n (𝚺, m)) : Hierarchy 𝚺 m p.val := p.sigma_prop - -@[simp] lemma hierarchy_pi (p : HSemiformula L ξ n (𝚷, m)) : Hierarchy 𝚷 m p.val := p.pi_prop - -@[simp] lemma hierarchy_zero {Γ Γ' m} (p : HSemiformula L ξ n (Γ, 0)) : Hierarchy Γ' m p.val := by - cases Γ - · exact Hierarchy.of_zero p.sigma_prop - · exact Hierarchy.of_zero p.pi_prop - · cases p - simp; exact Hierarchy.of_zero (HSemiformula.sigma_prop _) - -end HSemiformula - -end Definability - -open LO.Arith - -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] [Structure L M] [Structure.ORing L M] [Structure.Monotone L M] - -variable {Γ : HierarchySymbol} - -open Definability - -def Defined (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → HSemisentence L k Γ → Prop - | (𝚺, _), p => FirstOrder.Defined R p.val - | (𝚷, _), p => FirstOrder.Defined R p.val - | (𝚫, _), p => p.ProperOn M ∧ FirstOrder.Defined R p.val - -def DefinedWithParam (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → HSemiformula L M k Γ → Prop - | (𝚺, _), p => FirstOrder.DefinedWithParam R p.val - | (𝚷, _), p => FirstOrder.DefinedWithParam R p.val - | (𝚫, _), p => p.ProperWithParamOn M ∧ FirstOrder.DefinedWithParam R p.val - -variable (L Γ) - -class Definable {k} (P : (Fin k → M) → Prop) : Prop where - definable : ∃ p : HSemiformula L M k Γ, DefinedWithParam P p - -abbrev DefinedPred (P : M → Prop) (p : HSemisentence L 1 Γ) : Prop := - Defined (λ v ↦ P (v 0)) p - -abbrev DefinedRel (R : M → M → Prop) (p : HSemisentence L 2 Γ) : Prop := - Defined (λ v ↦ R (v 0) (v 1)) p - -abbrev DefinedRel₃ (R : M → M → M → Prop) (p : HSemisentence L 3 Γ) : Prop := - Defined (λ v ↦ R (v 0) (v 1) (v 2)) p - -abbrev DefinedRel₄ (R : M → M → M → M → Prop) (p : HSemisentence L 4 Γ) : Prop := - Defined (λ v ↦ R (v 0) (v 1) (v 2) (v 3)) p - -variable {L Γ} - -abbrev DefinedFunction {k} (f : (Fin k → M) → M) (p : HSemisentence L (k + 1) Γ) : Prop := - Defined (fun v => v 0 = f (v ·.succ)) p - -variable (L Γ) - -abbrev DefinedFunction₁ (f : M → M) (p : HSemisentence L 2 Γ) : Prop := - DefinedFunction (fun v => f (v 0)) p - -abbrev DefinedFunction₂ (f : M → M → M) (p : HSemisentence L 3 Γ) : Prop := - DefinedFunction (fun v => f (v 0) (v 1)) p - -abbrev DefinedFunction₃ (f : M → M → M → M) (p : HSemisentence L 4 Γ) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2)) p - -abbrev DefinedFunction₄ (f : M → M → M → M → M) (p : HSemisentence L 5 Γ) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3)) p - -abbrev DefinedFunction₅ (f : M → M → M → M → M → M) (p : HSemisentence L 6 Γ) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3) (v 4)) p - -abbrev DefinablePred (P : M → Prop) : Prop := Definable L Γ (k := 1) (fun v ↦ P (v 0)) - -abbrev DefinableRel (P : M → M → Prop) : Prop := Definable L Γ (k := 2) (fun v ↦ P (v 0) (v 1)) - -abbrev DefinableRel₃ (P : M → M → M → Prop) : Prop := Definable L Γ (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) - -abbrev DefinableRel₄ (P : M → M → M → M → Prop) : Prop := Definable L Γ (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) - -abbrev DefinableRel₅ (P : M → M → M → M → M → Prop) : Prop := Definable L Γ (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) - -abbrev DefinableRel₆ (P : M → M → M → M → M → M → Prop) : Prop := Definable L Γ (k := 6) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) - -abbrev DefinableFunction (f : (Fin k → M) → M) : Prop := Definable L Γ (k := k + 1) (fun v ↦ v 0 = f (v ·.succ)) - -abbrev DefinableFunction₁ (f : M → M) : Prop := DefinableFunction L Γ (k := 1) (fun v ↦ f (v 0)) - -abbrev DefinableFunction₂ (f : M → M → M) : Prop := DefinableFunction L Γ (k := 2) (fun v ↦ f (v 0) (v 1)) - -abbrev DefinableFunction₃ (f : M → M → M → M) : Prop := DefinableFunction L Γ (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) - -abbrev DefinableFunction₄ (f : M → M → M → M → M) : Prop := DefinableFunction L Γ (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) - -abbrev DefinableFunction₅ (f : M → M → M → M → M → M) : Prop := DefinableFunction L Γ (k := 5) (fun v ↦ f (v 0) (v 1) (v 2) (v 3) (v 4)) - -notation Γ "-Predicate " P " via " p => DefinedPred ℒₒᵣ Γ P p - -notation Γ "-Relation " P " via " p => DefinedRel ℒₒᵣ Γ P p - -notation Γ "-Relation₃ " P " via " p => DefinedRel₃ ℒₒᵣ Γ P p - -notation Γ "-Relation₄ " P " via " p => DefinedRel₄ ℒₒᵣ Γ P p - -notation Γ "-Function₁ " f " via " p => DefinedFunction₁ ℒₒᵣ Γ f p - -notation Γ "-Function₂ " f " via " p => DefinedFunction₂ ℒₒᵣ Γ f p - -notation Γ "-Function₃ " f " via " p => DefinedFunction₃ ℒₒᵣ Γ f p - -notation Γ "-Function₄ " f " via " p => DefinedFunction₄ ℒₒᵣ Γ f p - -notation Γ "-Function₅ " f " via " p => DefinedFunction₅ ℒₒᵣ Γ f p - -notation Γ "-Predicate " P => DefinablePred ℒₒᵣ Γ P - -notation Γ "-Relation " P => DefinableRel ℒₒᵣ Γ P - -notation Γ "-Relation₃ " P => DefinableRel₃ ℒₒᵣ Γ P - -notation Γ "-Relation₄ " P => DefinableRel₄ ℒₒᵣ Γ P - -notation Γ "-Relation₅ " P => DefinableRel₅ ℒₒᵣ Γ P - -notation Γ "-Function₁ " f => DefinableFunction₁ ℒₒᵣ Γ f - -notation Γ "-Function₂ " f => DefinableFunction₂ ℒₒᵣ Γ f - -notation Γ "-Function₃ " f => DefinableFunction₃ ℒₒᵣ Γ f - -notation Γ "-Function₄ " f => DefinableFunction₄ ℒₒᵣ Γ f - -variable {L Γ} - -section - -variable {k} {P Q : (Fin k → M) → Prop} - -namespace Defined - -lemma df {R : (Fin k → M) → Prop} {Γ} {p : HSemisentence L k Γ} (h : Defined R p) : FirstOrder.Defined R p.val := - match Γ with - | (𝚺, _) => h - | (𝚷, _) => h - | (𝚫, _) => h.2 - -lemma proper {R : (Fin k → M) → Prop} {m} {p : HSemisentence L k (𝚫, m)} (h : Defined R p) : p.ProperOn M := h.1 - -lemma of_zero {R : (Fin k → M) → Prop} {Γ} {p : HSemisentence L k (𝚺, 0)} (h : Defined R p) : Defined R (p.ofZero Γ) := - match Γ with - | (𝚺, m) => by intro _; simp [h.iff] - | (𝚷, m) => by intro _; simp [h.iff] - | (𝚫, m) => ⟨by simp, by intro _; simp [h.iff]⟩ - -lemma emb {R : (Fin k → M) → Prop} {Γ} {p : HSemisentence ℒₒᵣ k Γ} (h : Defined R p) : Defined R (p.emb L) := - match Γ with - | (𝚺, m) => by intro _; simp [h.iff] - | (𝚷, m) => by intro _; simp [h.iff] - | (𝚫, m) => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ - -lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) - {p : HSemisentence L k Γ} (H : Defined Q p) : Defined P p := by - rwa [show P = Q from by funext v; simp [h]] - -lemma to_definable (p : HSemisentence L k Γ) (hP : Defined P p) : Definable L Γ P := ⟨p.rew Rew.emb, by - match Γ with - | (𝚺, _) => intro; simp [hP.iff] - | (𝚷, _) => intro; simp [hP.iff] - | (𝚫, _) => exact ⟨ - fun v ↦ by rcases p; simpa [HSemiformula.rew] using hP.proper.rew Rew.emb v, - by intro; simp [hP.df.iff]⟩⟩ - -lemma to_definable₀ (p : HSemisentence L k (𝚺, 0)) (hP : Defined P p) : - Definable L Γ P := Defined.to_definable (p.ofZero Γ) hP.of_zero - -lemma to_definable_oRing (p : HSemisentence ℒₒᵣ k Γ) (hP : Defined P p) : - Definable L Γ P := Defined.to_definable (p.emb L) hP.emb - -lemma to_definable_oRing₀ (p : 𝚺₀-Semisentence k) (hP : Defined P p) : - Definable L Γ P := Defined.to_definable₀ (p.emb L) hP.emb - -end Defined - -namespace DefinedFunction - -lemma of_eq {f g : (Fin k → M) → M} (h : ∀ x, f x = g x) - {p : HSemisentence L (k + 1) Γ} (H : DefinedFunction f p) : DefinedFunction g p := - Defined.of_iff (by intro; simp [h]) H - -lemma graph_delta {f : (Fin k → M) → M} {p : HSemisentence L (k + 1) (𝚺, m)} - (h : DefinedFunction f p) : DefinedFunction f p.graphDelta := - ⟨by cases' m with m <;> simp [HSemiformula.graphDelta] - intro e; simp [Empty.eq_elim, h.df.iff] - rw [eq_comm], - by intro v; simp [h.df.iff]⟩ - -end DefinedFunction - -namespace DefinedWithParam - -lemma df {R : (Fin k → M) → Prop} {Γ} {p : HSemiformula L M k Γ} (h : DefinedWithParam R p) : FirstOrder.DefinedWithParam R p.val := - match Γ with - | (𝚺, _) => h - | (𝚷, _) => h - | (𝚫, _) => h.2 - -lemma proper {R : (Fin k → M) → Prop} {m} {p : HSemiformula L M k (𝚫, m)} (h : DefinedWithParam R p) : p.ProperWithParamOn M := h.1 - -lemma of_zero {R : (Fin k → M) → Prop} {Γ} {p : HSemiformula L M k (Γ', 0)} - (h : DefinedWithParam R p) : DefinedWithParam R (p.ofZero Γ) := - match Γ with - | (𝚺, m) => by intro _; simp [h.df.iff] - | (𝚷, m) => by intro _; simp [h.df.iff] - | (𝚫, m) => ⟨by simp , by intro _; simp [h.df.iff]⟩ - -lemma of_deltaOne {R : (Fin k → M) → Prop} {Γ m} {p : HSemiformula L M k 𝚫₁} - (h : DefinedWithParam R p) : DefinedWithParam R (p.ofDeltaOne Γ m) := - match Γ with - | 𝚺 => by intro _; simp [HSemiformula.ofDeltaOne, h.df.iff, HSemiformula.val_sigma] - | 𝚷 => by intro _; simp [HSemiformula.ofDeltaOne, h.df.iff, h.proper.iff'] - | 𝚫 => ⟨by intro _; simp [HSemiformula.ofDeltaOne, h.df.iff, HSemiformula.val_sigma, h.proper.iff'], - by intro _; simp [HSemiformula.ofDeltaOne, h.df.iff, HSemiformula.val_sigma]⟩ - -lemma emb {R : (Fin k → M) → Prop} {Γ} {p : HSemiformula ℒₒᵣ M k Γ} - (h : DefinedWithParam R p) : DefinedWithParam R (p.emb L) := - match Γ with - | (𝚺, m) => by intro _; simp [h.iff] - | (𝚷, m) => by intro _; simp [h.iff] - | (𝚫, m) => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ - -lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) - {p : HSemiformula L M k Γ} (H : DefinedWithParam Q p) : DefinedWithParam P p := by - rwa [show P = Q from by funext v; simp [h]] - -lemma to_definable {p : HSemiformula L M k Γ} (h : DefinedWithParam P p) : Definable L Γ P := ⟨p, h⟩ - -lemma to_definable₀ {p : HSemiformula L M k (Γ', 0)} - (h : DefinedWithParam P p) : Definable L Γ P := ⟨p.ofZero Γ, h.of_zero⟩ - -lemma to_definable_deltaOne {p : HSemiformula L M k 𝚫₁} {Γ m} - (h : DefinedWithParam P p) : Definable L (Γ, m + 1) P := ⟨p.ofDeltaOne Γ m, h.of_deltaOne⟩ - -lemma retraction {p : HSemiformula L M k Γ} (hp : DefinedWithParam P p) (f : Fin k → Fin l) : - DefinedWithParam (fun v ↦ P fun i ↦ v (f i)) (p.rew <| Rew.substs fun x ↦ #(f x)) := - match Γ with - | (𝚺, _) => by intro; simp [hp.df.iff] - | (𝚷, _) => by intro; simp [hp.df.iff] - | (𝚫, _) => ⟨hp.proper.rew _, by intro; simp [hp.df.iff]⟩ - -@[simp] lemma verum : - DefinedWithParam (fun _ ↦ True) (⊤ : HSemiformula L M k Γ) := - match Γ with - | (𝚺, m) => by intro v; simp - | (𝚷, m) => by intro v; simp - | (𝚫, m) => ⟨by simp, by intro v; simp⟩ - -@[simp] lemma falsum : - DefinedWithParam (fun _ ↦ False) (⊥ : HSemiformula L M k Γ) := - match Γ with -| (𝚺, m) => by intro v; simp - | (𝚷, m) => by intro v; simp - | (𝚫, m) => ⟨by simp, by intro v; simp⟩ - -lemma and {p q : HSemiformula L M k Γ} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ∧ Q x) (p ⋏ q) := - match Γ with - | (𝚺, m) => by intro v; simp [hp.iff, hq.iff] - | (𝚷, m) => by intro v; simp [hp.iff, hq.iff] - | (𝚫, m) => ⟨hp.proper.and hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ - -lemma or {p q : HSemiformula L M k Γ} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ∨ Q x) (p ⋎ q) := - match Γ with - | (𝚺, m) => by intro v; simp [hp.iff, hq.iff] - | (𝚷, m) => by intro v; simp [hp.iff, hq.iff] - | (𝚫, m) => ⟨hp.proper.or hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ - -lemma negSigma {p : HSemiformula L M k (𝚺, m)} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) p.negSigma := by intro v; simp [hp.iff] - -lemma negPi {p : HSemiformula L M k (𝚷, m)} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) p.negPi := by intro v; simp [hp.iff] - -lemma not {p : HSemiformula L M k (𝚫, m)} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) (~p) := ⟨hp.proper.neg, by intro v; simp [hp.proper.eval_neg, hp.df.iff]⟩ - -lemma imp {p q : HSemiformula L M k (𝚫, m)} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x → Q x) (p ⟶ q) := (hp.not.or hq).of_iff (by intro x; simp [imp_iff_not_or]) - -lemma iff {p q : HSemiformula L M k (𝚫, m)} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ↔ Q x) (p ⟷ q) := ((hp.imp hq).and (hq.imp hp)).of_iff <| by intro v; simp [iff_iff_implies_and_implies] - -lemma ball {P : (Fin (k + 1) → M) → Prop} {p : HSemiformula L M (k + 1) Γ} - (hp : DefinedWithParam P p) (t : Semiterm L M k) : - DefinedWithParam (fun v ↦ ∀ x < t.valm M v id, P (x :> v)) (HSemiformula.ball t p) := - match Γ with - | (𝚺, m) => by intro v; simp [hp.df.iff] - | (𝚷, m) => by intro v; simp [hp.df.iff] - | (𝚫, m) => ⟨hp.proper.ball, by intro v; simp [hp.df.iff]⟩ - -lemma bex {P : (Fin (k + 1) → M) → Prop} {p : HSemiformula L M (k + 1) Γ} - (hp : DefinedWithParam P p) (t : Semiterm L M k) : - DefinedWithParam (fun v ↦ ∃ x < t.valm M v id, P (x :> v)) (HSemiformula.bex t p) := - match Γ with - | (𝚺, m) => by intro v; simp [hp.df.iff] - | (𝚷, m) => by intro v; simp [hp.df.iff] - | (𝚫, m) => ⟨hp.proper.bex, by intro v; simp [hp.df.iff]⟩ - -lemma ex {P : (Fin (k + 1) → M) → Prop} {p : HSemiformula L M (k + 1) (𝚺, m + 1)} - (hp : DefinedWithParam P p) : - DefinedWithParam (fun v ↦ ∃ x, P (x :> v)) p.ex := by intro _; simp [hp.df.iff] - -lemma all {P : (Fin (k + 1) → M) → Prop} {p : HSemiformula L M (k + 1) (𝚷, m + 1)} - (hp : DefinedWithParam P p) : - DefinedWithParam (fun v ↦ ∀ x, P (x :> v)) p.all := by intro _; simp [hp.df.iff] - -end DefinedWithParam - -namespace Definable - -lemma mkPolarity {P : (Fin k → M) → Prop} {Γ : Polarity} - (p : Semiformula L M k) (hp : Hierarchy Γ m p) (hP : ∀ v, P v ↔ Semiformula.Evalm M v id p) : Definable L (Γ, m) P := - match Γ with - | 𝚺 => ⟨.mkSigma p hp, by intro v; simp [hP]⟩ - | 𝚷 => ⟨.mkPi p hp, by intro v; simp [hP]⟩ - -lemma of_iff (Q : (Fin k → M) → Prop) (h : ∀ x, P x ↔ Q x) (H : Definable L Γ Q) : Definable L Γ P := by - rwa [show P = Q from by funext v; simp [h]] - -lemma of_oRing (h : Definable ℒₒᵣ Γ P) : Definable L Γ P := by - rcases h with ⟨p, hP⟩; exact ⟨p.emb L, hP.emb⟩ - -lemma of_delta (h : Definable L (𝚫, m) P) {Γ} : Definable L (Γ, m) P := by - rcases h with ⟨p, h⟩ - match Γ with - | 𝚺 => exact ⟨p.sigma, by intro v; simp [HSemiformula.val_sigma, h.df.iff]⟩ - | 𝚷 => exact ⟨p.pi, by intro v; simp [←h.proper v, HSemiformula.val_sigma, h.df.iff]⟩ - | 𝚫 => exact ⟨p, h⟩ - -instance [Definable L (𝚫, m) P] (Γ) : Definable L (Γ, m) P := of_delta inferInstance - -lemma of_sigma_of_pi (hσ : Definable L (𝚺, m) P) (hπ : Definable L (𝚷, m) P) : Definable L (𝚫, m) P := by - rcases hσ with ⟨p, hp⟩; rcases hπ with ⟨q, hq⟩ - exact ⟨.mkDelta p q, by intro v; simp [hp.df.iff, hq.df.iff], by intro v; simp [hp.df.iff]⟩ - -instance [Definable ℒₒᵣ Γ P] : Definable L Γ P := Definable.of_oRing inferInstance - -lemma of_zero (h : Definable L (Γ', 0) P) (Γ) : Definable L Γ P := by - rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable₀ - -lemma of_deltaOne (h : Definable L 𝚫₁ P) (Γ m) : Definable L (Γ, m + 1) P := by - rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable_deltaOne - -instance [Definable L 𝚺₀ P] (Γ) : Definable L Γ P := Definable.of_zero (Γ' := 𝚺) inferInstance Γ - -lemma retraction (h : Definable L Γ P) (f : Fin k → Fin n) : - Definable L Γ fun v ↦ P (fun i ↦ v (f i)) := by - rcases h with ⟨p, h⟩ - exact ⟨p.rew (Rew.substs (fun i ↦ #(f i))), - match Γ with - | (𝚺, _) => by intro; simp [h.df.iff] - | (𝚷, _) => by intro; simp [h.df.iff] - | (𝚫, _) => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ - -lemma retractiont (h : Definable L Γ P) (f : Fin k → Semiterm L M n) : - Definable L Γ fun v ↦ P (fun i ↦ Semiterm.valm M v id (f i)) := by - rcases h with ⟨p, h⟩ - exact ⟨p.rew (Rew.substs f), - match Γ with - | (𝚺, _) => by intro; simp [h.df.iff] - | (𝚷, _) => by intro; simp [h.df.iff] - | (𝚫, _) => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ - -lemma const {P : Prop} : Definable L Γ (fun _ : Fin k → M ↦ P) := of_zero (by - by_cases hP : P - · exact ⟨.mkSigma ⊤ (by simp), by intro; simp[hP]⟩ - · exact ⟨.mkSigma ⊥ (by simp), by intro; simp[hP]⟩) Γ - -lemma and (h₁ : Definable L Γ P) (h₂ : Definable L Γ Q) : - Definable L Γ (fun v ↦ P v ∧ Q v) := by - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁ ⋏ p₂, h₁.and h₂⟩ - -lemma or (h₁ : Definable L Γ P) (h₂ : Definable L Γ Q) : - Definable L Γ (fun v ↦ P v ∨ Q v) := by - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁ ⋎ p₂, h₁.or h₂⟩ - -lemma not {Γ} (h : Definable L (SigmaPiDelta.alt Γ, m) P) : - Definable L (Γ, m) (fun v ↦ ¬P v) := by - match Γ with - | 𝚺 => rcases h with ⟨p, h⟩; exact ⟨p.negPi, h.negPi⟩ - | 𝚷 => rcases h with ⟨p, h⟩; exact ⟨p.negSigma, h.negSigma⟩ - | 𝚫 => rcases h with ⟨p, h⟩; exact ⟨p.negDelta, h.not⟩ - -lemma imp {Γ} (h₁ : Definable L (SigmaPiDelta.alt Γ, m) P) (h₂ : Definable L (Γ, m) Q) : - Definable L (Γ, m) (fun v ↦ P v → Q v) := by - match Γ with - | 𝚺 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁.negPi.or p₂, (h₁.negPi.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ - | 𝚷 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁.negSigma.or p₂, (h₁.negSigma.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ - | 𝚫 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩; exact ⟨p₁ ⟶ p₂, h₁.imp h₂⟩ - -lemma iff (h₁ : Definable L (𝚫, m) P) (h₂ : Definable L (𝚫, m) Q) {Γ} : - Definable L (Γ, m) (fun v ↦ P v ↔ Q v) := - .of_delta (by rcases h₁ with ⟨p, hp⟩; rcases h₂ with ⟨q, hq⟩; exact ⟨p ⟷ q, hp.iff hq⟩) - -lemma all {P : (Fin k → M) → M → Prop} (h : Definable L (𝚷, s + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (𝚷, s + 1) (fun v ↦ ∀ x, P v x) := by - rcases h with ⟨p, hp⟩ - exact ⟨.mkPi (∀' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ - -lemma ex {P : (Fin k → M) → M → Prop} (h : Definable L (𝚺, s + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (𝚺, s + 1) (fun v ↦ ∃ x, P v x) := by - rcases h with ⟨p, hp⟩ - exact ⟨.mkSigma (∃' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ - -lemma comp₁ {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : DefinableFunction L (𝚺, m + 1) f) - {Γ : SigmaPiDelta} (hP : DefinablePred L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f v)) := by - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact ⟨(pf ⋏ (p.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact ⟨(pf.negSigma ⋎ (p.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact of_sigma_of_pi - ⟨(pf ⋏ (p.sigma.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff, HSemiformula.val_sigma] ⟩ - ⟨(pf.negSigma ⋎ (p.pi.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₁' {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : DefinableFunction L (𝚺, m + 1) f) - {Γ : SigmaPiDelta} [DefinablePred L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f v)) := - comp₁ hf inferInstance - -lemma comp₂ {k} {P : M → M → Prop} {f g : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (hg : DefinableFunction L (𝚺, m + 1) g) - {Γ : SigmaPiDelta} (hP : DefinableRel L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f v) (g v)) := by - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.rew (Rew.substs ![#0, #1]))).ex.ex, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.rew (Rew.substs ![#0, #1]))).all.all, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact of_sigma_of_pi - ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.sigma.rew (Rew.substs ![#0, #1]))).ex.ex, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, HSemiformula.val_sigma]⟩ - ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma - ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.pi.rew (Rew.substs ![#0, #1]))).all.all, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₂' {k} {P : M → M → Prop} {f g : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (hg : DefinableFunction L (𝚺, m + 1) g) - {Γ : SigmaPiDelta} [DefinableRel L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f v) (g v)) := - comp₂ hf hg inferInstance - -lemma comp₃ {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) - {Γ : SigmaPiDelta} (hP : DefinableRel₃ L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, HSemiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₃' {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) - {Γ : SigmaPiDelta} [DefinableRel₃ L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := - comp₃ hf₁ hf₂ hf₃ inferInstance - -lemma comp₄ {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - {Γ : SigmaPiDelta} (hP : DefinableRel₄ L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩; rcases hf₄ with ⟨pf₄, hpf₄⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, HSemiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₄' {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - {Γ : SigmaPiDelta} [DefinableRel₄ L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := - comp₄ hf₁ hf₂ hf₃ hf₄ inferInstance - -lemma comp₅ {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - (hf₅ : DefinableFunction L (𝚺, m + 1) f₅) - {Γ : SigmaPiDelta} (hP : DefinableRel₅ L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, HSemiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₅' {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - (hf₅ : DefinableFunction L (𝚺, m + 1) f₅) - {Γ : SigmaPiDelta} [DefinableRel₅ L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := - comp₅ hf₁ hf₂ hf₃ hf₄ hf₅ inferInstance - -lemma comp₆ {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - (hf₅ : DefinableFunction L (𝚺, m + 1) f₅) (hf₆ : DefinableFunction L (𝚺, m + 1) f₆) - {Γ : SigmaPiDelta} (hP : DefinableRel₆ L (Γ, m + 1) P) : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩; rcases hf₆ with ⟨pf₆, hpf₆⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, HSemiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₆' {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} - (hf₁ : DefinableFunction L (𝚺, m + 1) f₁) (hf₂ : DefinableFunction L (𝚺, m + 1) f₂) - (hf₃ : DefinableFunction L (𝚺, m + 1) f₃) (hf₄ : DefinableFunction L (𝚺, m + 1) f₄) - (hf₅ : DefinableFunction L (𝚺, m + 1) f₅) (hf₆ : DefinableFunction L (𝚺, m + 1) f₆) - {Γ : SigmaPiDelta} [DefinableRel₆ L (Γ, m + 1) P] : Definable L (Γ, m + 1) (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := - comp₆ hf₁ hf₂ hf₃ hf₄ hf₅ hf₆ inferInstance - -end Definable - -lemma DefinablePred.of_iff {P : M → Prop} (Q) (h : ∀ x, P x ↔ Q x) (H : DefinablePred L Γ Q) : DefinablePred L Γ P := by - rwa [show P = Q from by funext v; simp [h]] - -instance DefinableFunction₁.graph {f : M → M} [h : DefinableFunction₁ L Γ f] : - DefinableRel L Γ (Function.Graph f) := h - -instance DefinableFunction₂.graph {f : M → M → M} [h : DefinableFunction₂ L Γ f] : - DefinableRel₃ L Γ (Function.Graph₂ f) := h - -instance DefinableFunction₃.graph {f : M → M → M → M} [h : DefinableFunction₃ L Γ f] : - DefinableRel₄ L Γ (Function.Graph₃ f) := h - -namespace DefinableFunction - -lemma graph_delta {k} {f : (Fin k → M) → M} - (h : DefinableFunction L (𝚺, m) f) : DefinableFunction L (𝚫, m) f := by - rcases h with ⟨p, h⟩ - exact ⟨p.graphDelta, by - cases' m with m <;> simp [HSemiformula.graphDelta] - intro e; simp [Empty.eq_elim, h.df.iff] - exact eq_comm, by - intro v; simp [h.df.iff]⟩ - -instance {k} {f : (Fin k → M) → M} [h : DefinableFunction L (𝚺, m) f] : DefinableFunction L (𝚫, m) f := - DefinableFunction.graph_delta h - -instance {k} {f : (Fin k → M) → M} [DefinableFunction L 𝚺₀ f] (Γ) : DefinableFunction L Γ f := inferInstance - -lemma of_sigmaOne {k} {f : (Fin k → M) → M} - (h : DefinableFunction L 𝚺₁ f) (Γ m) : DefinableFunction L (Γ, m + 1) f := Definable.of_deltaOne (graph_delta h) Γ m - -@[simp] lemma var {k} (i : Fin k) : DefinableFunction L Γ (fun v : Fin k → M ↦ v i) := - .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!#i.succ” (by simp), by intro _; simp⟩ _ - -@[simp] lemma const {k} (c : M) : DefinableFunction L Γ (fun _ : Fin k → M ↦ c) := - .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | #0 = &c” (by simp), by intro v; simp⟩ _ - -@[simp] lemma term_retraction (t : Semiterm L M n) (e : Fin n → Fin k) : - DefinableFunction L Γ fun v : Fin k → M ↦ Semiterm.valm M (fun x ↦ v (e x)) id t := - .of_zero (Γ' := 𝚺) - ⟨.mkSigma “x | x = !!(Rew.substs (fun x ↦ #(e x).succ) t)” (by simp), by intro v; simp [Semiterm.val_substs]⟩ _ - -@[simp] lemma term (t : Semiterm L M k) : - DefinableFunction L Γ fun v : Fin k → M ↦ Semiterm.valm M v id t := - .of_zero (Γ' := 𝚺) - ⟨.mkSigma “x | x = !!(Rew.bShift t)” (by simp), by intro v; simp [Semiterm.val_bShift']⟩ _ - -lemma of_eq {f : (Fin k → M) → M} (g) (h : ∀ v, f v = g v) (H : DefinableFunction L Γ f) : DefinableFunction L Γ g := by - rwa [show g = f from by funext v; simp [h]] - -lemma retraction {f : (Fin k → M) → M} (hf : DefinableFunction L Γ f) (e : Fin k → Fin n) : - DefinableFunction L Γ fun v ↦ f (fun i ↦ v (e i)) := by - have := Definable.retraction (n := n + 1) hf (0 :> fun i ↦ (e i).succ); simp at this - exact this.of_iff _ (by intro x; simp) - -lemma retractiont {f : (Fin k → M) → M} (hf : DefinableFunction L Γ f) (t : Fin k → Semiterm L M n) : - DefinableFunction L Γ fun v ↦ f (fun i ↦ Semiterm.valm M v id (t i)) := by - have := Definable.retractiont (n := n + 1) hf (#0 :> fun i ↦ Rew.bShift (t i)); simp at this - exact this.of_iff _ (by intro x; simp [Semiterm.val_bShift']) - -lemma rel {f : (Fin k → M) → M} (h : DefinableFunction L Γ f) : - Definable L Γ (fun v ↦ v 0 = f (v ·.succ)) := h - -end DefinableFunction - -lemma DefinableFunction₁.comp {Γ} {k} {f : M → M} {g : (Fin k → M) → M} - (hf : DefinableFunction₁ L (Γ, m + 1) f) (hg : DefinableFunction L (𝚺, m + 1) g) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g v)) := by - simpa using Definable.comp₂ (P := Function.Graph f) (DefinableFunction.var 0) (hg.retraction Fin.succ) hf - -lemma DefinableFunction₂.comp {Γ} {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} - (hf : DefinableFunction₂ L (Γ, m + 1) f) (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v)) := by - simpa using Definable.comp₃ (P := Function.Graph₂ f) (DefinableFunction.var 0) (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) hf - -lemma DefinableFunction₃.comp {Γ} {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} - (hf : DefinableFunction₃ L (Γ, m + 1) f) (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) - (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := by - simpa using Definable.comp₄ (P := Function.Graph₃ f) (DefinableFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) hf - -lemma DefinableFunction₄.comp {Γ} {k} {f : M → M → M → M → M} {g₁ g₂ g₃ g₄ : (Fin k → M) → M} - (hf : DefinableFunction₄ L (Γ, m + 1) f) (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) - (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) - (hg₄ : DefinableFunction L (𝚺, m + 1) g₄) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := by - simpa using Definable.comp₅ (P := Function.Graph₄ f) (DefinableFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) (hg₄.retraction Fin.succ) hf - -lemma DefinableFunction₅.comp {Γ} {k} {f : M → M → M → M → M → M} {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} - (hf : DefinableFunction₅ L (Γ, m + 1) f) (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) - (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) - (hg₄ : DefinableFunction L (𝚺, m + 1) g₄) (hg₅ : DefinableFunction L (𝚺, m + 1) g₅) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := by - simpa using Definable.comp₆ (P := Function.Graph₅ f) (DefinableFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) - (hg₄.retraction Fin.succ) (hg₅.retraction Fin.succ) hf - -lemma DefinableFunction.comp₁ {Γ} {k} {f : M → M} [DefinableFunction₁ L (Γ, m + 1) f] - {g : (Fin k → M) → M} (hg : DefinableFunction L (𝚺, m + 1) g) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g v)) := - DefinableFunction₁.comp inferInstance hg - -lemma DefinableFunction.comp₂ {Γ} {k} {f : M → M → M} [DefinableFunction₂ L (Γ, m + 1) f] - {g₁ g₂ : (Fin k → M) → M} (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v)) := - DefinableFunction₂.comp inferInstance hg₁ hg₂ - -lemma DefinableFunction.comp₃ {Γ} {k} {f : M → M → M → M} [DefinableFunction₃ L (Γ, m + 1) f] - {g₁ g₂ g₃ : (Fin k → M) → M} - (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := - DefinableFunction₃.comp inferInstance hg₁ hg₂ hg₃ - -lemma DefinableFunction.comp₄ {Γ} {k} {f : M → M → M → M → M} [DefinableFunction₄ L (Γ, m + 1) f] - {g₁ g₂ g₃ g₄ : (Fin k → M) → M} - (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) - (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) (hg₄ : DefinableFunction L (𝚺, m + 1) g₄) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := - DefinableFunction₄.comp inferInstance hg₁ hg₂ hg₃ hg₄ - -lemma DefinableFunction.comp₅ {Γ} {k} {f : M → M → M → M → M → M} [DefinableFunction₅ L (Γ, m + 1) f] - {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} - (hg₁ : DefinableFunction L (𝚺, m + 1) g₁) (hg₂ : DefinableFunction L (𝚺, m + 1) g₂) - (hg₃ : DefinableFunction L (𝚺, m + 1) g₃) (hg₄ : DefinableFunction L (𝚺, m + 1) g₄) - (hg₅ : DefinableFunction L (𝚺, m + 1) g₅) : - DefinableFunction L (Γ, m + 1) (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := - DefinableFunction₅.comp inferInstance hg₁ hg₂ hg₃ hg₄ hg₅ - -namespace DefinableRel - -@[simp] instance eq : DefinableRel L Γ (Eq : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1” (by simp)) (by intro _; simp) - -@[simp] instance lt : DefinableRel L Γ (LT.lt : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 < #1” (by simp)) (by intro _; simp) - -@[simp] instance le : DefinableRel L Γ (LE.le : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 ≤ #1” (by simp)) (by intro _; simp) - -end DefinableRel - -namespace DefinableFunction₂ - -@[simp] instance add : DefinableFunction₂ L Γ ((· + ·) : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) - -@[simp] instance mul : DefinableFunction₂ L Γ ((· * ·) : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) - -@[simp] instance hAdd : DefinableFunction₂ L Γ (HAdd.hAdd : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) - -@[simp] instance hMul : DefinableFunction₂ L Γ (HMul.hMul : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) - -end DefinableFunction₂ - -namespace Definable - -lemma ball_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∀ x < f v, P v x) := by - rcases hf with ⟨bf, hbf⟩ - rcases h with ⟨p, hp⟩ - match Γ with - | 𝚺 => exact - ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚷 => exact - ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚫 => - exact .of_sigma_of_pi - ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, HSemiformula.val_sigma] ⟩ - ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ - -lemma bex_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∃ x < f v, P v x) := by - rcases hf with ⟨bf, hbf⟩ - rcases h with ⟨p, hp⟩ - match Γ with - | 𝚺 => exact - ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚷 => exact - ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚫 => - exact .of_sigma_of_pi - ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, HSemiformula.val_sigma] ⟩ - ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ - -lemma ball_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∀ x ≤ f v, P v x) := by - have : Definable L (Γ, m + 1) (fun v ↦ ∀ x < f v + 1, P v x) := ball_lt (DefinableFunction₂.comp (by simp) hf (by simp)) h - exact this.of_iff <| by intro v; simp [lt_succ_iff_le] - -lemma bex_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∃ x ≤ f v, P v x) := by - have : Definable L (Γ, m + 1) (fun v ↦ ∃ x < f v + 1, P v x) := bex_lt (DefinableFunction₂.comp (by simp) hf (by simp)) h - exact this.of_iff <| by intro v; simp [lt_succ_iff_le] - -lemma ball_lt' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∀ {x}, x < f v → P v x) := ball_lt hf h - -lemma ball_le' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction L (𝚺, m + 1) f) (h : Definable L (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable L (Γ, m + 1) (fun v ↦ ∀ {x}, x ≤ f v → P v x) := ball_le hf h - -end Definable - -end - -end Arith - -end LO.FirstOrder diff --git a/Arithmetization/Definability/Definability.lean b/Arithmetization/Definability/PolyBounded.lean similarity index 99% rename from Arithmetization/Definability/Definability.lean rename to Arithmetization/Definability/PolyBounded.lean index 8140745..23c3e63 100644 --- a/Arithmetization/Definability/Definability.lean +++ b/Arithmetization/Definability/PolyBounded.lean @@ -347,18 +347,18 @@ attribute [aesop (rule_sets := [Definability]) norm] Definable.const attribute [aesop 5 (rule_sets := [Definability]) safe] - DefinableFunction.comp₁ - DefinableFunction.comp₂ - DefinableFunction.comp₃ + DefinableFunction.comp₁_infer + DefinableFunction.comp₂_infer + DefinableFunction.comp₃_infer DefinableBoundedFunction.comp₁ DefinableBoundedFunction.comp₂ DefinableBoundedFunction.comp₃ attribute [aesop 6 (rule_sets := [Definability]) safe] - Definable.comp₁' - Definable.comp₂' - Definable.comp₃' - Definable.comp₄' + Definable.comp₁_infer + Definable.comp₂_infer + Definable.comp₃_infer + Definable.comp₄_infer Definable.const attribute [aesop 7 (rule_sets := [Definability]) safe] diff --git a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean index 12f6464..f969aad 100644 --- a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean +++ b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean @@ -138,7 +138,7 @@ lemma mem_limSeq_succ_iff {x s : M} : lemma limSeq_cumulative {s s' : M} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' := by induction s' using induction_iSigmaOne generalizing s · apply Definable.ball_le' (by definability) - apply Definable.comp₂' + apply Definable.comp₂_infer · exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ · exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #2 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ case zero => @@ -156,9 +156,9 @@ lemma mem_limSeq_self [c.StrongFinite] {u s : M} : induction u using order_induction_piOne generalizing s · apply Definable.all apply Definable.imp - · apply Definable.comp₂' (by definability) + · apply Definable.comp₂_infer (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ - · apply Definable.comp₂' (by definability) + · apply Definable.comp₂_infer (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> ‘#2 + 1’ :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ case ind u ih => rcases zero_or_succ s with (rfl | ⟨s, rfl⟩) @@ -194,14 +194,14 @@ lemma finite_upperbound (m : M) : ∃ s, ∀ z < m, c.Fixpoint v z → z ∈ c.l have : ∃ F : M, ∀ x, x ∈ F ↔ x < m ∧ c.Fixpoint v x := by have : 𝚺₁-Predicate fun x ↦ x < m ∧ c.Fixpoint v x := Definable.and (by definability) - (Definable.ex (Definable.comp₂' (by definability) + (Definable.ex (Definable.comp₂_infer (by definability) ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩)) exact finite_comprehension₁! this ⟨m, fun i hi ↦ hi.1⟩ |>.exists rcases this with ⟨F, hF⟩ have : ∀ x ∈ F, ∃ u, x ∈ c.limSeq v u := by intro x hx; exact hF x |>.mp hx |>.2 have : ∃ f, IsMapping f ∧ domain f = F ∧ ∀ (x y : M), ⟪x, y⟫ ∈ f → x ∈ c.limSeq v y := sigmaOne_skolem - (by apply Definable.comp₂' (by definability) + (by apply Definable.comp₂_infer (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #2 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩) this rcases this with ⟨f, mf, rfl, hf⟩ exact ⟨f, by diff --git a/Arithmetization/ISigmaOne/HFS/Seq.lean b/Arithmetization/ISigmaOne/HFS/Seq.lean index 289d668..3f8adea 100644 --- a/Arithmetization/ISigmaOne/HFS/Seq.lean +++ b/Arithmetization/ISigmaOne/HFS/Seq.lean @@ -437,32 +437,32 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function ∀ l < k, ∀ m < W, ∀ m' < W, ⟪l, m⟫ ∈ W → ⟪l + 1, m'⟫ ∈ W → ∀ x' ≤ x - l, ∀ y' ≤ m, f x' y' ≤ m' := by intro k hk induction k using induction_iSigmaOne - · apply Definable.imp (Definable.comp₂' (DefinableFunction.var _) (DefinableFunction.const _)) + · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) apply Definable.ex - apply Definable.and (Definable.comp₁' (DefinableFunction.var _)) + apply Definable.and (Definable.comp₁_infer (DefinableFunction.var _)) apply Definable.and - (Definable.comp₂' - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.const _)) - (DefinableFunction.comp₁ <| DefinableFunction.var _)) + (Definable.comp₂_infer + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) + (DefinableFunction.comp₁_infer <| DefinableFunction.var _)) apply Definable.and - (Definable.comp₂' (DefinableFunction.comp₂ (DefinableFunction.const _) (DefinableFunction.const _)) (DefinableFunction.var _)) + (Definable.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.const _)) (DefinableFunction.var _)) apply Definable.ball_lt (DefinableFunction.var _) apply Definable.ball_lt (DefinableFunction.var _) apply Definable.ball_lt (DefinableFunction.var _) apply Definable.imp - (Definable.comp₂' - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.var _)) + (Definable.comp₂_infer + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _)) (DefinableFunction.var _)) apply Definable.imp - (Definable.comp₂' - (DefinableFunction.comp₂ - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.const _)) + (Definable.comp₂_infer + (DefinableFunction.comp₂_infer + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) (DefinableFunction.var _)) (DefinableFunction.var _)) apply Definable.ball_le - (Definable.comp₂' (DefinableFunction.var _) (DefinableFunction.comp₂ (DefinableFunction.const _) (DefinableFunction.var _))) + (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.var _))) apply Definable.ball_le (DefinableFunction.var _) - apply Definable.comp₂' (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.var _)) (DefinableFunction.var _) + apply Definable.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _)) (DefinableFunction.var _) case zero => exact ⟨!⟦y⟧, by simp⟩ case succ k ih => rcases ih (le_trans le_self_add hk) with ⟨W, SW, hkW, hW₀, hWₛ⟩ @@ -482,15 +482,15 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function have : ∀ i ≤ x, ∀ m < W, ⟪x - i, m⟫ ∈ W → ∀ x' ≤ i, ∀ y' ≤ m, P x' y' := by intro i induction i using induction_iSigmaOne - · apply Definable.imp (Definable.comp₂' (DefinableFunction.var _) (DefinableFunction.const _)) + · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) apply Definable.ball_lt (DefinableFunction.const _) apply Definable.imp - (Definable.comp₂' - (DefinableFunction.comp₂ (DefinableFunction.comp₂ (DefinableFunction.const _) (DefinableFunction.var _)) (DefinableFunction.var _)) + (Definable.comp₂_infer + (DefinableFunction.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.var _)) (DefinableFunction.var _)) (DefinableFunction.const _)) apply Definable.ball_le (DefinableFunction.var _) apply Definable.ball_le (DefinableFunction.var _) - apply Definable.comp₂' (DefinableFunction.var _) (DefinableFunction.var _) + apply Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _) case zero => intro _ _ _ _ _ h y' _ rcases nonpos_iff_eq_zero.mp h diff --git a/Arithmetization/ISigmaOne/Ind.lean b/Arithmetization/ISigmaOne/Ind.lean index 8e3315e..7458dd9 100644 --- a/Arithmetization/ISigmaOne/Ind.lean +++ b/Arithmetization/ISigmaOne/Ind.lean @@ -56,15 +56,15 @@ lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P · clear hp hq ind apply LO.FirstOrder.Arith.Definable.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.Definable.comp₂' + apply LO.FirstOrder.Arith.Definable.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] · simp_all only [Fin.isValue] apply LO.FirstOrder.Arith.Definable.or - · apply LO.FirstOrder.Arith.Definable.comp₂' + · apply LO.FirstOrder.Arith.Definable.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] - · apply LO.FirstOrder.Arith.Definable.comp₂' + · apply LO.FirstOrder.Arith.Definable.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] case ind z ih => diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index d3aadd7..57e7a15 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -1203,18 +1203,18 @@ lemma uformula_result_induction {P : V → V → V → Prop} (hP : 𝚺₁-Relat 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 := - DefinableFunction.comp₂ (f := Max.max) + DefinableFunction.comp₂_infer (f := Max.max) (DefinableFunction.var _) - (DefinableFunction.comp₂ - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.comp₁ (DefinableFunction.var _))) - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.comp₁ (DefinableFunction.var _)))) + (DefinableFunction.comp₂_infer + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _))) + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _)))) apply sigma₁_order_ball_induction hf ?_ ?_ p param · apply Definable.imp - (Definable.comp₁' (DefinableFunction.var _)) - (Definable.comp₃' + (Definable.comp₁_infer (DefinableFunction.var _)) + (Definable.comp₃_infer (DefinableFunction.var _) (DefinableFunction.var _) - (DefinableFunction.comp₂ (DefinableFunction.var _) (DefinableFunction.var _))) + (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _))) intro p param ih hp rcases hp.case with (⟨n, k, r, v, hkr, hv, rfl⟩ | ⟨n, k, r, v, hkr, hv, rfl⟩ | @@ -1260,15 +1260,15 @@ lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺 Fin.succ_zero_eq_one] apply LO.FirstOrder.Arith.Definable.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.Definable.comp₂' + apply LO.FirstOrder.Arith.Definable.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue] - apply LO.FirstOrder.Arith.DefinableFunction.comp₁ + apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · apply LO.FirstOrder.Arith.Definable.comp₄' + · apply LO.FirstOrder.Arith.Definable.comp₄_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue] - apply LO.FirstOrder.Arith.DefinableFunction.comp₁ + apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index 3908ab1..031b711 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -459,9 +459,9 @@ lemma uformula_subst_induction {P : V → V → V → V → Prop} (hP : 𝚺₁- 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 Definable.comp₄' - (DefinableFunction.comp₁ (DefinableFunction.var _)) - (DefinableFunction.comp₁ (DefinableFunction.var _)) + · apply Definable.comp₄_infer + (DefinableFunction.comp₁_infer (DefinableFunction.var _)) + (DefinableFunction.comp₁_infer (DefinableFunction.var _)) (DefinableFunction.var _) (DefinableFunction.var _) · intro param n k R v hkR hv; simpa using hRel n (π₁ param) (π₂ param) k R v hkR hv @@ -498,10 +498,10 @@ lemma semiformula_subst_induction {P : V → V → V → V → V → Prop} (hP : 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 Definable.comp₅' + · apply Definable.comp₅_infer (DefinableFunction.var _) - (DefinableFunction.comp₁ (DefinableFunction.var _)) - (DefinableFunction.comp₁ (DefinableFunction.var _)) + (DefinableFunction.comp₁_infer (DefinableFunction.var _)) + (DefinableFunction.comp₁_infer (DefinableFunction.var _)) (DefinableFunction.var _) (DefinableFunction.var _) · intro param n k R v hkR hv; simpa using hRel n (π₁ param) (π₂ param) k R v hkR hv @@ -611,8 +611,8 @@ lemma substs_substs {p} (hp : L.Semiformula l p) : apply Definable.all apply Definable.imp (by definability) apply Definable.imp (by definability) - apply Definable.comp₂' (by simp; definability) - apply DefinableFunction.comp₃ (by definability) ?_ (by definability) + apply Definable.comp₂_infer (by simp; definability) + apply DefinableFunction.comp₃_infer (by definability) ?_ (by definability) apply DefinableFunction₅.comp (termSubstVec_definable _) <;> 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] From ad18d045117609779c20f6424edb7834bdf618b4 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sat, 10 Aug 2024 02:08:11 +0900 Subject: [PATCH 02/10] add SigmaPiDeltaSystem --- .../Definability/SigmaPiDeltaSystem.lean | 281 ++++++++++++++++++ 1 file changed, 281 insertions(+) create mode 100644 Arithmetization/Definability/SigmaPiDeltaSystem.lean diff --git a/Arithmetization/Definability/SigmaPiDeltaSystem.lean b/Arithmetization/Definability/SigmaPiDeltaSystem.lean new file mode 100644 index 0000000..b25c9e5 --- /dev/null +++ b/Arithmetization/Definability/SigmaPiDeltaSystem.lean @@ -0,0 +1,281 @@ +import Arithmetization.Vorspiel.Lemmata +import Logic.FirstOrder.Arith.StrictHierarchy + +/-! + +# Arithmetical Formula Sorted by Arithmetical Hierarchy + +This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of first-order logic. + +- `𝚺-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚺-[m]`. +- `𝚷-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚷-[m]`. +- `𝚫-[m].Semiformula ξ n` is a pair of `𝚺-[m].Semiformula ξ n` and `𝚷-[m].Semiformula ξ n`. +- `ProperOn` : `p.ProperOn M` iff `p`'s two element `p.sigma` and `p.pi` are equivalent on model `M`. + +-/ + +namespace LO + +class SigmaPiDeltaLike (Ω : Type*) [SigmaSymbol Ω] [PiSymbol Ω] [DeltaSymbol Ω] where + alt : Ω → Ω + +variable {V : Type*} + +class SigmaPiDeltaSystem (V : Type*) where + VecPr : SigmaPiDelta → {k : ℕ} → ((Fin k → V) → Prop) → Prop + vecPr_delta_iff_sigma_and_pi {k} {P : (Fin k → V) → Prop} : VecPr 𝚫 P ↔ VecPr 𝚺 P ∧ VecPr 𝚷 P + verum' (Γ k) : VecPr Γ fun _ : Fin k → V ↦ ⊤ + and' {k} {P Q : (Fin k → V) → Prop} : VecPr Γ P → VecPr Γ Q → VecPr Γ fun x ↦ P x ∧ Q x + not' {k} {P : (Fin k → V) → Prop} : VecPr Γ.alt P → VecPr Γ fun x ↦ ¬P x + all' {k} {P : (Fin k → V) → V → Prop} : VecPr 𝚷 (fun x ↦ P (x ·.succ) (x 0)) → VecPr 𝚷 fun x ↦ ∀ z, P x z + retraction' {k l} {P : (Fin k → V) → Prop} (hP : VecPr Γ P) (f : Fin k → Fin l) : VecPr Γ fun v ↦ P fun i ↦ v (f i) + equal' (Γ) : VecPr Γ fun v : Fin 2 → V ↦ v 0 = v 1 + +abbrev SigmaPiDeltaSystem.VecFunc (𝔖 : SigmaPiDeltaSystem V) + (Γ : SigmaPiDelta) (f : (Fin k → V) → V) : Prop := 𝔖.VecPr Γ fun v ↦ v 0 = f (v ·.succ) + +namespace SigmaPiDeltaSystem + +variable {𝔖 : SigmaPiDeltaSystem V} {Γ : SigmaPiDelta} {k} {P Q : (Fin k → V) → Prop} + +namespace VecPr + +alias verum := verum' + +alias and := and' + +alias not := not' + +alias all := all' + +alias retraction := retraction' + +alias equal := equal' + +lemma of_iff (hP : 𝔖.VecPr Γ P) (h : ∀ x, P x ↔ Q x) : 𝔖.VecPr Γ Q := by + have : P = Q := funext <| by simpa + rcases this + exact hP + +lemma of_sigma_of_pi (hσ : 𝔖.VecPr 𝚺 P) (hπ : 𝔖.VecPr 𝚷 P) : 𝔖.VecPr Γ P := + match Γ with + | 𝚺 => hσ + | 𝚷 => hπ + | 𝚫 => vecPr_delta_iff_sigma_and_pi.mpr ⟨hσ, hπ⟩ + +lemma of_delta (h : 𝔖.VecPr 𝚫 P) {Γ} : 𝔖.VecPr Γ P := + of_sigma_of_pi + (vecPr_delta_iff_sigma_and_pi.mp h |>.1) + (vecPr_delta_iff_sigma_and_pi.mp h |>.2) + +lemma not' (h : 𝔖.VecPr Γ P) : 𝔖.VecPr Γ.alt fun x ↦ ¬P x := not (by simpa) + +lemma of_not (h : 𝔖.VecPr Γ.alt (fun x ↦ ¬P x)) : 𝔖.VecPr Γ P := by simpa using not' h + +lemma falsum (Γ : SigmaPiDelta) (k : ℕ) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ ⊥ := + of_sigma_of_pi (by simpa using not' (verum 𝚷 k)) (by simpa using not' (verum 𝚺 k)) + +@[simp] lemma constant (Γ : SigmaPiDelta) (k : ℕ) (P : Prop) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ P := by + by_cases h : P <;> simp [h] + · apply verum + · apply falsum + +lemma or (hP : 𝔖.VecPr Γ P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x ∨ Q x := + of_not <| by + simp only [not_or]; apply and + · apply not' hP + · apply not' hQ + +lemma imply (hP : 𝔖.VecPr Γ.alt P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x → Q x := by + simp [imp_iff_not_or]; apply or + · apply not hP + · exact hQ + +lemma ex {k} {P : (Fin k → V) → V → Prop} (h : 𝔖.VecPr 𝚺 fun x ↦ P (x ·.succ) (x 0)) : 𝔖.VecPr 𝚺 fun x ↦ ∃ z, P x z := of_not <| by + simpa using all (by apply not' h) + +lemma iff (hP : 𝔖.VecPr 𝚫 P) (hQ : 𝔖.VecPr 𝚫 Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x ↔ Q x := of_delta <| by + simp only [iff_iff_implies_and_implies] + apply and + · exact imply hP hQ + · exact imply hQ hP + +lemma equal' (Γ) (i j : Fin k) : 𝔖.VecPr Γ fun v ↦ v i = v j := by + simpa using retraction (equal Γ) ![i, j] + +lemma VecFunc.of_sigma {f : (Fin k → V) → V} (h : 𝔖.VecFunc 𝚺 f) {Γ} : 𝔖.VecFunc Γ f := by + apply of_sigma_of_pi + · exact h + · have : 𝔖.VecPr 𝚷 fun v ↦ ∀ y, y = f (v ·.succ) → v 0 = y := all <| imply + (by simpa using retraction h (0 :> (·.succ.succ))) + (by simpa using equal' 𝚷 1 0) + exact of_iff this (fun v ↦ by simp) + +lemma conj {k l} {P : Fin l → (Fin k → V) → Prop} + (h : ∀ i, 𝔖.VecPr Γ fun w : Fin k → V ↦ P i w) : + 𝔖.VecPr Γ fun v : Fin k → V ↦ ∀ i, P i v := by + induction l + case zero => simp + case succ l ih => + suffices 𝔖.VecPr Γ fun v : Fin k → V ↦ P 0 v ∧ ∀ i : Fin l, P i.succ v by + apply of_iff this; intro x + constructor + · rintro ⟨h0, hs⟩ + intro i; cases' i using Fin.cases with i + · exact h0 + · exact hs i + · intro h + exact ⟨h 0, fun i ↦ h i.succ⟩ + apply and (h 0); apply ih + intro i; exact h i.succ + +lemma exVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} + (h : 𝔖.VecPr 𝚺 fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝔖.VecPr 𝚺 fun v : Fin k → V ↦ ∃ ys : Fin l → V, P v ys := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq] using h + case succ l ih => + suffices 𝔖.VecPr 𝚺 fun v : Fin k → V ↦ ∃ y, ∃ ys : Fin l → V, P v (y :> ys) by + apply of_iff this; intro x + constructor + · rintro ⟨y, ys, h⟩; exact ⟨_, h⟩ + · rintro ⟨ys, h⟩; exact ⟨ys 0, (ys ·.succ), by simpa using h⟩ + apply ex; apply ih + let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (retraction h g) (by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · congr 1; ext; simp [Matrix.vecAppend_eq_ite]) + +lemma allVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} + (h : 𝔖.VecPr 𝚷 fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝔖.VecPr 𝚷 fun v : Fin k → V ↦ ∀ ys : Fin l → V, P v ys := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq] using h + case succ l ih => + suffices 𝔖.VecPr 𝚷 fun v : Fin k → V ↦ ∀ y, ∀ ys : Fin l → V, P v (y :> ys) by + apply of_iff this; intro x + constructor + · intro h ys; simpa using h (ys 0) (ys ·.succ) + · intro h y ys; apply h + apply all; apply ih + let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (retraction h g) (by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · congr 1; ext; simp [Matrix.vecAppend_eq_ite]) + +private lemma substitution_sigma {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr 𝚺 P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : + 𝔖.VecPr 𝚺 fun z ↦ P (fun i ↦ f i z) := by + have : 𝔖.VecPr 𝚺 fun z ↦ ∃ ys : Fin k → V, (∀ i, ys i = f i z) ∧ P ys := by + apply exVec; apply and + · apply conj; intro i + simpa using retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact retraction hP (Fin.natAdd l) + exact of_iff this <| by + intro v + constructor + · rintro ⟨ys, hys, hP⟩ + have : ys = fun i ↦ f i v := funext hys + rcases this; exact hP + · intro hP + exact ⟨(f · v), by simp, hP⟩ + +private lemma substitution_pi {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr 𝚷 P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : + 𝔖.VecPr 𝚷 fun z ↦ P (fun i ↦ f i z) := by + have : 𝔖.VecPr 𝚷 fun z ↦ ∀ ys : Fin k → V, (∀ i, ys i = f i z) → P ys := by + apply allVec; apply imply + · apply conj; intro i + simpa using retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact retraction hP (Fin.natAdd l) + exact of_iff this <| by + intro v + constructor + · intro h; apply h _ (by simp) + · intro h ys e + have : ys = (f · v) := funext e + rcases this; exact h + +lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr Γ P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : 𝔖.VecPr Γ fun z ↦ P (fun i ↦ f i z) := + match Γ with + | 𝚺 => substitution_sigma hP hf + | 𝚷 => substitution_pi hP hf + | 𝚫 => of_sigma_of_pi (substitution_sigma (of_delta hP) hf) (substitution_pi (of_delta hP) hf) + +end VecPr + +namespace VecFunc + +variable {F : (Fin k → V) → V} + +open VecPr + +lemma nth (Γ) (i : Fin k) : 𝔖.VecFunc Γ fun w ↦ w i := VecPr.equal' Γ 0 i.succ + +lemma substitution {f : Fin k → (Fin l → V) → V} (hF : 𝔖.VecFunc Γ F) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : + 𝔖.VecFunc Γ fun z ↦ F (fun i ↦ f i z) := by + simp only [VecFunc, Nat.succ_eq_add_one] + simpa using VecPr.substitution (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF + (by intro i + cases' i using Fin.cases with i + · simpa using nth 𝚺 0 + · simpa using retraction (hf i) (0 :> (·.succ.succ))) + +end VecFunc + +end SigmaPiDeltaSystem + +/- +class EmbeddingType (V : outParam Type*) (β : Type*) where + embedding : β ↪ V + +namespace EmbeddingType + +instance : EmbeddingType V V := ⟨Function.Embedding.refl V⟩ + +instance (p : V → Prop) : EmbeddingType V (Subtype p) := ⟨Function.Embedding.subtype p⟩ + +end EmbeddingType + +namespace SigmaPiDeltaSystem + +class Class {V : Type*} (𝔖 : SigmaPiDeltaSystem V) (α : Type*) [EmbeddingType V α] where + delta : 𝔖.VecPr 𝚫 fun x : Fin 1 → V ↦ x 0 ∈ Set.range (EmbeddingType.embedding : α ↪ V) + +section Class + +instance (𝔖 : SigmaPiDeltaSystem V) : 𝔖.Class V where + delta := VecPr.of_iff (𝔖.verum' 𝚫 1) <| by intro v; simp; exact ⟨v 0, by rfl⟩ + +variable {𝔖 : SigmaPiDeltaSystem V} + +variable {α β γ δ ε ζ : Type*} + [EmbeddingType V α] [EmbeddingType V β] [EmbeddingType V γ] [EmbeddingType V δ] [EmbeddingType V ε] [EmbeddingType V ζ] + [𝔖.Class α] [𝔖.Class β] [𝔖.Class γ] [𝔖.Class δ] [𝔖.Class ε] [𝔖.Class ζ] + +def Pr₁ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → Prop) : Prop := 𝔖.VecPr Γ fun x : Fin 1 → V ↦ ∃ a : α, x 0 = EmbeddingType.embedding a ∧ P a +def Pr₂ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → Prop) : Prop := + 𝔖.VecPr Γ fun x : Fin 2 → V ↦ ∃ (a : α) (b : β), x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ P a b +def Pr₃ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → γ → Prop) : Prop := + 𝔖.VecPr Γ fun x : Fin 3 → V ↦ ∃ (a : α) (b : β) (c : γ), + x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ x 2 = EmbeddingType.embedding c ∧ P a b c +def Pr₄ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → γ → δ → Prop) : Prop := + 𝔖.VecPr Γ fun x : Fin 4 → V ↦ ∃ (a : α) (b : β) (c : γ) (d : δ), + x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ x 2 = EmbeddingType.embedding c ∧ P a b c d + + +end Class + +end SigmaPiDeltaSystem + +-/ + +end LO From 115d159b5f4a3cfbead216ffb86e0ccdef43be5c Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sat, 10 Aug 2024 16:39:47 +0900 Subject: [PATCH 03/10] add: `BooleanSystem` --- .../Definability/BooleanSystem.lean | 229 ++++++++++++++++++ .../Definability/SigmaPiDeltaSystem.lean | 98 +++++--- Arithmetization/Vorspiel/Graph.lean | 2 + Arithmetization/Vorspiel/Pair.lean | 23 ++ Arithmetization/Vorspiel/Vorspiel.lean | 59 +++++ 5 files changed, 371 insertions(+), 40 deletions(-) create mode 100644 Arithmetization/Definability/BooleanSystem.lean create mode 100644 Arithmetization/Vorspiel/Pair.lean diff --git a/Arithmetization/Definability/BooleanSystem.lean b/Arithmetization/Definability/BooleanSystem.lean new file mode 100644 index 0000000..be9ccad --- /dev/null +++ b/Arithmetization/Definability/BooleanSystem.lean @@ -0,0 +1,229 @@ +import Arithmetization.Vorspiel.Lemmata +import Arithmetization.Vorspiel.Graph +import Logic.FirstOrder.Arith.StrictHierarchy + +/-! + +# Arithmetical Formula Sorted by Arithmetical Hierarchy + +This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of first-order logic. + +- `𝚺-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚺-[m]`. +- `𝚷-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚷-[m]`. +- `𝚫-[m].Semiformula ξ n` is a pair of `𝚺-[m].Semiformula ξ n` and `𝚷-[m].Semiformula ξ n`. +- `ProperOn` : `p.ProperOn M` iff `p`'s two element `p.sigma` and `p.pi` are equivalent on model `M`. + +-/ + +namespace LO.Arith + +variable {V : Type*} + +structure BooleanSystem (V : Type*) where + VecPr : {k : ℕ} → ((Fin k → V) → Prop) → Prop + verum : VecPr fun _ : Fin k → V ↦ ⊤ + and {P Q : (Fin k → V) → Prop} : VecPr P → VecPr Q → VecPr fun v ↦ P v ∧ Q v + not {P : (Fin k → V) → Prop} : VecPr P → VecPr fun v ↦ ¬P v + equal : VecPr fun v : Fin 2 → V ↦ v 0 = v 1 + replace {k l} {P : (Fin k → V) → Prop} (hP : VecPr P) (f : Fin k → Fin l) : VecPr fun v ↦ P fun i ↦ v (f i) + +variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] + +variable (V) + +class BooleanSystem.BoundedQuantifier (𝔅 : BooleanSystem V) where + IsPolynomial : {k : ℕ} → ((Fin k → V) → V) → Prop + polynomial_comp {k} (F : (Fin k → V) → V) (fs : Fin k → (Fin l → V) → V) : + IsPolynomial F → (∀ i, IsPolynomial (fs i)) → IsPolynomial (fun v ↦ F (fun i ↦ fs i v)) + polynomial_replace {p : (Fin k → V) → V} (hp : IsPolynomial p) (f : Fin k → Fin l) : IsPolynomial (fun v ↦ p (fun i ↦ v (f i))) + polynomial_nth (i : Fin k) : IsPolynomial (· i) + polynomial_monotone {p : (Fin k → V) → V} (h : IsPolynomial p) {v w} : v ≤ w → p v ≤ p w + ball_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} : + IsPolynomial f → 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0)) → 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x + lessThan : 𝔅.VecPr fun v : Fin 2 → V ↦ v 0 < v 1 + +variable {V} + +namespace BooleanSystem + +variable {𝔅 : BooleanSystem V} {P Q : (Fin k → V) → Prop} + +lemma of_iff (hP : 𝔅.VecPr P) (h : ∀ x, P x ↔ Q x) : 𝔅.VecPr Q := by + have : P = Q := funext fun x ↦ by simp [h] + rcases this; exact hP + +lemma of_not (h : 𝔅.VecPr fun v ↦ ¬P v) : 𝔅.VecPr P := by simpa using 𝔅.not h + +lemma falsum : 𝔅.VecPr fun _ : Fin k → V ↦ ⊥ := of_not <| by simpa using 𝔅.verum + +@[simp] lemma constant (P : Prop) : 𝔅.VecPr fun _ : Fin k → V ↦ P := by + by_cases h : P <;> simp [h] + · exact 𝔅.verum + · exact 𝔅.falsum + +lemma or (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ∨ Q v := of_not <| by + simp; apply 𝔅.and + · apply 𝔅.not hP + · apply 𝔅.not hQ + +lemma imply (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v → Q v := by + simp only [imp_iff_not_or]; apply or + · apply 𝔅.not hP + · exact hQ + +lemma iff (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ↔ Q v := by + simp only [iff_iff_implies_and_implies] + exact 𝔅.and (imply hP hQ) (imply hQ hP) + +lemma conj {P : Fin l → (Fin k → V) → Prop} (hP : ∀ i, 𝔅.VecPr (P i)) : 𝔅.VecPr fun v ↦ ∀ i, P i v := by + induction l + case zero => simp + case succ l ih => + simp [forall_fin_iff_zero_and_forall_succ] + apply and + · exact hP 0 + · exact ih (fun i ↦ hP i.succ) + +lemma equal' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i = v j := by + simpa using 𝔅.replace 𝔅.equal ![i, j] + +section BoundedQuantifier + +open BoundedQuantifier + +variable [𝔅.BoundedQuantifier] + +variable (𝔅) + +def BoundedVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) ∧ ∃ p, IsPolynomial 𝔅 p ∧ f ≤ p + +variable {𝔅} + +lemma BoundedVecFunc.vecPr {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : 𝔅.VecPr (Function.Graphᵥ f) := h.1 + +lemma BoundedVecFunc.le_poly {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : ∃ p, IsPolynomial 𝔅 p ∧ f ≤ p := h.2 + +@[simp] lemma BoundedVecFunc.nth (i : Fin k) : 𝔅.BoundedVecFunc (· i) := by + constructor + · apply equal' + · exact ⟨(· i), polynomial_nth i, by simp⟩ + +lemma BoundedVecFunc.replace {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) (c : Fin k → Fin l) : + 𝔅.BoundedVecFunc (fun v ↦ f (fun i ↦ v (c i))) := by + constructor + · apply of_iff (𝔅.replace (l := l + 1) hf.vecPr (0 :> fun x ↦ (c x).succ)) <| by + intro v; simp [Function.Graphᵥ] + · rcases hf.le_poly with ⟨p, pp, hfp⟩ + refine ⟨fun v ↦ p (fun i ↦ v (c i)), by apply polynomial_replace pp, by intro v; simpa using hfp _⟩ + +lemma lessThan' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i < v j := by + simpa using 𝔅.replace lessThan ![i, j] + +lemma lessThanOrEq (i j : Fin k) : 𝔅.VecPr fun v ↦ v i ≤ v j := by + simp [le_iff_lt_or_eq] + apply or + · apply lessThan' + · apply equal' + +lemma BoundedQuantifier.bex_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (pf : IsPolynomial 𝔅 f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by + simp only [not_exists, not_and]; exact ball_poly pf (𝔅.not hP) + +lemma BoundedQuantifier.bex_vec_poly {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} + (pp : ∀ i, IsPolynomial 𝔅 (p i)) (hP : 𝔅.VecPr fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), P v w := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq (α := V)] using hP + case succ l ih => + simp only [exists_le_vec_iff_exists_le_exists_vec] + apply bex_poly (pp 0) + apply ih + · intro i; apply polynomial_replace (pp i.succ) + · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (𝔅.replace hP g) <| by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + +lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔅.VecPr P) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : + 𝔅.VecPr fun z ↦ P (fun i ↦ f i z) := by + choose p hp using fun i ↦ (hf i).le_poly + have : 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), (∀ i, w i = f i v) ∧ P w := by + apply BoundedQuantifier.bex_vec_poly (fun i ↦ (hp i).1) + apply and + · apply conj; intro i + simpa using 𝔅.replace (hf i).vecPr (i.natAdd l :> Fin.castAdd k) + · apply 𝔅.replace hP + apply of_iff this <| by + intro v; constructor + · rintro ⟨w, hw, e, h⟩ + rcases funext e + exact h + · intro h; exact ⟨(f · v), by intro i; simpa using (hp i).2 v, by simp, h⟩ + +lemma BoundedVecFunc.substitution {F : (Fin k → V) → V} {f : Fin k → (Fin l → V) → V} + (hF : 𝔅.BoundedVecFunc F) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : + 𝔅.BoundedVecFunc fun v ↦ F (fun i ↦ f i v) := by + constructor + · simpa [Function.Graphᵥ] using + BooleanSystem.substitution (l := l + 1) hF.vecPr (f := (· 0) :> fun i v ↦ f i (v ·.succ)) + (by intro i; cases' i using Fin.cases with i + · simp + · simpa using BoundedVecFunc.replace (hf i) _) + · rcases hF.le_poly with ⟨p, hp, hFp⟩ + choose ps hps using fun i ↦ (hf i).le_poly + refine ⟨fun v ↦ p fun i ↦ ps i v, polynomial_comp p ps hp (fun i ↦ (hps i).1), ?_⟩ + intro v; exact le_trans (hFp (f · v)) (polynomial_monotone hp (fun i ↦ (hps i).2 v)) + +lemma ball_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x := by + rcases hf.le_poly with ⟨p, hp, hfp⟩ + have : 𝔅.VecPr fun v ↦ ∀ x ≤ p v, x ≤ f v → P v x := by + apply ball_poly hp + apply imply + · simpa using substitution (𝔅.lessThanOrEq 0 1) + (f := ![(· 0), fun v ↦ f (v ·.succ)]) (by simpa using hf.replace Fin.succ) + · exact hP + exact of_iff this <| by + intro v; constructor + · intro h x hx + exact h x (le_trans hx (hfp v)) hx + · intro h x _ hx + exact h x hx + +lemma bex_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by + simp only [not_exists, not_and] + exact ball_le hf (𝔅.not hP) + +lemma ball_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∀ x < f v, P v x := by + have : 𝔅.VecPr fun v ↦ ∀ x ≤ f v, x < f v → P v x := by + apply ball_le hf + apply imply ?_ hP + simpa using substitution (𝔅.lessThan' 0 1) + (f := ![(· 0), fun v ↦ f (v ·.succ)]) (by simpa using hf.replace Fin.succ) + exact of_iff this <| by + intro v; constructor + · intro h x hx + exact h x (le_of_lt hx) hx + · intro h x _ hx + exact h x hx + +lemma bex_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x < f v, P v x := of_not <| by + simp only [not_exists, not_and] + exact ball_lt hf (𝔅.not hP) + +end BoundedQuantifier + +end BooleanSystem diff --git a/Arithmetization/Definability/SigmaPiDeltaSystem.lean b/Arithmetization/Definability/SigmaPiDeltaSystem.lean index b25c9e5..d5f64a9 100644 --- a/Arithmetization/Definability/SigmaPiDeltaSystem.lean +++ b/Arithmetization/Definability/SigmaPiDeltaSystem.lean @@ -14,22 +14,22 @@ This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of fi -/ -namespace LO +namespace LO.Arith class SigmaPiDeltaLike (Ω : Type*) [SigmaSymbol Ω] [PiSymbol Ω] [DeltaSymbol Ω] where alt : Ω → Ω variable {V : Type*} -class SigmaPiDeltaSystem (V : Type*) where +structure SigmaPiDeltaSystem (V : Type*) where VecPr : SigmaPiDelta → {k : ℕ} → ((Fin k → V) → Prop) → Prop vecPr_delta_iff_sigma_and_pi {k} {P : (Fin k → V) → Prop} : VecPr 𝚫 P ↔ VecPr 𝚺 P ∧ VecPr 𝚷 P - verum' (Γ k) : VecPr Γ fun _ : Fin k → V ↦ ⊤ - and' {k} {P Q : (Fin k → V) → Prop} : VecPr Γ P → VecPr Γ Q → VecPr Γ fun x ↦ P x ∧ Q x - not' {k} {P : (Fin k → V) → Prop} : VecPr Γ.alt P → VecPr Γ fun x ↦ ¬P x - all' {k} {P : (Fin k → V) → V → Prop} : VecPr 𝚷 (fun x ↦ P (x ·.succ) (x 0)) → VecPr 𝚷 fun x ↦ ∀ z, P x z - retraction' {k l} {P : (Fin k → V) → Prop} (hP : VecPr Γ P) (f : Fin k → Fin l) : VecPr Γ fun v ↦ P fun i ↦ v (f i) - equal' (Γ) : VecPr Γ fun v : Fin 2 → V ↦ v 0 = v 1 + verum (Γ k) : VecPr Γ fun _ : Fin k → V ↦ ⊤ + and {k} {P Q : (Fin k → V) → Prop} : VecPr Γ P → VecPr Γ Q → VecPr Γ fun x ↦ P x ∧ Q x + not {k} {P : (Fin k → V) → Prop} : VecPr Γ.alt P → VecPr Γ fun x ↦ ¬P x + all {k} {P : (Fin k → V) → V → Prop} : VecPr 𝚷 (fun x ↦ P (x ·.succ) (x 0)) → VecPr 𝚷 fun x ↦ ∀ z, P x z + retraction {k l} {P : (Fin k → V) → Prop} (hP : VecPr Γ P) (f : Fin k → Fin l) : VecPr Γ fun v ↦ P fun i ↦ v (f i) + equal (Γ) : VecPr Γ fun v : Fin 2 → V ↦ v 0 = v 1 abbrev SigmaPiDeltaSystem.VecFunc (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (f : (Fin k → V) → V) : Prop := 𝔖.VecPr Γ fun v ↦ v 0 = f (v ·.succ) @@ -38,19 +38,28 @@ namespace SigmaPiDeltaSystem variable {𝔖 : SigmaPiDeltaSystem V} {Γ : SigmaPiDelta} {k} {P Q : (Fin k → V) → Prop} -namespace VecPr +/- +variable (𝔖 Γ) + +abbrev Pred (P : V → Prop) : Prop := 𝔖.VecPr Γ (k := 1) (fun v ↦ P (v 0)) + +abbrev Rel (P : V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 2) (fun v ↦ P (v 0) (v 1)) -alias verum := verum' +abbrev Rel₃ (P : V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) -alias and := and' +abbrev Rel₄ (P : V → V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) -alias not := not' +abbrev Rel₅ (P : V → V → V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) -alias all := all' +abbrev Function (f : V → V) : Prop := 𝔖.VecFunc Γ (k := 1) (fun v ↦ f (v 0)) -alias retraction := retraction' +abbrev Function₂ (f : V → V → V) : Prop := 𝔖.VecFunc Γ (k := 2) (fun v ↦ f (v 0) (v 1)) -alias equal := equal' +abbrev Function₃ (f : V → V → V → V) : Prop := 𝔖.VecFunc Γ (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) + +abbrev Function₄ (f : V → V → V → V → V) : Prop := 𝔖.VecFunc Γ (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) + +-/ lemma of_iff (hP : 𝔖.VecPr Γ P) (h : ∀ x, P x ↔ Q x) : 𝔖.VecPr Γ Q := by have : P = Q := funext <| by simpa @@ -61,19 +70,19 @@ lemma of_sigma_of_pi (hσ : 𝔖.VecPr 𝚺 P) (hπ : 𝔖.VecPr 𝚷 P) : 𝔖. match Γ with | 𝚺 => hσ | 𝚷 => hπ - | 𝚫 => vecPr_delta_iff_sigma_and_pi.mpr ⟨hσ, hπ⟩ + | 𝚫 => (vecPr_delta_iff_sigma_and_pi _).mpr ⟨hσ, hπ⟩ lemma of_delta (h : 𝔖.VecPr 𝚫 P) {Γ} : 𝔖.VecPr Γ P := of_sigma_of_pi - (vecPr_delta_iff_sigma_and_pi.mp h |>.1) - (vecPr_delta_iff_sigma_and_pi.mp h |>.2) + ((vecPr_delta_iff_sigma_and_pi _).mp h |>.1) + ((vecPr_delta_iff_sigma_and_pi _).mp h |>.2) -lemma not' (h : 𝔖.VecPr Γ P) : 𝔖.VecPr Γ.alt fun x ↦ ¬P x := not (by simpa) +lemma not' (h : 𝔖.VecPr Γ P) : 𝔖.VecPr Γ.alt fun x ↦ ¬P x := 𝔖.not (by simpa) lemma of_not (h : 𝔖.VecPr Γ.alt (fun x ↦ ¬P x)) : 𝔖.VecPr Γ P := by simpa using not' h lemma falsum (Γ : SigmaPiDelta) (k : ℕ) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ ⊥ := - of_sigma_of_pi (by simpa using not' (verum 𝚷 k)) (by simpa using not' (verum 𝚺 k)) + of_sigma_of_pi (by simpa using not' (𝔖.verum 𝚷 k)) (by simpa using not' (𝔖.verum 𝚺 k)) @[simp] lemma constant (Γ : SigmaPiDelta) (k : ℕ) (P : Prop) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ P := by by_cases h : P <;> simp [h] @@ -87,12 +96,12 @@ lemma or (hP : 𝔖.VecPr Γ P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : F · apply not' hQ lemma imply (hP : 𝔖.VecPr Γ.alt P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x → Q x := by - simp [imp_iff_not_or]; apply or - · apply not hP + simp only [imp_iff_not_or]; apply or + · apply 𝔖.not hP · exact hQ lemma ex {k} {P : (Fin k → V) → V → Prop} (h : 𝔖.VecPr 𝚺 fun x ↦ P (x ·.succ) (x 0)) : 𝔖.VecPr 𝚺 fun x ↦ ∃ z, P x z := of_not <| by - simpa using all (by apply not' h) + simpa using 𝔖.all (by apply not' h) lemma iff (hP : 𝔖.VecPr 𝚫 P) (hQ : 𝔖.VecPr 𝚫 Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x ↔ Q x := of_delta <| by simp only [iff_iff_implies_and_implies] @@ -101,13 +110,13 @@ lemma iff (hP : 𝔖.VecPr 𝚫 P) (hQ : 𝔖.VecPr 𝚫 Q) : 𝔖.VecPr Γ fun · exact imply hQ hP lemma equal' (Γ) (i j : Fin k) : 𝔖.VecPr Γ fun v ↦ v i = v j := by - simpa using retraction (equal Γ) ![i, j] + simpa using 𝔖.retraction (𝔖.equal Γ) ![i, j] lemma VecFunc.of_sigma {f : (Fin k → V) → V} (h : 𝔖.VecFunc 𝚺 f) {Γ} : 𝔖.VecFunc Γ f := by apply of_sigma_of_pi · exact h - · have : 𝔖.VecPr 𝚷 fun v ↦ ∀ y, y = f (v ·.succ) → v 0 = y := all <| imply - (by simpa using retraction h (0 :> (·.succ.succ))) + · have : 𝔖.VecPr 𝚷 fun v ↦ ∀ y, y = f (v ·.succ) → v 0 = y := 𝔖.all <| imply + (by simpa using 𝔖.retraction h (0 :> (·.succ.succ))) (by simpa using equal' 𝚷 1 0) exact of_iff this (fun v ↦ by simp) @@ -126,7 +135,7 @@ lemma conj {k l} {P : Fin l → (Fin k → V) → Prop} · exact hs i · intro h exact ⟨h 0, fun i ↦ h i.succ⟩ - apply and (h 0); apply ih + apply 𝔖.and (h 0); apply ih intro i; exact h i.succ lemma exVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} @@ -142,7 +151,7 @@ lemma exVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} · rintro ⟨ys, h⟩; exact ⟨ys 0, (ys ·.succ), by simpa using h⟩ apply ex; apply ih let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (retraction h g) (by + exact of_iff (𝔖.retraction h g) (by intro v; simp [g] apply iff_of_eq; congr · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] @@ -164,7 +173,7 @@ lemma allVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} · intro h y ys; apply h apply all; apply ih let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (retraction h g) (by + exact of_iff (𝔖.retraction h g) (by intro v; simp [g] apply iff_of_eq; congr · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] @@ -178,8 +187,8 @@ private lemma substitution_sigma {f : Fin k → (Fin l → V) → V} (hP : 𝔖. have : 𝔖.VecPr 𝚺 fun z ↦ ∃ ys : Fin k → V, (∀ i, ys i = f i z) ∧ P ys := by apply exVec; apply and · apply conj; intro i - simpa using retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) - · exact retraction hP (Fin.natAdd l) + simpa using 𝔖.retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact 𝔖.retraction hP (Fin.natAdd l) exact of_iff this <| by intro v constructor @@ -194,8 +203,8 @@ private lemma substitution_pi {f : Fin k → (Fin l → V) → V} (hP : 𝔖.Vec have : 𝔖.VecPr 𝚷 fun z ↦ ∀ ys : Fin k → V, (∀ i, ys i = f i z) → P ys := by apply allVec; apply imply · apply conj; intro i - simpa using retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) - · exact retraction hP (Fin.natAdd l) + simpa using 𝔖.retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact 𝔖.retraction hP (Fin.natAdd l) exact of_iff this <| by intro v constructor @@ -210,27 +219,36 @@ lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr Γ P) (h | 𝚷 => substitution_pi hP hf | 𝚫 => of_sigma_of_pi (substitution_sigma (of_delta hP) hf) (substitution_pi (of_delta hP) hf) -end VecPr - namespace VecFunc variable {F : (Fin k → V) → V} -open VecPr - -lemma nth (Γ) (i : Fin k) : 𝔖.VecFunc Γ fun w ↦ w i := VecPr.equal' Γ 0 i.succ +lemma nth (Γ) (i : Fin k) : 𝔖.VecFunc Γ fun w ↦ w i := 𝔖.equal' Γ 0 i.succ lemma substitution {f : Fin k → (Fin l → V) → V} (hF : 𝔖.VecFunc Γ F) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : 𝔖.VecFunc Γ fun z ↦ F (fun i ↦ f i z) := by simp only [VecFunc, Nat.succ_eq_add_one] - simpa using VecPr.substitution (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF + simpa using 𝔖.substitution (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF (by intro i cases' i using Fin.cases with i · simpa using nth 𝚺 0 - · simpa using retraction (hf i) (0 :> (·.succ.succ))) + · simpa using 𝔖.retraction (hf i) (0 :> (·.succ.succ))) end VecFunc +variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] + +class Arithmetical (𝔖 : SigmaPiDeltaSystem V) where + zero' (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 0 + one' (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 1 + add' (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 + v 1 + mul' (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 * v 1 + lt' (Γ) : 𝔖.VecPr Γ fun v : Fin 2 → V ↦ v 0 < v 1 + ball' {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∀ x ≤ v 0, P (v ·.succ) x + bex' {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∃ x ≤ v 0, P (v ·.succ) x + + + end SigmaPiDeltaSystem /- diff --git a/Arithmetization/Vorspiel/Graph.lean b/Arithmetization/Vorspiel/Graph.lean index d0746c5..f095285 100644 --- a/Arithmetization/Vorspiel/Graph.lean +++ b/Arithmetization/Vorspiel/Graph.lean @@ -4,6 +4,8 @@ namespace Function variable {σ α β : Sort*} +def Graphᵥ (f : (Fin k → α) → α) : (Fin (k + 1) → α) → Prop := fun v ↦ v 0 = f (v ·.succ) + def Graph (f : α → σ) : σ → α → Prop := fun y x ↦ y = f x def Graph₂ (f : α → β → σ) : σ → α → β → Prop := fun y x₁ x₂ ↦ y = f x₁ x₂ diff --git a/Arithmetization/Vorspiel/Pair.lean b/Arithmetization/Vorspiel/Pair.lean new file mode 100644 index 0000000..2de04d8 --- /dev/null +++ b/Arithmetization/Vorspiel/Pair.lean @@ -0,0 +1,23 @@ +import Arithmetization.Vorspiel.Vorspiel + +class Pair (α : Type*) where + pair : α → α → α + right : α → α + left : α → α + left_pair (x y : α) : left (pair x y) = x + right_pair (x y : α) : right (pair x y) = y + pair_left_right (x y : α) : pair (left x) (right x) = x + +attribute [simp] Pair.left_pair Pair.right_pair Pair.pair_left_right + +/-- `!⟪x, y, z, ...⟫` notation for `Seq` -/ +syntax "⟪" term,* "⟫" : term + +macro_rules + | `(⟪$term:term, $terms:term,*⟫) => `(Pair.pair $term ⟪$terms,*⟫) + | `(⟪$term:term⟫) => `($term) + +@[app_unexpander Pair.pair] +def pairUnexpander : Lean.PrettyPrinter.Unexpander + | `($_ $term $term2) => `(⟪$term, $term2⟫) + | _ => throw () diff --git a/Arithmetization/Vorspiel/Vorspiel.lean b/Arithmetization/Vorspiel/Vorspiel.lean index 3bf06e7..21039cc 100644 --- a/Arithmetization/Vorspiel/Vorspiel.lean +++ b/Arithmetization/Vorspiel/Vorspiel.lean @@ -52,6 +52,65 @@ lemma eq_vecCons' (s : Fin (n + 1) → C) : s 0 :> (s ·.succ) = s := end Matrix +lemma forall_fin_iff_zero_and_forall_succ {P : Fin (k + 1) → Prop} : (∀ i, P i) ↔ P 0 ∧ ∀ i : Fin k, P i.succ := + ⟨fun h ↦ ⟨h 0, fun i ↦ h i.succ⟩, by + rintro ⟨hz, hs⟩ i + cases' i using Fin.cases with i + · exact hz + · exact hs i⟩ + +lemma exists_fin_iff_zero_or_exists_succ {P : Fin (k + 1) → Prop} : (∃ i, P i) ↔ P 0 ∨ ∃ i : Fin k, P i.succ := + ⟨by rintro ⟨i, hi⟩ + cases i using Fin.cases + · left; exact hi + · right; exact ⟨_, hi⟩, + by rintro (hz | ⟨i, h⟩) + · exact ⟨0, hz⟩ + · exact ⟨_, h⟩⟩ + +lemma forall_vec_iff_forall_forall_vec {P : (Fin (k + 1) → α) → Prop} : + (∀ v : Fin (k + 1) → α, P v) ↔ ∀ x, ∀ v : Fin k → α, P (x :> v) := by + constructor + · intro h x v; exact h _ + · intro h v; simpa using h (v 0) (v ·.succ) + +lemma exists_vec_iff_exists_exists_vec {P : (Fin (k + 1) → α) → Prop} : + (∃ v : Fin (k + 1) → α, P v) ↔ ∃ x, ∃ v : Fin k → α, P (x :> v) := by + constructor + · rintro ⟨v, h⟩; exact ⟨v 0, (v ·.succ), by simpa using h⟩ + · rintro ⟨x, v, h⟩; exact ⟨_, h⟩ + +lemma exists_le_vec_iff_exists_le_exists_vec [LE α] {P : (Fin (k + 1) → α) → Prop} {f : Fin (k + 1) → α} : + (∃ v ≤ f, P v) ↔ ∃ x ≤ f 0, ∃ v ≤ (f ·.succ), P (x :> v) := by + constructor + · rintro ⟨w, hw, h⟩ + exact ⟨w 0, hw 0, (w ·.succ), fun i ↦ hw i.succ, by simpa using h⟩ + · rintro ⟨x, hx, v, hv, h⟩ + refine ⟨x :> v, ?_, h⟩ + intro i; cases' i using Fin.cases with i + · exact hx + · exact hv i + +lemma forall_le_vec_iff_forall_le_forall_vec [LE α] {P : (Fin (k + 1) → α) → Prop} {f : Fin (k + 1) → α} : + (∀ v ≤ f, P v) ↔ ∀ x ≤ f 0, ∀ v ≤ (f ·.succ), P (x :> v) := by + constructor + · intro h x hx v hv + refine h (x :> v) ?_ + intro i; cases' i using Fin.cases with i + · exact hx + · exact hv i + · intro h v hv + simpa using h (v 0) (hv 0) (v ·.succ) (hv ·.succ) + +@[simp] lemma forall_fin_two_iff {P : Fin 2 → Prop} : (∀ i, P i) ↔ P 0 ∧ P 1 := by + constructor + · intro h; exact ⟨h 0, h 1⟩ + · rintro ⟨h0, h1⟩ + intro i + match i with + | 0 => exact h0 + | 1 => exact h1 + instance : ToString Empty := ⟨Empty.elim⟩ class Hash (α : Type*) where From 2ee3fecdc204f3da8dd0f04b54cab2f34bc796d4 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sun, 11 Aug 2024 00:30:09 +0900 Subject: [PATCH 04/10] wip --- Arithmetization/Definability/Basic.lean | 29 ++ .../Definability/BooleanSystem.lean | 229 ---------- .../Definability/BooleanSystem/Basic.lean | 417 ++++++++++++++++++ .../Definability/BooleanSystem/DeltaZero.lean | 104 +++++ .../Definability/BooleanSystem/Open.lean | 37 ++ .../Definability/SigmaPiDeltaSystem.lean | 12 +- Arithmetization/Vorspiel/Vorspiel.lean | 9 - 7 files changed, 593 insertions(+), 244 deletions(-) create mode 100644 Arithmetization/Definability/Basic.lean delete mode 100644 Arithmetization/Definability/BooleanSystem.lean create mode 100644 Arithmetization/Definability/BooleanSystem/Basic.lean create mode 100644 Arithmetization/Definability/BooleanSystem/DeltaZero.lean create mode 100644 Arithmetization/Definability/BooleanSystem/Open.lean diff --git a/Arithmetization/Definability/Basic.lean b/Arithmetization/Definability/Basic.lean new file mode 100644 index 0000000..6f1a088 --- /dev/null +++ b/Arithmetization/Definability/Basic.lean @@ -0,0 +1,29 @@ +import Arithmetization.Vorspiel.Lemmata +import Arithmetization.Vorspiel.Graph +import Logic.FirstOrder.Arith.StrictHierarchy + +namespace LO.FirstOrder + +variable {L : Language} {V : Type*} [Structure L V] + +def Defined {k} (R : (Fin k → V) → Prop) (p : Semisentence L k) : Prop := + ∀ v, R v ↔ Semiformula.Evalbm V v p + +def DefinedWithParam {k} (R : (Fin k → V) → Prop) (p : Semiformula L V k) : Prop := + ∀ v, R v ↔ Semiformula.Evalm V v id p + +lemma Defined.iff {k} {R : (Fin k → V) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : + Semiformula.Evalbm V v p ↔ R v := (h v).symm + +lemma DefinedWithParam.iff {k} {R : (Fin k → V) → Prop} {p : Semiformula L V k} (h : DefinedWithParam R p) (v) : + Semiformula.Evalm V v id p ↔ R v := (h v).symm + +variable (L) + +def Definable {k} (C : Semisentence L k → Prop) (R : (Fin k → V) → Prop) : Prop := + ∃ p, C p ∧ Defined R p + +def DefinableWithParam {k} (C : Semiformula L V k → Prop) (R : (Fin k → V) → Prop) : Prop := + ∃ p, C p ∧ DefinedWithParam R p + +end LO.FirstOrder diff --git a/Arithmetization/Definability/BooleanSystem.lean b/Arithmetization/Definability/BooleanSystem.lean deleted file mode 100644 index be9ccad..0000000 --- a/Arithmetization/Definability/BooleanSystem.lean +++ /dev/null @@ -1,229 +0,0 @@ -import Arithmetization.Vorspiel.Lemmata -import Arithmetization.Vorspiel.Graph -import Logic.FirstOrder.Arith.StrictHierarchy - -/-! - -# Arithmetical Formula Sorted by Arithmetical Hierarchy - -This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of first-order logic. - -- `𝚺-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚺-[m]`. -- `𝚷-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚷-[m]`. -- `𝚫-[m].Semiformula ξ n` is a pair of `𝚺-[m].Semiformula ξ n` and `𝚷-[m].Semiformula ξ n`. -- `ProperOn` : `p.ProperOn M` iff `p`'s two element `p.sigma` and `p.pi` are equivalent on model `M`. - --/ - -namespace LO.Arith - -variable {V : Type*} - -structure BooleanSystem (V : Type*) where - VecPr : {k : ℕ} → ((Fin k → V) → Prop) → Prop - verum : VecPr fun _ : Fin k → V ↦ ⊤ - and {P Q : (Fin k → V) → Prop} : VecPr P → VecPr Q → VecPr fun v ↦ P v ∧ Q v - not {P : (Fin k → V) → Prop} : VecPr P → VecPr fun v ↦ ¬P v - equal : VecPr fun v : Fin 2 → V ↦ v 0 = v 1 - replace {k l} {P : (Fin k → V) → Prop} (hP : VecPr P) (f : Fin k → Fin l) : VecPr fun v ↦ P fun i ↦ v (f i) - -variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] - -variable (V) - -class BooleanSystem.BoundedQuantifier (𝔅 : BooleanSystem V) where - IsPolynomial : {k : ℕ} → ((Fin k → V) → V) → Prop - polynomial_comp {k} (F : (Fin k → V) → V) (fs : Fin k → (Fin l → V) → V) : - IsPolynomial F → (∀ i, IsPolynomial (fs i)) → IsPolynomial (fun v ↦ F (fun i ↦ fs i v)) - polynomial_replace {p : (Fin k → V) → V} (hp : IsPolynomial p) (f : Fin k → Fin l) : IsPolynomial (fun v ↦ p (fun i ↦ v (f i))) - polynomial_nth (i : Fin k) : IsPolynomial (· i) - polynomial_monotone {p : (Fin k → V) → V} (h : IsPolynomial p) {v w} : v ≤ w → p v ≤ p w - ball_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} : - IsPolynomial f → 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0)) → 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x - lessThan : 𝔅.VecPr fun v : Fin 2 → V ↦ v 0 < v 1 - -variable {V} - -namespace BooleanSystem - -variable {𝔅 : BooleanSystem V} {P Q : (Fin k → V) → Prop} - -lemma of_iff (hP : 𝔅.VecPr P) (h : ∀ x, P x ↔ Q x) : 𝔅.VecPr Q := by - have : P = Q := funext fun x ↦ by simp [h] - rcases this; exact hP - -lemma of_not (h : 𝔅.VecPr fun v ↦ ¬P v) : 𝔅.VecPr P := by simpa using 𝔅.not h - -lemma falsum : 𝔅.VecPr fun _ : Fin k → V ↦ ⊥ := of_not <| by simpa using 𝔅.verum - -@[simp] lemma constant (P : Prop) : 𝔅.VecPr fun _ : Fin k → V ↦ P := by - by_cases h : P <;> simp [h] - · exact 𝔅.verum - · exact 𝔅.falsum - -lemma or (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ∨ Q v := of_not <| by - simp; apply 𝔅.and - · apply 𝔅.not hP - · apply 𝔅.not hQ - -lemma imply (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v → Q v := by - simp only [imp_iff_not_or]; apply or - · apply 𝔅.not hP - · exact hQ - -lemma iff (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ↔ Q v := by - simp only [iff_iff_implies_and_implies] - exact 𝔅.and (imply hP hQ) (imply hQ hP) - -lemma conj {P : Fin l → (Fin k → V) → Prop} (hP : ∀ i, 𝔅.VecPr (P i)) : 𝔅.VecPr fun v ↦ ∀ i, P i v := by - induction l - case zero => simp - case succ l ih => - simp [forall_fin_iff_zero_and_forall_succ] - apply and - · exact hP 0 - · exact ih (fun i ↦ hP i.succ) - -lemma equal' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i = v j := by - simpa using 𝔅.replace 𝔅.equal ![i, j] - -section BoundedQuantifier - -open BoundedQuantifier - -variable [𝔅.BoundedQuantifier] - -variable (𝔅) - -def BoundedVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) ∧ ∃ p, IsPolynomial 𝔅 p ∧ f ≤ p - -variable {𝔅} - -lemma BoundedVecFunc.vecPr {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : 𝔅.VecPr (Function.Graphᵥ f) := h.1 - -lemma BoundedVecFunc.le_poly {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : ∃ p, IsPolynomial 𝔅 p ∧ f ≤ p := h.2 - -@[simp] lemma BoundedVecFunc.nth (i : Fin k) : 𝔅.BoundedVecFunc (· i) := by - constructor - · apply equal' - · exact ⟨(· i), polynomial_nth i, by simp⟩ - -lemma BoundedVecFunc.replace {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) (c : Fin k → Fin l) : - 𝔅.BoundedVecFunc (fun v ↦ f (fun i ↦ v (c i))) := by - constructor - · apply of_iff (𝔅.replace (l := l + 1) hf.vecPr (0 :> fun x ↦ (c x).succ)) <| by - intro v; simp [Function.Graphᵥ] - · rcases hf.le_poly with ⟨p, pp, hfp⟩ - refine ⟨fun v ↦ p (fun i ↦ v (c i)), by apply polynomial_replace pp, by intro v; simpa using hfp _⟩ - -lemma lessThan' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i < v j := by - simpa using 𝔅.replace lessThan ![i, j] - -lemma lessThanOrEq (i j : Fin k) : 𝔅.VecPr fun v ↦ v i ≤ v j := by - simp [le_iff_lt_or_eq] - apply or - · apply lessThan' - · apply equal' - -lemma BoundedQuantifier.bex_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (pf : IsPolynomial 𝔅 f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by - simp only [not_exists, not_and]; exact ball_poly pf (𝔅.not hP) - -lemma BoundedQuantifier.bex_vec_poly {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} - (pp : ∀ i, IsPolynomial 𝔅 (p i)) (hP : 𝔅.VecPr fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : - 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), P v w := by - induction l generalizing k - case zero => simpa [Matrix.empty_eq (α := V)] using hP - case succ l ih => - simp only [exists_le_vec_iff_exists_le_exists_vec] - apply bex_poly (pp 0) - apply ih - · intro i; apply polynomial_replace (pp i.succ) - · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (𝔅.replace hP g) <| by - intro v; simp [g] - apply iff_of_eq; congr - · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · ext i - cases' i using Fin.cases with i - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - -lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔅.VecPr P) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : - 𝔅.VecPr fun z ↦ P (fun i ↦ f i z) := by - choose p hp using fun i ↦ (hf i).le_poly - have : 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), (∀ i, w i = f i v) ∧ P w := by - apply BoundedQuantifier.bex_vec_poly (fun i ↦ (hp i).1) - apply and - · apply conj; intro i - simpa using 𝔅.replace (hf i).vecPr (i.natAdd l :> Fin.castAdd k) - · apply 𝔅.replace hP - apply of_iff this <| by - intro v; constructor - · rintro ⟨w, hw, e, h⟩ - rcases funext e - exact h - · intro h; exact ⟨(f · v), by intro i; simpa using (hp i).2 v, by simp, h⟩ - -lemma BoundedVecFunc.substitution {F : (Fin k → V) → V} {f : Fin k → (Fin l → V) → V} - (hF : 𝔅.BoundedVecFunc F) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : - 𝔅.BoundedVecFunc fun v ↦ F (fun i ↦ f i v) := by - constructor - · simpa [Function.Graphᵥ] using - BooleanSystem.substitution (l := l + 1) hF.vecPr (f := (· 0) :> fun i v ↦ f i (v ·.succ)) - (by intro i; cases' i using Fin.cases with i - · simp - · simpa using BoundedVecFunc.replace (hf i) _) - · rcases hF.le_poly with ⟨p, hp, hFp⟩ - choose ps hps using fun i ↦ (hf i).le_poly - refine ⟨fun v ↦ p fun i ↦ ps i v, polynomial_comp p ps hp (fun i ↦ (hps i).1), ?_⟩ - intro v; exact le_trans (hFp (f · v)) (polynomial_monotone hp (fun i ↦ (hps i).2 v)) - -lemma ball_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x := by - rcases hf.le_poly with ⟨p, hp, hfp⟩ - have : 𝔅.VecPr fun v ↦ ∀ x ≤ p v, x ≤ f v → P v x := by - apply ball_poly hp - apply imply - · simpa using substitution (𝔅.lessThanOrEq 0 1) - (f := ![(· 0), fun v ↦ f (v ·.succ)]) (by simpa using hf.replace Fin.succ) - · exact hP - exact of_iff this <| by - intro v; constructor - · intro h x hx - exact h x (le_trans hx (hfp v)) hx - · intro h x _ hx - exact h x hx - -lemma bex_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by - simp only [not_exists, not_and] - exact ball_le hf (𝔅.not hP) - -lemma ball_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∀ x < f v, P v x := by - have : 𝔅.VecPr fun v ↦ ∀ x ≤ f v, x < f v → P v x := by - apply ball_le hf - apply imply ?_ hP - simpa using substitution (𝔅.lessThan' 0 1) - (f := ![(· 0), fun v ↦ f (v ·.succ)]) (by simpa using hf.replace Fin.succ) - exact of_iff this <| by - intro v; constructor - · intro h x hx - exact h x (le_of_lt hx) hx - · intro h x _ hx - exact h x hx - -lemma bex_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x < f v, P v x := of_not <| by - simp only [not_exists, not_and] - exact ball_lt hf (𝔅.not hP) - -end BoundedQuantifier - -end BooleanSystem diff --git a/Arithmetization/Definability/BooleanSystem/Basic.lean b/Arithmetization/Definability/BooleanSystem/Basic.lean new file mode 100644 index 0000000..96ff21c --- /dev/null +++ b/Arithmetization/Definability/BooleanSystem/Basic.lean @@ -0,0 +1,417 @@ +import Arithmetization.Definability.Basic +import Arithmetization.Definability.Init + +/-! + +# Boolean System + + +-/ + +namespace LO + +variable {V : Type*} + +structure BooleanSystem (V : Type*) where + VecPr : {k : ℕ} → ((Fin k → V) → Prop) → Prop + verum : VecPr fun _ : Fin k → V ↦ ⊤ + and {P Q : (Fin k → V) → Prop} : VecPr P → VecPr Q → VecPr fun v ↦ P v ∧ Q v + not {P : (Fin k → V) → Prop} : VecPr P → VecPr fun v ↦ ¬P v + equal : VecPr fun v : Fin 2 → V ↦ v 0 = v 1 + replace {k l} {P : (Fin k → V) → Prop} (hP : VecPr P) (f : Fin k → Fin l) : VecPr fun v ↦ P fun i ↦ v (f i) + +namespace BooleanSystem + +variable {𝔅 : BooleanSystem V} {P Q : (Fin k → V) → Prop} + +lemma of_iff (hP : 𝔅.VecPr P) (h : ∀ x, P x ↔ Q x) : 𝔅.VecPr Q := by + have : P = Q := funext fun x ↦ by simp [h] + rcases this; exact hP + +lemma of_not (h : 𝔅.VecPr fun v ↦ ¬P v) : 𝔅.VecPr P := by simpa using 𝔅.not h + +lemma falsum : 𝔅.VecPr fun _ : Fin k → V ↦ ⊥ := of_not <| by simpa using 𝔅.verum + +@[simp] lemma constant (P : Prop) : 𝔅.VecPr fun _ : Fin k → V ↦ P := by + by_cases h : P <;> simp [h] + · exact 𝔅.verum + · exact 𝔅.falsum + +lemma or (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ∨ Q v := of_not <| by + simp; apply 𝔅.and + · apply 𝔅.not hP + · apply 𝔅.not hQ + +lemma imply (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v → Q v := by + simp only [imp_iff_not_or]; apply or + · apply 𝔅.not hP + · exact hQ + +lemma iff (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ↔ Q v := by + simp only [iff_iff_implies_and_implies] + exact 𝔅.and (imply hP hQ) (imply hQ hP) + +lemma conj {P : Fin l → (Fin k → V) → Prop} (hP : ∀ i, 𝔅.VecPr (P i)) : 𝔅.VecPr fun v ↦ ∀ i, P i v := by + induction l + case zero => simp + case succ l ih => + simp [forall_fin_iff_zero_and_forall_succ] + apply and + · exact hP 0 + · exact ih (fun i ↦ hP i.succ) + +lemma equal' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i = v j := by + simpa using 𝔅.replace 𝔅.equal ![i, j] + +variable (𝔅) + +class Quantifer where + all {k} {P : (Fin k → V) → V → Prop} : 𝔅.VecPr (fun x ↦ P (x ·.succ) (x 0)) → 𝔅.VecPr fun x ↦ ∀ z, P x z + ex {k} {P : (Fin k → V) → V → Prop} : 𝔅.VecPr (fun x ↦ P (x ·.succ) (x 0)) → 𝔅.VecPr fun x ↦ ∃ z, P x z + +variable {𝔅} + +end BooleanSystem + +namespace Arith + +variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] + +variable (V) + +structure BoundedSystem extends BooleanSystem V where + Polynomial : {k : ℕ} → ((Fin k → V) → V) → Prop + polynomial_comp {k} (F : (Fin k → V) → V) (fs : Fin k → (Fin l → V) → V) : + Polynomial F → (∀ i, Polynomial (fs i)) → Polynomial (fun v ↦ F (fun i ↦ fs i v)) + polynomial_replace {p : (Fin k → V) → V} (hp : Polynomial p) (f : Fin k → Fin l) : Polynomial (fun v ↦ p (fun i ↦ v (f i))) + polynomial_nth (i : Fin k) : Polynomial (· i) + polynomial_monotone {p : (Fin k → V) → V} (h : Polynomial p) {v w} : v ≤ w → p v ≤ p w + ball_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} : + Polynomial f → VecPr (fun v ↦ P (v ·.succ) (v 0)) → VecPr fun v ↦ ∀ x ≤ f v, P v x + lessThan : VecPr fun v : Fin 2 → V ↦ v 0 < v 1 + +variable {V} + +namespace BoundedSystem + +open LO.BooleanSystem + +variable {𝔅 : Arith.BoundedSystem V} {P Q : (Fin k → V) → Prop} + +variable (𝔅) + +abbrev Pred (P : V → Prop) : Prop := 𝔅.VecPr (k := 1) fun v ↦ P (v 0) + +abbrev Rel (R : V → V → Prop) : Prop := 𝔅.VecPr (k := 2) fun v ↦ R (v 0) (v 1) + +abbrev Rel₃ (R : V → V → V → Prop) : Prop := 𝔅.VecPr (k := 3) fun v ↦ R (v 0) (v 1) (v 2) + +abbrev Rel₄ (R : V → V → V → V → Prop) : Prop := 𝔅.VecPr (k := 4) fun v ↦ R (v 0) (v 1) (v 2) (v 3) + +abbrev UBVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) + +abbrev UBConstant (c : V) : Prop := 𝔅.UBVecFunc (k := 0) fun _ ↦ c + +abbrev UBFunction (f : V → V) : Prop := 𝔅.UBVecFunc (k := 1) fun v ↦ f (v 0) + +abbrev UBFunction₂ (f : V → V → V) : Prop := 𝔅.UBVecFunc (k := 2) fun v ↦ f (v 0) (v 1) + +abbrev UBFunction₃ (f : V → V → V → V) : Prop := 𝔅.UBVecFunc (k := 3) fun v ↦ f (v 0) (v 1) (v 2) + +def BoundedVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) ∧ ∃ p, 𝔅.Polynomial p ∧ f ≤ p + +abbrev BoundedConstant (c : V) : Prop := 𝔅.BoundedVecFunc (k := 0) fun _ ↦ c + +abbrev BoundedFunction (f : V → V) : Prop := 𝔅.BoundedVecFunc (k := 1) fun v ↦ f (v 0) + +abbrev BoundedFunction₂ (f : V → V → V) : Prop := 𝔅.BoundedVecFunc (k := 2) fun v ↦ f (v 0) (v 1) + +abbrev BoundedFunction₃ (f : V → V → V → V) : Prop := 𝔅.BoundedVecFunc (k := 3) fun v ↦ f (v 0) (v 1) (v 2) + +variable {𝔅} + +lemma BoundedVecFunc.vecPr {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : 𝔅.VecPr (Function.Graphᵥ f) := h.1 + +lemma BoundedVecFunc.le_poly {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : ∃ p, 𝔅.Polynomial p ∧ f ≤ p := h.2 + +lemma UBVecFunc.boundedVecFunc_of_le {f g : (Fin k → V) → V} (hf : 𝔅.UBVecFunc f) (hg : 𝔅.BoundedVecFunc g) + (h : f ≤ g) : 𝔅.BoundedVecFunc f := by + constructor + · exact hf + · rcases hg.le_poly with ⟨p, hp, hgp⟩ + exact ⟨p, hp, le_trans h hgp⟩ + +lemma UBFunction.boundedVecFunc_of_le {f g : V → V} (hf : 𝔅.UBFunction f) (hg : 𝔅.BoundedFunction g) (h : f ≤ g) : 𝔅.BoundedFunction f := + UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _) + +lemma UBFunction₂.boundedVecFunc_of_le {f g : V → V → V} (hf : 𝔅.UBFunction₂ f) (hg : 𝔅.BoundedFunction₂ g) (h : f ≤ g) : 𝔅.BoundedFunction₂ f := + UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _ _) + +lemma UBFunction₃.boundedVecFunc_of_le {f g : V → V → V → V} (hf : 𝔅.UBFunction₃ f) (hg : 𝔅.BoundedFunction₃ g) (h : f ≤ g) : 𝔅.BoundedFunction₃ f := + UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _ _ _) + +@[simp] lemma BoundedVecFunc.nth (i : Fin k) : 𝔅.BoundedVecFunc (· i) := by + constructor + · apply equal' + · exact ⟨(· i), 𝔅.polynomial_nth i, by simp⟩ + +lemma BoundedVecFunc.replace {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) (c : Fin k → Fin l) : + 𝔅.BoundedVecFunc (fun v ↦ f (fun i ↦ v (c i))) := by + constructor + · apply of_iff (𝔅.replace (l := l + 1) hf.vecPr (0 :> fun x ↦ (c x).succ)) <| by + intro v; simp [Function.Graphᵥ] + · rcases hf.le_poly with ⟨p, pp, hfp⟩ + refine ⟨fun v ↦ p (fun i ↦ v (c i)), by apply 𝔅.polynomial_replace pp, by intro v; simpa using hfp _⟩ + +lemma lessThan' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i < v j := by + simpa using 𝔅.replace 𝔅.lessThan ![i, j] + +lemma lessThanOrEq (i j : Fin k) : 𝔅.VecPr fun v ↦ v i ≤ v j := by + simp [le_iff_lt_or_eq] + apply 𝔅.or + · apply lessThan' + · apply equal' + +lemma BoundedQuantifier.bex_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (pf : Polynomial 𝔅 f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by + simp only [not_exists, not_and]; exact 𝔅.ball_poly pf (𝔅.not hP) + +lemma BoundedQuantifier.bex_vec_poly {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} + (pp : ∀ i, Polynomial 𝔅 (p i)) (hP : 𝔅.VecPr fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), P v w := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq (α := V)] using hP + case succ l ih => + simp only [exists_le_vec_iff_exists_le_exists_vec] + apply bex_poly (pp 0) + apply ih + · intro i; apply 𝔅.polynomial_replace (pp i.succ) + · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (𝔅.replace hP g) <| by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · congr 1; ext; simp [Matrix.vecAppend_eq_ite] + +lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔅.VecPr P) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : + 𝔅.VecPr fun z ↦ P (fun i ↦ f i z) := by + choose p hp using fun i ↦ (hf i).le_poly + have : 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), (∀ i, w i = f i v) ∧ P w := by + apply BoundedQuantifier.bex_vec_poly (fun i ↦ (hp i).1) + apply 𝔅.and + · apply conj; intro i + simpa using 𝔅.replace (hf i).vecPr (i.natAdd l :> Fin.castAdd k) + · apply 𝔅.replace hP + apply of_iff this <| by + intro v; constructor + · rintro ⟨w, hw, e, h⟩ + rcases funext e + exact h + · intro h; exact ⟨(f · v), by intro i; simpa using (hp i).2 v, by simp, h⟩ + +lemma BoundedVecFunc.substitution {F : (Fin k → V) → V} {f : Fin k → (Fin l → V) → V} + (hF : 𝔅.BoundedVecFunc F) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : + 𝔅.BoundedVecFunc fun v ↦ F (fun i ↦ f i v) := by + constructor + · simpa [Function.Graphᵥ] using + BoundedSystem.substitution (l := l + 1) hF.vecPr (f := (· 0) :> fun i v ↦ f i (v ·.succ)) + (by intro i; cases' i using Fin.cases with i + · simp + · simpa using BoundedVecFunc.replace (hf i) _) + · rcases hF.le_poly with ⟨p, hp, hFp⟩ + choose ps hps using fun i ↦ (hf i).le_poly + refine ⟨fun v ↦ p fun i ↦ ps i v, 𝔅.polynomial_comp p ps hp (fun i ↦ (hps i).1), ?_⟩ + intro v; exact le_trans (hFp (f · v)) (𝔅.polynomial_monotone hp (fun i ↦ (hps i).2 v)) + +lemma ball_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x := by + rcases hf.le_poly with ⟨p, hp, hfp⟩ + have : 𝔅.VecPr fun v ↦ ∀ x ≤ p v, x ≤ f v → P v x := by + apply 𝔅.ball_poly hp + apply imply + · simpa using substitution (𝔅.lessThanOrEq 0 1) + (f := ![(· 0), fun v ↦ f (v ·.succ)]) + (by simpa [forall_fin_iff_zero_and_forall_succ] using hf.replace Fin.succ) + · exact hP + exact of_iff this <| by + intro v; constructor + · intro h x hx + exact h x (le_trans hx (hfp v)) hx + · intro h x _ hx + exact h x hx + +lemma bex_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by + simp only [not_exists, not_and] + exact ball_le hf (𝔅.not hP) + +lemma ball_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∀ x < f v, P v x := by + have : 𝔅.VecPr fun v ↦ ∀ x ≤ f v, x < f v → P v x := by + apply ball_le hf + apply imply ?_ hP + simpa using substitution (𝔅.lessThan' 0 1) + (f := ![(· 0), fun v ↦ f (v ·.succ)]) + (by simpa [forall_fin_iff_zero_and_forall_succ] using hf.replace Fin.succ) + exact of_iff this <| by + intro v; constructor + · intro h x hx + exact h x (le_of_lt hx) hx + · intro h x _ hx + exact h x hx + +lemma bex_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} + (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : + 𝔅.VecPr fun v ↦ ∃ x < f v, P v x := of_not <| by + simp only [not_exists, not_and] + exact ball_lt hf (𝔅.not hP) + +variable (𝔅) + +abbrev FactVecPr {k} (P : (Fin k → V) → Prop) : Prop := Fact (𝔅.VecPr P) + +abbrev IsPred (P : V → Prop) : Prop := Fact (𝔅.Pred P) + +abbrev IsRel (R : V → V → Prop) : Prop := Fact (𝔅.Rel R) + +abbrev IsRel₃ (R : V → V → V → Prop) : Prop := Fact (𝔅.Rel₃ R) + +abbrev IsBoundedConstant (c : V) : Prop := Fact (𝔅.BoundedConstant c) + +abbrev IsBoundedFunction (f : V → V) : Prop := Fact (𝔅.BoundedFunction f) + +abbrev IsBoundedFunction₂ (f : V → V → V) : Prop := Fact (𝔅.BoundedFunction₂ f) + +abbrev IsBoundedFunction₃ (f : V → V → V → V) : Prop := Fact (𝔅.BoundedFunction₃ f) + +variable {𝔅} + +instance : 𝔅.IsRel (· = ·) := ⟨equal' 0 1⟩ + +instance : 𝔅.IsRel (· < ·) := ⟨lessThan' 0 1⟩ + +instance : 𝔅.IsRel (· ≤ ·) := ⟨lessThanOrEq 0 1⟩ + +lemma Pred.comp {P : V → Prop} [hP : 𝔅.IsPred P] {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) : + 𝔅.VecPr fun v ↦ P (f v) := by + simpa using substitution hP.out (f := ![f]) (by simp [hf]) + +lemma Rel.comp {R : V → V → Prop} [hR : 𝔅.IsRel R] {f₁ f₂ : (Fin k → V) → V} (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) : + 𝔅.VecPr fun v ↦ R (f₁ v) (f₂ v) := by + simpa using substitution hR.out (f := ![f₁, f₂]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂]) + +lemma Rel₃.comp {R : V → V → V → Prop} [hR : 𝔅.IsRel₃ R] {f₁ f₂ f₃ : (Fin k → V) → V} + (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) (hf₃ : 𝔅.BoundedVecFunc f₃) : + 𝔅.VecPr fun v ↦ R (f₁ v) (f₂ v) (f₃ v) := by + simpa using substitution hR.out (f := ![f₁, f₂, f₃]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃]) + +@[simp] lemma Constant.comp (c : V) [hc : 𝔅.IsBoundedConstant c] : + 𝔅.BoundedVecFunc fun _ : Fin k → V ↦ c := + BoundedVecFunc.substitution (l := k) hc.out (f := ![]) (by simp) + +lemma Function.comp {F : V → V} [hF : 𝔅.IsBoundedFunction F] {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) : + 𝔅.BoundedVecFunc fun v ↦ F (f v) := by + simpa using BoundedVecFunc.substitution hF.out (f := ![f]) (by simp [hf]) + +lemma Function₂.comp {F : V → V → V} [hF : 𝔅.IsBoundedFunction₂ F] {f₁ f₂ : (Fin k → V) → V} + (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) : + 𝔅.BoundedVecFunc fun v ↦ F (f₁ v) (f₂ v) := by + simpa using BoundedVecFunc.substitution hF.out (f := ![f₁, f₂]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂]) + +lemma Function₃.comp {F : V → V → V → V} [hF : 𝔅.IsBoundedFunction₃ F] {f₁ f₂ f₃ : (Fin k → V) → V} + (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) (hf₃ : 𝔅.BoundedVecFunc f₃) : + 𝔅.BoundedVecFunc fun v ↦ F (f₁ v) (f₂ v) (f₃ v) := by + simpa using BoundedVecFunc.substitution hF.out (f := ![f₁, f₂, f₃]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃]) + +variable (𝔅) + +class Arithmetical where + zero : 𝔅.BoundedConstant 0 + one : 𝔅.BoundedConstant 1 + add : 𝔅.BoundedFunction₂ (· + ·) + mul : 𝔅.BoundedFunction₂ (· * ·) + +variable {𝔅} + +section Arithmetical + +variable [𝔅.Arithmetical] + +instance : 𝔅.IsBoundedConstant 0 := ⟨Arithmetical.zero⟩ + +instance : 𝔅.IsBoundedConstant 1 := ⟨Arithmetical.one⟩ + +instance : 𝔅.IsBoundedFunction₂ (· + ·) := ⟨Arithmetical.add⟩ + +instance : 𝔅.IsBoundedFunction₂ (· * ·) := ⟨Arithmetical.mul⟩ + +instance (n : ℕ) : 𝔅.IsBoundedConstant n := ⟨by + induction n + case zero => simp + case succ n ih => + simpa using Function₂.comp ih (by simp)⟩ + +end Arithmetical + +variable (𝔅) + +class Boldface where + const (z : V) : 𝔅.BoundedVecFunc (k := 0) fun _ ↦ z + +variable {𝔅} + +instance [𝔅.Boldface] (z : V) : 𝔅.IsBoundedConstant z := ⟨Boldface.const z⟩ + +end BoundedSystem + +section + +open Lean.Parser.Tactic (config) + +attribute [aesop (rule_sets := [Definability]) norm] + sq + Arith.pow_three + pow_four + +attribute [aesop 5 (rule_sets := [Definability]) safe] + BoundedSystem.Function.comp + BoundedSystem.Function₂.comp + BoundedSystem.Function₃.comp + BoundedSystem.Pred.comp + BoundedSystem.Rel.comp + BoundedSystem.Rel₃.comp + +attribute [aesop 8 (rule_sets := [Definability]) safe] + BoundedSystem.ball_le + BoundedSystem.bex_le + BoundedSystem.ball_lt + BoundedSystem.bex_lt + +attribute [aesop 10 (rule_sets := [Definability]) safe] + BooleanSystem.not + BooleanSystem.imply + BooleanSystem.iff + +attribute [aesop 11 (rule_sets := [Definability]) safe] + BooleanSystem.and + BooleanSystem.or + +macro "definability" : attr => + `(attr|aesop 10 (rule_sets := [$(Lean.mkIdent `Definability):ident]) safe) + +macro "definability" (config)? : tactic => + `(tactic| aesop (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) + +macro "definability?" (config)? : tactic => + `(tactic| aesop? (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) + +end + +end LO.Arith diff --git a/Arithmetization/Definability/BooleanSystem/DeltaZero.lean b/Arithmetization/Definability/BooleanSystem/DeltaZero.lean new file mode 100644 index 0000000..7214c53 --- /dev/null +++ b/Arithmetization/Definability/BooleanSystem/DeltaZero.lean @@ -0,0 +1,104 @@ +import Arithmetization.Definability.BooleanSystem.Open + +namespace LO.Arith + +open FirstOrder + +variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] + +def DeltaZero.Lightface : BoundedSystem V where + VecPr {k} (R) := Definable ℒₒᵣ (Arith.Hierarchy 𝚺 0) R + verum {k} := ⟨⊤, by simp, by intro v; simp⟩ + and {k P Q} := by + rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ + refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ + not {k P} := by + rintro ⟨p, hp, hP⟩ + exact ⟨~p, by simp [Arith.Hierarchy.pi_zero_iff_sigma_zero, hp], by intro v; simp [hP v]⟩ + equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ + replace {k l P} := by + rintro ⟨p, hp, hP⟩ f + refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ + Polynomial {k} (f) := ∃ t : Semiterm ℒₒᵣ Empty k, f = t.valbm V + polynomial_comp {l k F f} := by + rintro ⟨T, rfl⟩ ht + choose t ht using ht + rcases funext ht + exact ⟨Rew.substs t T, by ext v; simp [Semiterm.val_substs]⟩ + polynomial_replace {k l _} := by + rintro ⟨t, rfl⟩ f; exact ⟨Rew.substs (fun x ↦ #(f x)) t, by ext v; simp [Semiterm.val_substs]⟩ + polynomial_monotone {k _} := by + rintro ⟨t, rfl⟩ v w h + apply Structure.Monotone.term_monotone + · exact h + · simp + polynomial_nth {k i} := ⟨#i, by simp⟩ + ball_poly {k f P} := by + rintro ⟨t, rfl⟩ ⟨p, hp, hP⟩ + exact ⟨“∀ x <⁺ !t ⋯, !p x ⋯”, by simp [hp], by intro v; simp [Semiterm.val_substs, hP.iff]⟩ + lessThan := ⟨“x y | x < y”, by simp, by intro v; simp⟩ + +notation "𝛥₀[" V "]" => DeltaZero.Lightface (V := V) + +notation "𝛥₀" => DeltaZero.Lightface + +instance : 𝛥₀[V].Arithmetical where + zero := ⟨⟨“x | x = 0”, by simp, by intro; simp [Function.Graphᵥ]⟩, 0, ⟨‘0’, by ext v; simp⟩, by intro; simp⟩ + one := ⟨⟨“x | x = 1”, by simp, by intro; simp [Function.Graphᵥ]⟩, 1, ⟨‘1’, by ext v; simp⟩, by intro; simp⟩ + add := ⟨⟨“z x y | z = x + y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 + v 1, ⟨‘#0 + #1’, by ext v; simp⟩, by simp⟩ + mul := ⟨⟨“z x y | z = x * y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 * v 1, ⟨‘#0 * #1’, by ext v; simp⟩, by simp⟩ + +def DeltaZero.Boldface : BoundedSystem V where + VecPr {k} (R) := DefinableWithParam ℒₒᵣ (Arith.Hierarchy 𝚺 0) R + verum {k} := ⟨⊤, by simp, by intro v; simp⟩ + and {k P Q} := by + rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ + refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ + not {k P} := by + rintro ⟨p, hp, hP⟩ + exact ⟨~p, by simp [Arith.Hierarchy.pi_zero_iff_sigma_zero, hp], by intro v; simp [hP v]⟩ + equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ + replace {k l P} := by + rintro ⟨p, hp, hP⟩ f + refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ + Polynomial {k} (f) := ∃ t : Semiterm ℒₒᵣ V k, f = fun v ↦ t.valm V v id + polynomial_comp {l k F f} := by + rintro ⟨T, rfl⟩ ht + choose t ht using ht + rcases funext ht + exact ⟨Rew.substs t T, by ext v; simp [Semiterm.val_substs]⟩ + polynomial_replace {k l _} := by + rintro ⟨t, rfl⟩ f; exact ⟨Rew.substs (fun x ↦ #(f x)) t, by ext v; simp [Semiterm.val_substs]⟩ + polynomial_monotone {k _} := by + rintro ⟨t, rfl⟩ v w h + apply Structure.Monotone.term_monotone + · exact h + · simp + polynomial_nth {k i} := ⟨#i, by simp⟩ + ball_poly {k f P} := by + rintro ⟨t, rfl⟩ ⟨p, hp, hP⟩ + exact ⟨“∀ x <⁺ !t ⋯, !p x ⋯”, by simp [hp], by intro v; simp [Semiterm.val_substs, hP.iff]⟩ + lessThan := ⟨“x y | x < y”, by simp, by intro v; simp⟩ + +notation "𝜟₀" => DeltaZero.Boldface + +notation "𝜟₀[" V "]" => DeltaZero.Boldface (V := V) + +instance : 𝜟₀[V].Arithmetical where + zero := ⟨⟨“x | x = 0”, by simp, by intro; simp [Function.Graphᵥ]⟩, 0, ⟨‘0’, by ext v; simp⟩, by intro; simp⟩ + one := ⟨⟨“x | x = 1”, by simp, by intro; simp [Function.Graphᵥ]⟩, 1, ⟨‘1’, by ext v; simp⟩, by intro; simp⟩ + add := ⟨⟨“z x y | z = x + y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 + v 1, ⟨‘#0 + #1’, by ext v; simp⟩, by simp⟩ + mul := ⟨⟨“z x y | z = x * y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 * v 1, ⟨‘#0 * #1’, by ext v; simp⟩, by simp⟩ + +instance : 𝜟₀[V].Boldface where + const (z) := ⟨⟨“x | x = &z”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun _ ↦ z, ⟨&z, by simp⟩, by simp⟩ + +section + +example : 𝛥₀.Rel₃ fun x y z : V ↦ x = 32 → ¬y < z ∧ ∀ w < x + 2 * z, w < x := by definability? + +example (c : V) : 𝜟₀.Rel₃ fun x y z : V ↦ x = 32 → ¬y < z ∧ ∀ w < x + z, c < x := by definability + +end + +namespace LO.Arith diff --git a/Arithmetization/Definability/BooleanSystem/Open.lean b/Arithmetization/Definability/BooleanSystem/Open.lean new file mode 100644 index 0000000..770c4a3 --- /dev/null +++ b/Arithmetization/Definability/BooleanSystem/Open.lean @@ -0,0 +1,37 @@ +import Arithmetization.Definability.BooleanSystem.Basic + +namespace LO.Arith + +open FirstOrder + +variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] + +def Open.Lightface : BooleanSystem V where + VecPr {k} (R) := Definable ℒₒᵣ Semiformula.Open R + verum {k} := ⟨⊤, by simp, by intro v; simp⟩ + and {k P Q} := by + rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ + refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ + not {k P} := by + rintro ⟨p, hp, hP⟩ + exact ⟨~p, by simp [hp], by intro v; simp [hP v]⟩ + equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ + replace {k l P} := by + rintro ⟨p, hp, hP⟩ f + refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ + +def Open.Boldface : BooleanSystem V where + VecPr {k} (R) := DefinableWithParam ℒₒᵣ Semiformula.Open R + verum {k} := ⟨⊤, by simp, by intro v; simp⟩ + and {k P Q} := by + rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ + refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ + not {k P} := by + rintro ⟨p, hp, hP⟩ + exact ⟨~p, by simp [hp], by intro v; simp [hP v]⟩ + equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ + replace {k l P} := by + rintro ⟨p, hp, hP⟩ f + refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ + +namespace LO.Arith diff --git a/Arithmetization/Definability/SigmaPiDeltaSystem.lean b/Arithmetization/Definability/SigmaPiDeltaSystem.lean index d5f64a9..8472fb7 100644 --- a/Arithmetization/Definability/SigmaPiDeltaSystem.lean +++ b/Arithmetization/Definability/SigmaPiDeltaSystem.lean @@ -239,12 +239,12 @@ end VecFunc variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] class Arithmetical (𝔖 : SigmaPiDeltaSystem V) where - zero' (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 0 - one' (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 1 - add' (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 + v 1 - mul' (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 * v 1 - lt' (Γ) : 𝔖.VecPr Γ fun v : Fin 2 → V ↦ v 0 < v 1 - ball' {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∀ x ≤ v 0, P (v ·.succ) x + zero (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 0 + one (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 1 + add (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 + v 1 + mul (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 * v 1 + lt (Γ) : 𝔖.VecPr Γ fun v : Fin 2 → V ↦ v 0 < v 1 + ball {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∀ x ≤ v 0, P (v ·.succ) x bex' {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∃ x ≤ v 0, P (v ·.succ) x diff --git a/Arithmetization/Vorspiel/Vorspiel.lean b/Arithmetization/Vorspiel/Vorspiel.lean index 21039cc..15b76d0 100644 --- a/Arithmetization/Vorspiel/Vorspiel.lean +++ b/Arithmetization/Vorspiel/Vorspiel.lean @@ -102,15 +102,6 @@ lemma forall_le_vec_iff_forall_le_forall_vec [LE α] {P : (Fin (k + 1) → α) · intro h v hv simpa using h (v 0) (hv 0) (v ·.succ) (hv ·.succ) -@[simp] lemma forall_fin_two_iff {P : Fin 2 → Prop} : (∀ i, P i) ↔ P 0 ∧ P 1 := by - constructor - · intro h; exact ⟨h 0, h 1⟩ - · rintro ⟨h0, h1⟩ - intro i - match i with - | 0 => exact h0 - | 1 => exact h1 - instance : ToString Empty := ⟨Empty.elim⟩ class Hash (α : Type*) where From 537f87a4543fc2b9dc8e1ced66b3f0a7b1670e2e Mon Sep 17 00:00:00 2001 From: palalansouki Date: Fri, 16 Aug 2024 02:39:05 +0900 Subject: [PATCH 05/10] wip --- Arithmetization/Basic/PeanoMinus.lean | 2 +- .../Definability/Absoluteness.lean | 2 +- Arithmetization/Definability/Basic.lean | 29 - Arithmetization/Definability/Boldface.lean | 831 +++++++++++++++++ .../Definability/BooleanSystem/Basic.lean | 417 --------- .../Definability/BooleanSystem/DeltaZero.lean | 104 --- .../Definability/BooleanSystem/Open.lean | 37 - .../Definability/BoundedBoldface.lean | 427 +++++++++ Arithmetization/Definability/Defined.lean | 864 ------------------ .../{HSemiformula.lean => Hierarchy.lean} | 0 Arithmetization/Definability/PolyBounded.lean | 421 --------- .../Definability/SigmaPiDeltaSystem.lean | 299 ------ Arithmetization/Vorspiel/Lemmata.lean | 2 +- Arithmetization/Vorspiel/Vorspiel.lean | 4 +- lake-manifest.json | 2 +- 15 files changed, 1264 insertions(+), 2177 deletions(-) delete mode 100644 Arithmetization/Definability/Basic.lean create mode 100644 Arithmetization/Definability/Boldface.lean delete mode 100644 Arithmetization/Definability/BooleanSystem/Basic.lean delete mode 100644 Arithmetization/Definability/BooleanSystem/DeltaZero.lean delete mode 100644 Arithmetization/Definability/BooleanSystem/Open.lean create mode 100644 Arithmetization/Definability/BoundedBoldface.lean delete mode 100644 Arithmetization/Definability/Defined.lean rename Arithmetization/Definability/{HSemiformula.lean => Hierarchy.lean} (100%) delete mode 100644 Arithmetization/Definability/PolyBounded.lean delete mode 100644 Arithmetization/Definability/SigmaPiDeltaSystem.lean diff --git a/Arithmetization/Basic/PeanoMinus.lean b/Arithmetization/Basic/PeanoMinus.lean index 853f536..f354df7 100644 --- a/Arithmetization/Basic/PeanoMinus.lean +++ b/Arithmetization/Basic/PeanoMinus.lean @@ -1,4 +1,4 @@ -import Arithmetization.Definability.Absoluteness +import Arithmetization.Vorspiel.Lemmata namespace LO.Arith diff --git a/Arithmetization/Definability/Absoluteness.lean b/Arithmetization/Definability/Absoluteness.lean index 480cfb1..5073a9e 100644 --- a/Arithmetization/Definability/Absoluteness.lean +++ b/Arithmetization/Definability/Absoluteness.lean @@ -1,4 +1,4 @@ -import Arithmetization.Definability.Definability +import Arithmetization.Definability.Basic noncomputable section diff --git a/Arithmetization/Definability/Basic.lean b/Arithmetization/Definability/Basic.lean deleted file mode 100644 index 6f1a088..0000000 --- a/Arithmetization/Definability/Basic.lean +++ /dev/null @@ -1,29 +0,0 @@ -import Arithmetization.Vorspiel.Lemmata -import Arithmetization.Vorspiel.Graph -import Logic.FirstOrder.Arith.StrictHierarchy - -namespace LO.FirstOrder - -variable {L : Language} {V : Type*} [Structure L V] - -def Defined {k} (R : (Fin k → V) → Prop) (p : Semisentence L k) : Prop := - ∀ v, R v ↔ Semiformula.Evalbm V v p - -def DefinedWithParam {k} (R : (Fin k → V) → Prop) (p : Semiformula L V k) : Prop := - ∀ v, R v ↔ Semiformula.Evalm V v id p - -lemma Defined.iff {k} {R : (Fin k → V) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : - Semiformula.Evalbm V v p ↔ R v := (h v).symm - -lemma DefinedWithParam.iff {k} {R : (Fin k → V) → Prop} {p : Semiformula L V k} (h : DefinedWithParam R p) (v) : - Semiformula.Evalm V v id p ↔ R v := (h v).symm - -variable (L) - -def Definable {k} (C : Semisentence L k → Prop) (R : (Fin k → V) → Prop) : Prop := - ∃ p, C p ∧ Defined R p - -def DefinableWithParam {k} (C : Semiformula L V k → Prop) (R : (Fin k → V) → Prop) : Prop := - ∃ p, C p ∧ DefinedWithParam R p - -end LO.FirstOrder diff --git a/Arithmetization/Definability/Boldface.lean b/Arithmetization/Definability/Boldface.lean new file mode 100644 index 0000000..1336f97 --- /dev/null +++ b/Arithmetization/Definability/Boldface.lean @@ -0,0 +1,831 @@ +import Arithmetization.Definability.Hierarchy +import Arithmetization.Vorspiel.Graph + +namespace LO.FirstOrder.Arith + +end Arith + +def Defined {k} (R : (Fin k → V) → Prop) [Structure L V] (p : Semisentence L k) : Prop := + ∀ v, R v ↔ Semiformula.Evalbm V v p + +def DefinedWithParam {k} (R : (Fin k → V) → Prop) [Structure L V] (p : Semiformula L V k) : Prop := + ∀ v, R v ↔ Semiformula.Evalm V v id p + +lemma Defined.iff [Structure L V] {k} {R : (Fin k → V) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : + Semiformula.Evalbm V v p ↔ R v := (h v).symm + +lemma DefinedWithParam.iff [Structure L V] {k} {R : (Fin k → V) → Prop} {p : Semiformula L V k} (h : DefinedWithParam R p) (v) : + Semiformula.Evalm V v id p ↔ R v := (h v).symm + +namespace Arith.HierarchySymbol + +variable (ξ : Type*) (n : ℕ) + +open LO.Arith + +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐏𝐀⁻] + +def Defined (R : (Fin k → V) → Prop) : {ℌ : HierarchySymbol} → ℌ.Semisentence k → Prop + | 𝚺-[_], p => FirstOrder.Defined R p.val + | 𝚷-[_], p => FirstOrder.Defined R p.val + | 𝚫-[_], p => p.ProperOn V ∧ FirstOrder.Defined R p.val + +def DefinedWithParam (R : (Fin k → V) → Prop) : {ℌ : HierarchySymbol} → ℌ.Semiformula V k → Prop + | 𝚺-[_], p => FirstOrder.DefinedWithParam R p.val + | 𝚷-[_], p => FirstOrder.DefinedWithParam R p.val + | 𝚫-[_], p => p.ProperWithParamOn V ∧ FirstOrder.DefinedWithParam R p.val + +variable {ℌ : HierarchySymbol} {Γ : SigmaPiDelta} + +section + +variable (ℌ) + +class Definable {k} (P : (Fin k → V) → Prop) : Prop where + definable : ∃ p : ℌ.Semisentence k, Defined P p + +class Boldface {k} (P : (Fin k → V) → Prop) : Prop where + definable : ∃ p : ℌ.Semiformula V k, DefinedWithParam P p + +abbrev DefinedPred (P : V → Prop) (p : ℌ.Semisentence 1) : Prop := + Defined (λ v ↦ P (v 0)) p + +abbrev DefinedRel (R : V → V → Prop) (p : ℌ.Semisentence 2) : Prop := + Defined (λ v ↦ R (v 0) (v 1)) p + +abbrev DefinedRel₃ (R : V → V → V → Prop) (p : ℌ.Semisentence 3) : Prop := + Defined (λ v ↦ R (v 0) (v 1) (v 2)) p + +abbrev DefinedRel₄ (R : V → V → V → V → Prop) (p : ℌ.Semisentence 4) : Prop := + Defined (λ v ↦ R (v 0) (v 1) (v 2) (v 3)) p + +variable {ℌ} + +abbrev DefinedFunction {k} (f : (Fin k → V) → V) (p : ℌ.Semisentence (k + 1)) : Prop := + Defined (fun v => v 0 = f (v ·.succ)) p + +variable (ℌ) + +abbrev DefinedFunction₁ (f : V → V) (p : ℌ.Semisentence 2) : Prop := + DefinedFunction (fun v => f (v 0)) p + +abbrev DefinedFunction₂ (f : V → V → V) (p : ℌ.Semisentence 3) : Prop := + DefinedFunction (fun v => f (v 0) (v 1)) p + +abbrev DefinedFunction₃ (f : V → V → V → V) (p : ℌ.Semisentence 4) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2)) p + +abbrev DefinedFunction₄ (f : V → V → V → V → V) (p : ℌ.Semisentence 5) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3)) p + +abbrev DefinedFunction₅ (f : V → V → V → V → V → V) (p : ℌ.Semisentence 6) : Prop := + DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3) (v 4)) p + +abbrev BoldfacePred (P : V → Prop) : Prop := ℌ.Boldface (k := 1) (fun v ↦ P (v 0)) + +abbrev BoldfaceRel (P : V → V → Prop) : Prop := ℌ.Boldface (k := 2) (fun v ↦ P (v 0) (v 1)) + +abbrev BoldfaceRel₃ (P : V → V → V → Prop) : Prop := ℌ.Boldface (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) + +abbrev BoldfaceRel₄ (P : V → V → V → V → Prop) : Prop := ℌ.Boldface (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) + +abbrev BoldfaceRel₅ (P : V → V → V → V → V → Prop) : Prop := ℌ.Boldface (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) + +abbrev BoldfaceRel₆ (P : V → V → V → V → V → V → Prop) : Prop := ℌ.Boldface (k := 6) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) + +abbrev BoldfaceFunction (f : (Fin k → V) → V) : Prop := ℌ.Boldface (k := k + 1) (fun v ↦ v 0 = f (v ·.succ)) + +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)) + +abbrev BoldfaceFunction₃ (f : V → V → V → V) : Prop := ℌ.BoldfaceFunction (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) + +abbrev BoldfaceFunction₄ (f : V → V → V → V → V) : Prop := ℌ.BoldfaceFunction (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) + +abbrev BoldfaceFunction₅ (f : V → V → V → V → V → V) : Prop := ℌ.BoldfaceFunction (k := 5) (fun v ↦ f (v 0) (v 1) (v 2) (v 3) (v 4)) + +variable {ℌ} + +end + +section + +variable {k} {P Q : (Fin k → V) → Prop} + +namespace Defined + +lemma df {R : (Fin k → V) → Prop} {p : ℌ.Semisentence k} (h : Defined R p) : FirstOrder.Defined R p.val := + match ℌ with + | 𝚺-[_] => h + | 𝚷-[_] => h + | 𝚫-[_] => h.2 + +lemma proper {R : (Fin k → V) → Prop} {m} {p : 𝚫-[m].Semisentence k} (h : Defined R p) : p.ProperOn V := h.1 + +lemma of_zero {R : (Fin k → V) → Prop} {p : 𝚺₀.Semisentence k} (h : Defined R p) : Defined R (p.ofZero ℌ) := + match ℌ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simp, by intro _; simp [h.iff]⟩ + +lemma emb {R : (Fin k → V) → Prop} {p : ℌ.Semisentence k} (h : Defined R p) : Defined R p.emb := + match ℌ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ + +lemma of_iff {P Q : (Fin k → V) → Prop} (h : ∀ x, P x ↔ Q x) {p : ℌ.Semisentence k} (H : Defined Q p) : Defined P p := by + rwa [show P = Q from by funext v; simp [h]] + +lemma to_definable (p : ℌ.Semisentence k) (hP : Defined P p) : ℌ.Boldface P := ⟨p.rew Rew.emb, by + match ℌ with + | 𝚺-[_] => intro; simp [hP.iff] + | 𝚷-[_] => intro; simp [hP.iff] + | 𝚫-[_] => exact ⟨ + fun v ↦ by rcases p; simpa [HierarchySymbol.Semiformula.rew] using hP.proper.rew Rew.emb v, + by intro; simp [hP.df.iff]⟩⟩ + +lemma to_definable₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : + ℌ.Boldface P := Defined.to_definable (p.ofZero ℌ) hP.of_zero + +lemma to_definable_oRing (p : ℌ.Semisentence k) (hP : Defined P p) : + ℌ.Boldface P := Defined.to_definable p.emb hP.emb + +lemma to_definable_oRing₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : + ℌ.Boldface P := Defined.to_definable₀ p.emb hP.emb + +end Defined + +namespace DefinedFunction + +lemma of_eq {f g : (Fin k → V) → V} (h : ∀ x, f x = g x) + {p : ℌ.Semisentence (k + 1)} (H : DefinedFunction f p) : DefinedFunction g p := + Defined.of_iff (by intro; simp [h]) H + +lemma graph_delta {f : (Fin k → V) → V} {p : 𝚺-[m].Semisentence (k + 1)} + (h : DefinedFunction f p) : DefinedFunction f p.graphDelta := + ⟨by cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] + intro e; simp [Empty.eq_elim, h.df.iff] + rw [eq_comm], + by intro v; simp [h.df.iff]⟩ + +end DefinedFunction + +namespace DefinedWithParam + +lemma df {R : (Fin k → V) → Prop} {p : ℌ.Semiformula V k} (h : DefinedWithParam R p) : FirstOrder.DefinedWithParam R p.val := + match ℌ with + | 𝚺-[_] => h + | 𝚷-[_] => h + | 𝚫-[_] => h.2 + +lemma proper {R : (Fin k → V) → Prop} {m} {p : 𝚫-[m].Semiformula V k} (h : DefinedWithParam R p) : p.ProperWithParamOn V := h.1 + +lemma of_zero {R : (Fin k → V) → Prop} {Γ'} {p : Γ'-[0].Semiformula V k} + (h : DefinedWithParam R p) {Γ} : DefinedWithParam R (p.ofZero Γ) := + match Γ with + | 𝚺-[m] => by intro _; simp [h.df.iff] + | 𝚷-[m] => by intro _; simp [h.df.iff] + | 𝚫-[m] => ⟨by simp , by intro _; simp [h.df.iff]⟩ + +lemma of_deltaOne {R : (Fin k → V) → Prop} {Γ m} {p : 𝚫₁.Semiformula V k} + (h : DefinedWithParam R p) : DefinedWithParam R (p.ofDeltaOne Γ m) := + match Γ with + | 𝚺 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma] + | 𝚷 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, h.proper.iff'] + | 𝚫 => ⟨by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma, h.proper.iff'], + by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ + +lemma emb {R : (Fin k → V) → Prop} {p : ℌ.Semiformula V k} (h : DefinedWithParam R p) : DefinedWithParam R p.emb := + match ℌ with + | 𝚺-[m] => by intro _; simp [h.iff] + | 𝚷-[m] => by intro _; simp [h.iff] + | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ + +lemma of_iff {P Q : (Fin k → V) → Prop} (h : ∀ x, P x ↔ Q x) + {p : ℌ.Semiformula V k} (H : DefinedWithParam Q p) : DefinedWithParam P p := by + rwa [show P = Q from by funext v; simp [h]] + +lemma to_definable {p : ℌ.Semiformula V k} (h : DefinedWithParam P p) : ℌ.Boldface P := ⟨p, h⟩ + +lemma to_definable₀ {p : Γ'-[0].Semiformula V k} + (h : DefinedWithParam P p) : ℌ.Boldface P := ⟨p.ofZero ℌ, h.of_zero⟩ + +lemma to_definable_deltaOne {p : 𝚫₁.Semiformula V k} {Γ m} + (h : DefinedWithParam P p) : Γ-[m + 1].Boldface P := ⟨p.ofDeltaOne Γ m, h.of_deltaOne⟩ + +lemma retraction {p : ℌ.Semiformula V k} (hp : DefinedWithParam P p) (f : Fin k → Fin l) : + DefinedWithParam (fun v ↦ P fun i ↦ v (f i)) (p.rew <| Rew.substs fun x ↦ #(f x)) := + match ℌ with + | 𝚺-[_] => by intro; simp [hp.df.iff] + | 𝚷-[_] => by intro; simp [hp.df.iff] + | 𝚫-[_] => ⟨hp.proper.rew _, by intro; simp [hp.df.iff]⟩ + +@[simp] lemma verum : DefinedWithParam (fun _ ↦ True) (⊤ : ℌ.Semiformula V k) := + match ℌ with + | 𝚺-[m] => by intro v; simp + | 𝚷-[m] => by intro v; simp + | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ + +@[simp] lemma falsum : DefinedWithParam (fun _ ↦ False) (⊥ : ℌ.Semiformula V k) := + match ℌ with + | 𝚺-[m] => by intro v; simp + | 𝚷-[m] => by intro v; simp + | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ + +lemma and {p q : ℌ.Semiformula V k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ∧ Q x) (p ⋏ q) := + match ℌ with + | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚫-[m] => ⟨hp.proper.and hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ + +lemma or {p q : ℌ.Semiformula V k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ∨ Q x) (p ⋎ q) := + match ℌ with + | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] + | 𝚫-[m] => ⟨hp.proper.or hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ + +lemma negSigma {p : 𝚺-[m].Semiformula V k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) p.negSigma := by intro v; simp [hp.iff] + +lemma negPi {p : 𝚷-[m].Semiformula V k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) p.negPi := by intro v; simp [hp.iff] + +lemma not {p : 𝚫-[m].Semiformula V k} (hp : DefinedWithParam P p) : + DefinedWithParam (fun x ↦ ¬P x) (~p) := ⟨hp.proper.neg, by intro v; simp [hp.proper.eval_neg, hp.df.iff]⟩ + +lemma imp {p q : 𝚫-[m].Semiformula V k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x → Q x) (p ⟶ q) := (hp.not.or hq).of_iff (by intro x; simp [imp_iff_not_or]) + +lemma iff {p q : 𝚫-[m].Semiformula V k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : + DefinedWithParam (fun x ↦ P x ↔ Q x) (p ⟷ q) := ((hp.imp hq).and (hq.imp hp)).of_iff <| by intro v; simp [iff_iff_implies_and_implies] + +lemma ball {P : (Fin (k + 1) → V) → Prop} {p : ℌ.Semiformula V (k + 1)} + (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ V k) : + DefinedWithParam (fun v ↦ ∀ x < t.valm V v id, P (x :> v)) (HierarchySymbol.Semiformula.ball t p) := + match ℌ with + | 𝚺-[m] => by intro v; simp [hp.df.iff] + | 𝚷-[m] => by intro v; simp [hp.df.iff] + | 𝚫-[m] => ⟨hp.proper.ball, by intro v; simp [hp.df.iff]⟩ + +lemma bex {P : (Fin (k + 1) → V) → Prop} {p : ℌ.Semiformula V (k + 1)} + (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ V k) : + DefinedWithParam (fun v ↦ ∃ x < t.valm V v id, P (x :> v)) (HierarchySymbol.Semiformula.bex t p) := + match ℌ with + | 𝚺-[m] => by intro v; simp [hp.df.iff] + | 𝚷-[m] => by intro v; simp [hp.df.iff] + | 𝚫-[m] => ⟨hp.proper.bex, by intro v; simp [hp.df.iff]⟩ + +lemma ex {P : (Fin (k + 1) → V) → Prop} {p : 𝚺-[m + 1].Semiformula V (k + 1)} + (hp : DefinedWithParam P p) : + DefinedWithParam (fun v ↦ ∃ x, P (x :> v)) p.ex := by intro _; simp [hp.df.iff] + +lemma all {P : (Fin (k + 1) → V) → Prop} {p : 𝚷-[m + 1].Semiformula V (k + 1)} + (hp : DefinedWithParam P p) : + DefinedWithParam (fun v ↦ ∀ x, P (x :> v)) p.all := by intro _; simp [hp.df.iff] + +end DefinedWithParam + +namespace BoldfaceRel + +@[simp] instance eq : ℌ.BoldfaceRel (Eq : V → V → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1” (by simp)) (by intro _; simp) + +@[simp] instance lt : ℌ.BoldfaceRel (LT.lt : V → V → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 < #1” (by simp)) (by intro _; simp) + +@[simp] instance le : ℌ.BoldfaceRel (LE.le : V → V → Prop) := + Defined.to_definable_oRing₀ (.mkSigma “#0 ≤ #1” (by simp)) (by intro _; simp) + +end BoldfaceRel + +namespace BoldfaceFunction₂ + +@[simp] instance add : ℌ.BoldfaceFunction₂ ((· + ·) : V → V → V) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) + +@[simp] instance mul : ℌ.BoldfaceFunction₂ ((· * ·) : V → V → V) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) + +@[simp] instance hAdd : ℌ.BoldfaceFunction₂ (HAdd.hAdd : V → V → V) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) + +@[simp] instance hMul : ℌ.BoldfaceFunction₂ (HMul.hMul : V → V → V) := + Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) + +end BoldfaceFunction₂ + +namespace Boldface + +lemma mkPolarity {P : (Fin k → V) → Prop} {Γ : Polarity} + (p : Semiformula ℒₒᵣ V k) (hp : Hierarchy Γ m p) (hP : ∀ v, P v ↔ Semiformula.Evalm V v id p) : Γ-[m].Boldface P := + match Γ with + | 𝚺 => ⟨.mkSigma p hp, by intro v; simp [hP]⟩ + | 𝚷 => ⟨.mkPi p hp, by intro v; simp [hP]⟩ + +lemma of_iff (H : ℌ.Boldface Q) (h : ∀ x, P x ↔ Q x) : ℌ.Boldface P := by + rwa [show P = Q from by funext v; simp [h]] + +lemma of_oRing (h : ℌ.Boldface P) : ℌ.Boldface P := by + rcases h with ⟨p, hP⟩; exact ⟨p.emb, hP.emb⟩ + +lemma of_delta (h : 𝚫-[m].Boldface P) : Γ-[m].Boldface P := by + rcases h with ⟨p, h⟩ + match Γ with + | 𝚺 => exact ⟨p.sigma, by intro v; simp [HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ + | 𝚷 => exact ⟨p.pi, by intro v; simp [←h.proper v, HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ + | 𝚫 => exact ⟨p, h⟩ + +instance [𝚫-[m].Boldface P] (Γ) : Γ-[m].Boldface P := of_delta inferInstance + +lemma of_sigma_of_pi (hσ : 𝚺-[m].Boldface P) (hπ : 𝚷-[m].Boldface P) : Γ-[m].Boldface P := + match Γ with + | 𝚺 => hσ + | 𝚷 => hπ + | 𝚫 => by + rcases hσ with ⟨p, hp⟩; rcases hπ with ⟨q, hq⟩ + exact ⟨.mkDelta p q, by intro v; simp [hp.df.iff, hq.df.iff], by intro v; simp [hp.df.iff]⟩ + +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 + rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable_deltaOne + +instance [𝚺₀.Boldface P] (ℌ : HierarchySymbol) : ℌ.Boldface P := Boldface.of_zero (Γ' := 𝚺) (ℌ := ℌ) inferInstance + +lemma retraction (h : ℌ.Boldface P) {n} (f : Fin k → Fin n) : + ℌ.Boldface fun v ↦ P (fun i ↦ v (f i)) := by + rcases h with ⟨p, h⟩ + exact ⟨p.rew (Rew.substs (fun i ↦ #(f i))), + match ℌ with + | 𝚺-[_] => by intro; simp [h.df.iff] + | 𝚷-[_] => by intro; simp [h.df.iff] + | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ + +lemma retractiont (h : ℌ.Boldface P) (f : Fin k → Semiterm ℒₒᵣ V n) : + ℌ.Boldface fun v ↦ P (fun i ↦ Semiterm.valm V v id (f i)) := by + rcases h with ⟨p, h⟩ + exact ⟨p.rew (Rew.substs f), + match ℌ with + | 𝚺-[_] => by intro; simp [h.df.iff] + | 𝚷-[_] => by intro; simp [h.df.iff] + | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ + +@[simp] lemma const {P : Prop} : ℌ.Boldface (fun _ : Fin k → V ↦ P) := of_zero (by + by_cases hP : P + · exact ⟨.mkSigma ⊤ (by simp), by intro; simp[hP]⟩ + · exact ⟨.mkSigma ⊥ (by simp), by intro; simp[hP]⟩) + +lemma and (h₁ : ℌ.Boldface P) (h₂ : ℌ.Boldface Q) : + ℌ.Boldface (fun v ↦ P v ∧ Q v) := by + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁ ⋏ p₂, h₁.and h₂⟩ + +lemma conj {k l} {P : Fin l → (Fin k → V) → Prop} + (h : ∀ i, ℌ.Boldface fun w : Fin k → V ↦ P i w) : + ℌ.Boldface fun v : Fin k → V ↦ ∀ i, P i v := by + induction l + case zero => simp + case succ l ih => + suffices ℌ.Boldface fun v : Fin k → V ↦ P 0 v ∧ ∀ i : Fin l, P i.succ v by + apply of_iff this; intro x + constructor + · intro h + exact ⟨h 0, fun i ↦ h i.succ⟩ + · rintro ⟨h0, hs⟩ + intro i; cases' i using Fin.cases with i + · exact h0 + · exact hs i + apply and (h 0); apply ih + intro i; exact h i.succ + +lemma or (h₁ : ℌ.Boldface P) (h₂ : ℌ.Boldface Q) : + ℌ.Boldface (fun v ↦ P v ∨ Q v) := by + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁ ⋎ p₂, h₁.or h₂⟩ + +lemma not (h : Γ.alt-[m].Boldface P) : + Γ-[m].Boldface (fun v ↦ ¬P v) := by + match Γ with + | 𝚺 => rcases h with ⟨p, h⟩; exact ⟨p.negPi, h.negPi⟩ + | 𝚷 => rcases h with ⟨p, h⟩; exact ⟨p.negSigma, h.negSigma⟩ + | 𝚫 => rcases h with ⟨p, h⟩; exact ⟨p.negDelta, h.not⟩ + +lemma imp (h₁ : Γ.alt-[m].Boldface P) (h₂ : Γ-[m].Boldface Q) : + Γ-[m].Boldface (fun v ↦ P v → Q v) := by + match Γ with + | 𝚺 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁.negPi.or p₂, (h₁.negPi.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ + | 𝚷 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ + exact ⟨p₁.negSigma.or p₂, (h₁.negSigma.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ + | 𝚫 => + rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩; exact ⟨p₁ ⟶ p₂, h₁.imp h₂⟩ + +lemma iff (h₁ : 𝚫-[m].Boldface P) (h₂ : 𝚫-[m].Boldface Q) {Γ} : + Γ-[m].Boldface (fun v ↦ P v ↔ Q v) := + .of_delta (by rcases h₁ with ⟨p, hp⟩; rcases h₂ with ⟨q, hq⟩; exact ⟨p ⟷ q, hp.iff hq⟩) + +lemma all {P : (Fin k → V) → V → Prop} (h : 𝚷-[s + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + 𝚷-[s + 1].Boldface (fun v ↦ ∀ x, P v x) := by + rcases h with ⟨p, hp⟩ + exact ⟨.mkPi (∀' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ + +lemma ex {P : (Fin k → V) → V → Prop} (h : 𝚺-[s + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + 𝚺-[s + 1].Boldface (fun v ↦ ∃ x, P v x) := by + rcases h with ⟨p, hp⟩ + exact ⟨.mkSigma (∃' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ + +lemma equal' (i j : Fin k) : ℌ.Boldface fun v : Fin k → V ↦ v i = v j := by + simpa using retraction BoldfaceRel.eq ![i, j] + +lemma of_sigma {f : (Fin k → V) → V} (h : 𝚺-[m].BoldfaceFunction f) {Γ} : Γ-[m].BoldfaceFunction f := by + cases' m with m + · exact of_zero h + apply of_sigma_of_pi + · exact h + · have : 𝚷-[m + 1].Boldface fun v ↦ ∀ y, y = f (v ·.succ) → v 0 = y := all <| imp + (by simpa using retraction h (0 :> (·.succ.succ))) + (by simpa using equal' 1 0) + exact of_iff this (fun v ↦ by simp) + +lemma exVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} + (h : 𝚺-[m + 1].Boldface fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝚺-[m + 1].Boldface fun v : Fin k → V ↦ ∃ ys : Fin l → V, P v ys := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq] using h + case succ l ih => + suffices 𝚺-[m + 1].Boldface fun v : Fin k → V ↦ ∃ y, ∃ ys : Fin l → V, P v (y :> ys) by + apply of_iff this; intro x + constructor + · rintro ⟨ys, h⟩; exact ⟨ys 0, (ys ·.succ), by simpa using h⟩ + · rintro ⟨y, ys, h⟩; exact ⟨_, h⟩ + apply ex; apply ih + let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (retraction h g) (by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite]) + +lemma allVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} + (h : 𝚷-[m+1].Boldface fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + 𝚷-[m+1].Boldface fun v : Fin k → V ↦ ∀ ys : Fin l → V, P v ys := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq] using h + case succ l ih => + suffices 𝚷-[m+1].Boldface fun v : Fin k → V ↦ ∀ y, ∀ ys : Fin l → V, P v (y :> ys) by + apply of_iff this; intro x + constructor + · intro h y ys; apply h + · intro h ys; simpa using h (ys 0) (ys ·.succ) + apply all; apply ih + let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (retraction h g) (by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite]) + +private lemma substitution_sigma {f : Fin k → (Fin l → V) → V} (hP : 𝚺-[m+1].Boldface P) (hf : ∀ i, 𝚺-[m+1].BoldfaceFunction (f i)) : + 𝚺-[m+1].Boldface fun z ↦ P (fun i ↦ f i z) := by + have : 𝚺-[m+1].Boldface fun z ↦ ∃ ys : Fin k → V, (∀ i, ys i = f i z) ∧ P ys := by + apply exVec; apply and + · apply conj; intro i + simpa using retraction (of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact retraction hP (Fin.natAdd l) + exact of_iff this <| by + intro v + constructor + · intro hP + exact ⟨(f · v), by simp, hP⟩ + · rintro ⟨ys, hys, hP⟩ + have : ys = fun i ↦ f i v := funext hys + rcases this; exact hP + +private lemma substitution_pi {f : Fin k → (Fin l → V) → V} (hP : 𝚷-[m+1].Boldface P) (hf : ∀ i, 𝚺-[m+1].BoldfaceFunction (f i)) : + 𝚷-[m+1].Boldface fun z ↦ P (fun i ↦ f i z) := by + have : 𝚷-[m+1].Boldface fun z ↦ ∀ ys : Fin k → V, (∀ i, ys i = f i z) → P ys := by + apply allVec; apply imp + · apply conj; intro i + simpa using retraction (of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) + · exact retraction hP (Fin.natAdd l) + exact of_iff this <| by + intro v + constructor + · intro h ys e + have : ys = (f · v) := funext e + rcases this; exact h + · intro h; apply h _ (by simp) + +lemma substitution {f : Fin k → (Fin l → V) → V} + (hP : Γ-[m + 1].Boldface P) (hf : ∀ i, 𝚺-[m + 1].BoldfaceFunction (f i)) : + Γ-[m + 1].Boldface fun z ↦ P (fun i ↦ f i z) := + match Γ with + | 𝚺 => substitution_sigma hP hf + | 𝚷 => substitution_pi hP hf + | 𝚫 => of_sigma_of_pi (substitution_sigma (of_delta hP) hf) (substitution_pi (of_delta hP) hf) + +end Boldface + +lemma BoldfacePred.comp {P : V → Prop} {k} {f : (Fin k → V) → V} + (hP : Γ-[m + 1].BoldfacePred P) (hf : 𝚺-[m + 1].BoldfaceFunction f) : + Γ-[m + 1].Boldface (fun v ↦ P (f v)) := + Boldface.substitution (f := ![f]) hP (by simpa using hf) + +lemma BoldfaceRel.comp {P : V → V → Prop} {k} {f g : (Fin k → V) → V} + (hP : Γ-[m + 1].BoldfaceRel P) + (hf : 𝚺-[m + 1].BoldfaceFunction f) (hg : 𝚺-[m + 1].BoldfaceFunction g) : + Γ-[m + 1].Boldface fun v ↦ P (f v) (g v) := + Boldface.substitution (f := ![f, g]) hP (by simp [forall_fin_iff_zero_and_forall_succ, hf, hg]) + +lemma BoldfaceRel₃.comp {k} {P : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} + (hP : Γ-[m + 1].BoldfaceRel₃ P) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := + Boldface.substitution (f := ![f₁, f₂, f₃]) hP (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃]) + +lemma BoldfaceRel₄.comp {k} {P : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} + (hP : Γ-[m + 1].BoldfaceRel₄ P) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := + Boldface.substitution (f := ![f₁, f₂, f₃, f₄]) hP (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃, hf₄]) + +lemma BoldfaceRel₅.comp {k} {P : V → V → V → V → V → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → V) → V} + (hP : Γ-[m + 1].BoldfaceRel₅ P) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) + (hf₅ : 𝚺-[m + 1].BoldfaceFunction f₅) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := + Boldface.substitution (f := ![f₁, f₂, f₃, f₄, f₅]) hP (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃, hf₄, hf₅]) + +namespace Boldface + +lemma comp₁ {k} {P : V → Prop} {f : (Fin k → V) → V} + [Γ-[m + 1].BoldfacePred P] + (hf : 𝚺-[m + 1].BoldfaceFunction f) : Γ-[m + 1].Boldface fun v ↦ P (f v) := + BoldfacePred.comp inferInstance hf + +lemma comp₂ {k} {P : V → V → Prop} {f g : (Fin k → V) → V} + [Γ-[m + 1].BoldfaceRel P] + (hf : 𝚺-[m + 1].BoldfaceFunction f) (hg : 𝚺-[m + 1].BoldfaceFunction g) : + Γ-[m + 1].Boldface (fun v ↦ P (f v) (g v)) := + BoldfaceRel.comp inferInstance hf hg + +lemma comp₃ {k} {P : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} + [Γ-[m + 1].BoldfaceRel₃ P] + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := + BoldfaceRel₃.comp inferInstance hf₁ hf₂ hf₃ + +lemma comp₄ {k} {P : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} + [Γ-[m + 1].BoldfaceRel₄ P] + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := + BoldfaceRel₄.comp inferInstance hf₁ hf₂ hf₃ hf₄ + +lemma comp₅ {k} {P : V → V → V → V → V → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → V) → V} + [Γ-[m + 1].BoldfaceRel₅ P] + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) + (hf₅ : 𝚺-[m + 1].BoldfaceFunction f₅) : + Γ-[m + 1].Boldface (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := + BoldfaceRel₅.comp inferInstance hf₁ hf₂ hf₃ hf₄ hf₅ + +end Boldface + +section + +variable {ℌ : HierarchySymbol} + +lemma BoldfacePred.of_iff {P Q : V → Prop} + (H : ℌ.BoldfacePred Q) (h : ∀ x, P x ↔ Q x) : ℌ.BoldfacePred P := by + rwa [show P = Q from by funext v; simp [h]] + +instance BoldfaceFunction₁.graph {f : V → V} [h : ℌ.BoldfaceFunction₁ f] : + ℌ.BoldfaceRel (Function.Graph f) := h + +instance BoldfaceFunction₂.graph {f : V → V → V} [h : ℌ.BoldfaceFunction₂ f] : + ℌ.BoldfaceRel₃ (Function.Graph₂ f) := h + +instance BoldfaceFunction₃.graph {f : V → V → V → V} [h : ℌ.BoldfaceFunction₃ f] : + ℌ.BoldfaceRel₄ (Function.Graph₃ f) := h + +end + +namespace BoldfaceFunction + +variable {ℌ : HierarchySymbol} + +lemma graph_delta {k} {f : (Fin k → V) → V} + (h : 𝚺-[m].BoldfaceFunction f) : 𝚫-[m].BoldfaceFunction f := by + rcases h with ⟨p, h⟩ + exact ⟨p.graphDelta, by + cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] + intro e; simp [Empty.eq_elim, h.df.iff] + exact eq_comm, by + intro v; simp [h.df.iff]⟩ + +instance {k} {f : (Fin k → V) → V} [h : 𝚺-[m].BoldfaceFunction f] : 𝚫-[m].BoldfaceFunction f := + BoldfaceFunction.graph_delta h + +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 + +@[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⟩ + +@[simp] lemma const {k} (c : V) : ℌ.BoldfaceFunction (fun _ : Fin k → V ↦ c) := + .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | #0 = &c” (by simp), by intro v; simp⟩ + +@[simp] lemma term_retraction (t : Semiterm ℒₒᵣ V n) (e : Fin n → Fin k) : + ℌ.BoldfaceFunction fun v : Fin k → V ↦ Semiterm.valm V (fun x ↦ v (e x)) id t := + .of_zero (Γ' := 𝚺) + ⟨.mkSigma “x | x = !!(Rew.substs (fun x ↦ #(e x).succ) t)” (by simp), by intro v; simp [Semiterm.val_substs]⟩ + +@[simp] lemma term (t : Semiterm ℒₒᵣ V k) : + ℌ.BoldfaceFunction fun v : Fin k → V ↦ Semiterm.valm V v id t := + .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!(Rew.bShift t)” (by simp), by intro v; simp [Semiterm.val_bShift']⟩ + +lemma of_eq {f : (Fin k → V) → V} (g) (h : ∀ v, f v = g v) (H : ℌ.BoldfaceFunction f) : ℌ.BoldfaceFunction g := by + rwa [show g = f from by funext v; simp [h]] + +lemma retraction {n k} {f : (Fin k → V) → V} (hf : ℌ.BoldfaceFunction f) (e : Fin k → Fin n) : + ℌ.BoldfaceFunction fun v ↦ f (fun i ↦ v (e i)) := by + have := Boldface.retraction (n := n + 1) hf (0 :> fun i ↦ (e i).succ); simp at this + exact this.of_iff (by intro x; simp) + +lemma retractiont {f : (Fin k → V) → V} (hf : ℌ.BoldfaceFunction f) (t : Fin k → Semiterm ℒₒᵣ V n) : + ℌ.BoldfaceFunction fun v ↦ f (fun i ↦ Semiterm.valm V v id (t i)) := by + have := Boldface.retractiont (n := n + 1) hf (#0 :> fun i ↦ Rew.bShift (t i)); simp at this + exact this.of_iff (by intro x; simp [Semiterm.val_bShift']) + +lemma rel {f : (Fin k → V) → V} (h : ℌ.BoldfaceFunction f) : + ℌ.Boldface (fun v ↦ v 0 = f (v ·.succ)) := h + +@[simp] lemma nth (ℌ : HierarchySymbol) (i : Fin k) : ℌ.BoldfaceFunction fun w : Fin k → V ↦ w i := by + apply Boldface.of_zero (Γ' := 𝚺) + exact ⟨.mkSigma “x | x = #i.succ” (by simp), by intro v; simp⟩ + +lemma substitution {f : Fin k → (Fin l → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction F) (hf : ∀ i, 𝚺-[m + 1].BoldfaceFunction (f i)) : + Γ-[m + 1].BoldfaceFunction fun z ↦ F (fun i ↦ f i z) := by + simpa using Boldface.substitution (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF <| by + intro i + cases' i using Fin.cases with i + · simp + · simpa using Boldface.retraction (hf i) (0 :> (·.succ.succ)) + +end BoldfaceFunction + +lemma BoldfaceFunction₁.comp {k} {F : V → V} {f : (Fin k → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction₁ F) (hf : 𝚺-[m + 1].BoldfaceFunction f) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ F (f v)) := + BoldfaceFunction.substitution (f := ![f]) hF (by simp [hf]) + +lemma BoldfaceFunction₂.comp {k} {F : V → V → V} {f₁ f₂ : (Fin k → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction₂ F) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v)) := + BoldfaceFunction.substitution (f := ![f₁, f₂]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma BoldfaceFunction₃.comp {k} {F : V → V → V → V} {f₁ f₂ f₃ : (Fin k → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction₃ F) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v)) := + BoldfaceFunction.substitution (f := ![f₁, f₂, f₃]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma BoldfaceFunction₄.comp {k} {F : V → V → V → V → V} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction₄ F) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := + BoldfaceFunction.substitution (f := ![f₁, f₂, f₃, f₄]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma BoldfaceFunction₅.comp {k} {F : V → V → V → V → V → V} {f₁ f₂ f₃ f₄ f₅ : (Fin k → V) → V} + (hF : Γ-[m + 1].BoldfaceFunction₅ F) + (hf₁ : 𝚺-[m + 1].BoldfaceFunction f₁) (hf₂ : 𝚺-[m + 1].BoldfaceFunction f₂) + (hf₃ : 𝚺-[m + 1].BoldfaceFunction f₃) (hf₄ : 𝚺-[m + 1].BoldfaceFunction f₄) + (hf₅ : 𝚺-[m + 1].BoldfaceFunction f₅) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := + BoldfaceFunction.substitution (f := ![f₁, f₂, f₃, f₄, f₅]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +namespace BoldfaceFunction + +lemma comp₁ {k} {f : V → V} [Γ-[m + 1].BoldfaceFunction₁ f] + {g : (Fin k → V) → V} (hg : 𝚺-[m + 1].BoldfaceFunction g) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ f (g v)) := + BoldfaceFunction₁.comp inferInstance hg + +lemma comp₂{k} {f : V → V → V} [Γ-[m + 1].BoldfaceFunction₂ f] + {g₁ g₂ : (Fin k → V) → V} (hg₁ : 𝚺-[m + 1].BoldfaceFunction g₁) (hg₂ : 𝚺-[m + 1].BoldfaceFunction g₂) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ f (g₁ v) (g₂ v)) := + BoldfaceFunction₂.comp inferInstance hg₁ hg₂ + +lemma comp₃ {k} {f : V → V → V → V} [Γ-[m + 1].BoldfaceFunction₃ f] + {g₁ g₂ g₃ : (Fin k → V) → V} + (hg₁ : 𝚺-[m + 1].BoldfaceFunction g₁) (hg₂ : 𝚺-[m + 1].BoldfaceFunction g₂) (hg₃ : 𝚺-[m + 1].BoldfaceFunction g₃) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := + BoldfaceFunction₃.comp inferInstance hg₁ hg₂ hg₃ + +lemma comp₄ {k} {f : V → V → V → V → V} [Γ-[m + 1].BoldfaceFunction₄ f] + {g₁ g₂ g₃ g₄ : (Fin k → V) → V} + (hg₁ : 𝚺-[m + 1].BoldfaceFunction g₁) (hg₂ : 𝚺-[m + 1].BoldfaceFunction g₂) + (hg₃ : 𝚺-[m + 1].BoldfaceFunction g₃) (hg₄ : 𝚺-[m + 1].BoldfaceFunction g₄) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := + BoldfaceFunction₄.comp inferInstance hg₁ hg₂ hg₃ hg₄ + +lemma comp₅ {k} {f : V → V → V → V → V → V} [Γ-[m + 1].BoldfaceFunction₅ f] + {g₁ g₂ g₃ g₄ g₅ : (Fin k → V) → V} + (hg₁ : 𝚺-[m + 1].BoldfaceFunction g₁) (hg₂ : 𝚺-[m + 1].BoldfaceFunction g₂) + (hg₃ : 𝚺-[m + 1].BoldfaceFunction g₃) (hg₄ : 𝚺-[m + 1].BoldfaceFunction g₄) + (hg₅ : 𝚺-[m + 1].BoldfaceFunction g₅) : + Γ-[m + 1].BoldfaceFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := + BoldfaceFunction₅.comp inferInstance hg₁ hg₂ hg₃ hg₄ hg₅ + +end BoldfaceFunction + +namespace Boldface + +lemma ball_lt {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∀ x < f v, P v x) := by + rcases hf with ⟨bf, hbf⟩ + rcases h with ⟨p, hp⟩ + match Γ with + | 𝚺 => exact + ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚷 => exact + ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚫 => + exact .of_sigma_of_pi + ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ + ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ + +lemma bex_lt {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∃ x < f v, P v x) := by + rcases hf with ⟨bf, hbf⟩ + rcases h with ⟨p, hp⟩ + match Γ with + | 𝚺 => exact + ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚷 => exact + ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ + | 𝚫 => + exact .of_sigma_of_pi + ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ + ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), + by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ + +lemma ball_le {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∀ x ≤ f v, P v x) := by + have : Γ-[m + 1].Boldface (fun v ↦ ∀ x < f v + 1, P v x) := ball_lt (BoldfaceFunction₂.comp (by simp) hf (by simp)) h + exact this.of_iff <| by intro v; simp [lt_succ_iff_le] + +lemma bex_le {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∃ x ≤ f v, P v x) := by + have : Γ-[m + 1].Boldface (fun v ↦ ∃ x < f v + 1, P v x) := bex_lt (BoldfaceFunction₂.comp (by simp) hf (by simp)) h + exact this.of_iff <| by intro v; simp [lt_succ_iff_le] + +lemma ball_lt' {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∀ {x}, x < f v → P v x) := ball_lt hf h + +lemma ball_le' {Γ} {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∀ {x}, x ≤ f v → P v x) := ball_le hf h + +end Boldface + +end + +end Arith.HierarchySymbol + +end LO.FirstOrder diff --git a/Arithmetization/Definability/BooleanSystem/Basic.lean b/Arithmetization/Definability/BooleanSystem/Basic.lean deleted file mode 100644 index 96ff21c..0000000 --- a/Arithmetization/Definability/BooleanSystem/Basic.lean +++ /dev/null @@ -1,417 +0,0 @@ -import Arithmetization.Definability.Basic -import Arithmetization.Definability.Init - -/-! - -# Boolean System - - --/ - -namespace LO - -variable {V : Type*} - -structure BooleanSystem (V : Type*) where - VecPr : {k : ℕ} → ((Fin k → V) → Prop) → Prop - verum : VecPr fun _ : Fin k → V ↦ ⊤ - and {P Q : (Fin k → V) → Prop} : VecPr P → VecPr Q → VecPr fun v ↦ P v ∧ Q v - not {P : (Fin k → V) → Prop} : VecPr P → VecPr fun v ↦ ¬P v - equal : VecPr fun v : Fin 2 → V ↦ v 0 = v 1 - replace {k l} {P : (Fin k → V) → Prop} (hP : VecPr P) (f : Fin k → Fin l) : VecPr fun v ↦ P fun i ↦ v (f i) - -namespace BooleanSystem - -variable {𝔅 : BooleanSystem V} {P Q : (Fin k → V) → Prop} - -lemma of_iff (hP : 𝔅.VecPr P) (h : ∀ x, P x ↔ Q x) : 𝔅.VecPr Q := by - have : P = Q := funext fun x ↦ by simp [h] - rcases this; exact hP - -lemma of_not (h : 𝔅.VecPr fun v ↦ ¬P v) : 𝔅.VecPr P := by simpa using 𝔅.not h - -lemma falsum : 𝔅.VecPr fun _ : Fin k → V ↦ ⊥ := of_not <| by simpa using 𝔅.verum - -@[simp] lemma constant (P : Prop) : 𝔅.VecPr fun _ : Fin k → V ↦ P := by - by_cases h : P <;> simp [h] - · exact 𝔅.verum - · exact 𝔅.falsum - -lemma or (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ∨ Q v := of_not <| by - simp; apply 𝔅.and - · apply 𝔅.not hP - · apply 𝔅.not hQ - -lemma imply (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v → Q v := by - simp only [imp_iff_not_or]; apply or - · apply 𝔅.not hP - · exact hQ - -lemma iff (hP : 𝔅.VecPr P) (hQ : 𝔅.VecPr Q) : 𝔅.VecPr fun v ↦ P v ↔ Q v := by - simp only [iff_iff_implies_and_implies] - exact 𝔅.and (imply hP hQ) (imply hQ hP) - -lemma conj {P : Fin l → (Fin k → V) → Prop} (hP : ∀ i, 𝔅.VecPr (P i)) : 𝔅.VecPr fun v ↦ ∀ i, P i v := by - induction l - case zero => simp - case succ l ih => - simp [forall_fin_iff_zero_and_forall_succ] - apply and - · exact hP 0 - · exact ih (fun i ↦ hP i.succ) - -lemma equal' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i = v j := by - simpa using 𝔅.replace 𝔅.equal ![i, j] - -variable (𝔅) - -class Quantifer where - all {k} {P : (Fin k → V) → V → Prop} : 𝔅.VecPr (fun x ↦ P (x ·.succ) (x 0)) → 𝔅.VecPr fun x ↦ ∀ z, P x z - ex {k} {P : (Fin k → V) → V → Prop} : 𝔅.VecPr (fun x ↦ P (x ·.succ) (x 0)) → 𝔅.VecPr fun x ↦ ∃ z, P x z - -variable {𝔅} - -end BooleanSystem - -namespace Arith - -variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] - -variable (V) - -structure BoundedSystem extends BooleanSystem V where - Polynomial : {k : ℕ} → ((Fin k → V) → V) → Prop - polynomial_comp {k} (F : (Fin k → V) → V) (fs : Fin k → (Fin l → V) → V) : - Polynomial F → (∀ i, Polynomial (fs i)) → Polynomial (fun v ↦ F (fun i ↦ fs i v)) - polynomial_replace {p : (Fin k → V) → V} (hp : Polynomial p) (f : Fin k → Fin l) : Polynomial (fun v ↦ p (fun i ↦ v (f i))) - polynomial_nth (i : Fin k) : Polynomial (· i) - polynomial_monotone {p : (Fin k → V) → V} (h : Polynomial p) {v w} : v ≤ w → p v ≤ p w - ball_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} : - Polynomial f → VecPr (fun v ↦ P (v ·.succ) (v 0)) → VecPr fun v ↦ ∀ x ≤ f v, P v x - lessThan : VecPr fun v : Fin 2 → V ↦ v 0 < v 1 - -variable {V} - -namespace BoundedSystem - -open LO.BooleanSystem - -variable {𝔅 : Arith.BoundedSystem V} {P Q : (Fin k → V) → Prop} - -variable (𝔅) - -abbrev Pred (P : V → Prop) : Prop := 𝔅.VecPr (k := 1) fun v ↦ P (v 0) - -abbrev Rel (R : V → V → Prop) : Prop := 𝔅.VecPr (k := 2) fun v ↦ R (v 0) (v 1) - -abbrev Rel₃ (R : V → V → V → Prop) : Prop := 𝔅.VecPr (k := 3) fun v ↦ R (v 0) (v 1) (v 2) - -abbrev Rel₄ (R : V → V → V → V → Prop) : Prop := 𝔅.VecPr (k := 4) fun v ↦ R (v 0) (v 1) (v 2) (v 3) - -abbrev UBVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) - -abbrev UBConstant (c : V) : Prop := 𝔅.UBVecFunc (k := 0) fun _ ↦ c - -abbrev UBFunction (f : V → V) : Prop := 𝔅.UBVecFunc (k := 1) fun v ↦ f (v 0) - -abbrev UBFunction₂ (f : V → V → V) : Prop := 𝔅.UBVecFunc (k := 2) fun v ↦ f (v 0) (v 1) - -abbrev UBFunction₃ (f : V → V → V → V) : Prop := 𝔅.UBVecFunc (k := 3) fun v ↦ f (v 0) (v 1) (v 2) - -def BoundedVecFunc {k} (f : (Fin k → V) → V) : Prop := 𝔅.VecPr (Function.Graphᵥ f) ∧ ∃ p, 𝔅.Polynomial p ∧ f ≤ p - -abbrev BoundedConstant (c : V) : Prop := 𝔅.BoundedVecFunc (k := 0) fun _ ↦ c - -abbrev BoundedFunction (f : V → V) : Prop := 𝔅.BoundedVecFunc (k := 1) fun v ↦ f (v 0) - -abbrev BoundedFunction₂ (f : V → V → V) : Prop := 𝔅.BoundedVecFunc (k := 2) fun v ↦ f (v 0) (v 1) - -abbrev BoundedFunction₃ (f : V → V → V → V) : Prop := 𝔅.BoundedVecFunc (k := 3) fun v ↦ f (v 0) (v 1) (v 2) - -variable {𝔅} - -lemma BoundedVecFunc.vecPr {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : 𝔅.VecPr (Function.Graphᵥ f) := h.1 - -lemma BoundedVecFunc.le_poly {k} {f : (Fin k → V) → V} (h : 𝔅.BoundedVecFunc f) : ∃ p, 𝔅.Polynomial p ∧ f ≤ p := h.2 - -lemma UBVecFunc.boundedVecFunc_of_le {f g : (Fin k → V) → V} (hf : 𝔅.UBVecFunc f) (hg : 𝔅.BoundedVecFunc g) - (h : f ≤ g) : 𝔅.BoundedVecFunc f := by - constructor - · exact hf - · rcases hg.le_poly with ⟨p, hp, hgp⟩ - exact ⟨p, hp, le_trans h hgp⟩ - -lemma UBFunction.boundedVecFunc_of_le {f g : V → V} (hf : 𝔅.UBFunction f) (hg : 𝔅.BoundedFunction g) (h : f ≤ g) : 𝔅.BoundedFunction f := - UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _) - -lemma UBFunction₂.boundedVecFunc_of_le {f g : V → V → V} (hf : 𝔅.UBFunction₂ f) (hg : 𝔅.BoundedFunction₂ g) (h : f ≤ g) : 𝔅.BoundedFunction₂ f := - UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _ _) - -lemma UBFunction₃.boundedVecFunc_of_le {f g : V → V → V → V} (hf : 𝔅.UBFunction₃ f) (hg : 𝔅.BoundedFunction₃ g) (h : f ≤ g) : 𝔅.BoundedFunction₃ f := - UBVecFunc.boundedVecFunc_of_le hf hg (by intro v; simpa using h _ _ _) - -@[simp] lemma BoundedVecFunc.nth (i : Fin k) : 𝔅.BoundedVecFunc (· i) := by - constructor - · apply equal' - · exact ⟨(· i), 𝔅.polynomial_nth i, by simp⟩ - -lemma BoundedVecFunc.replace {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) (c : Fin k → Fin l) : - 𝔅.BoundedVecFunc (fun v ↦ f (fun i ↦ v (c i))) := by - constructor - · apply of_iff (𝔅.replace (l := l + 1) hf.vecPr (0 :> fun x ↦ (c x).succ)) <| by - intro v; simp [Function.Graphᵥ] - · rcases hf.le_poly with ⟨p, pp, hfp⟩ - refine ⟨fun v ↦ p (fun i ↦ v (c i)), by apply 𝔅.polynomial_replace pp, by intro v; simpa using hfp _⟩ - -lemma lessThan' (i j : Fin k) : 𝔅.VecPr fun v ↦ v i < v j := by - simpa using 𝔅.replace 𝔅.lessThan ![i, j] - -lemma lessThanOrEq (i j : Fin k) : 𝔅.VecPr fun v ↦ v i ≤ v j := by - simp [le_iff_lt_or_eq] - apply 𝔅.or - · apply lessThan' - · apply equal' - -lemma BoundedQuantifier.bex_poly {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (pf : Polynomial 𝔅 f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by - simp only [not_exists, not_and]; exact 𝔅.ball_poly pf (𝔅.not hP) - -lemma BoundedQuantifier.bex_vec_poly {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} - (pp : ∀ i, Polynomial 𝔅 (p i)) (hP : 𝔅.VecPr fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : - 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), P v w := by - induction l generalizing k - case zero => simpa [Matrix.empty_eq (α := V)] using hP - case succ l ih => - simp only [exists_le_vec_iff_exists_le_exists_vec] - apply bex_poly (pp 0) - apply ih - · intro i; apply 𝔅.polynomial_replace (pp i.succ) - · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (𝔅.replace hP g) <| by - intro v; simp [g] - apply iff_of_eq; congr - · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · ext i - cases' i using Fin.cases with i - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - -lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔅.VecPr P) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : - 𝔅.VecPr fun z ↦ P (fun i ↦ f i z) := by - choose p hp using fun i ↦ (hf i).le_poly - have : 𝔅.VecPr fun v ↦ ∃ w ≤ (p · v), (∀ i, w i = f i v) ∧ P w := by - apply BoundedQuantifier.bex_vec_poly (fun i ↦ (hp i).1) - apply 𝔅.and - · apply conj; intro i - simpa using 𝔅.replace (hf i).vecPr (i.natAdd l :> Fin.castAdd k) - · apply 𝔅.replace hP - apply of_iff this <| by - intro v; constructor - · rintro ⟨w, hw, e, h⟩ - rcases funext e - exact h - · intro h; exact ⟨(f · v), by intro i; simpa using (hp i).2 v, by simp, h⟩ - -lemma BoundedVecFunc.substitution {F : (Fin k → V) → V} {f : Fin k → (Fin l → V) → V} - (hF : 𝔅.BoundedVecFunc F) (hf : ∀ i, 𝔅.BoundedVecFunc (f i)) : - 𝔅.BoundedVecFunc fun v ↦ F (fun i ↦ f i v) := by - constructor - · simpa [Function.Graphᵥ] using - BoundedSystem.substitution (l := l + 1) hF.vecPr (f := (· 0) :> fun i v ↦ f i (v ·.succ)) - (by intro i; cases' i using Fin.cases with i - · simp - · simpa using BoundedVecFunc.replace (hf i) _) - · rcases hF.le_poly with ⟨p, hp, hFp⟩ - choose ps hps using fun i ↦ (hf i).le_poly - refine ⟨fun v ↦ p fun i ↦ ps i v, 𝔅.polynomial_comp p ps hp (fun i ↦ (hps i).1), ?_⟩ - intro v; exact le_trans (hFp (f · v)) (𝔅.polynomial_monotone hp (fun i ↦ (hps i).2 v)) - -lemma ball_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∀ x ≤ f v, P v x := by - rcases hf.le_poly with ⟨p, hp, hfp⟩ - have : 𝔅.VecPr fun v ↦ ∀ x ≤ p v, x ≤ f v → P v x := by - apply 𝔅.ball_poly hp - apply imply - · simpa using substitution (𝔅.lessThanOrEq 0 1) - (f := ![(· 0), fun v ↦ f (v ·.succ)]) - (by simpa [forall_fin_iff_zero_and_forall_succ] using hf.replace Fin.succ) - · exact hP - exact of_iff this <| by - intro v; constructor - · intro h x hx - exact h x (le_trans hx (hfp v)) hx - · intro h x _ hx - exact h x hx - -lemma bex_le {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x ≤ f v, P v x := of_not <| by - simp only [not_exists, not_and] - exact ball_le hf (𝔅.not hP) - -lemma ball_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∀ x < f v, P v x := by - have : 𝔅.VecPr fun v ↦ ∀ x ≤ f v, x < f v → P v x := by - apply ball_le hf - apply imply ?_ hP - simpa using substitution (𝔅.lessThan' 0 1) - (f := ![(· 0), fun v ↦ f (v ·.succ)]) - (by simpa [forall_fin_iff_zero_and_forall_succ] using hf.replace Fin.succ) - exact of_iff this <| by - intro v; constructor - · intro h x hx - exact h x (le_of_lt hx) hx - · intro h x _ hx - exact h x hx - -lemma bex_lt {k} {f : (Fin k → V) → V} {P : (Fin k → V) → V → Prop} - (hf : 𝔅.BoundedVecFunc f) (hP : 𝔅.VecPr (fun v ↦ P (v ·.succ) (v 0))) : - 𝔅.VecPr fun v ↦ ∃ x < f v, P v x := of_not <| by - simp only [not_exists, not_and] - exact ball_lt hf (𝔅.not hP) - -variable (𝔅) - -abbrev FactVecPr {k} (P : (Fin k → V) → Prop) : Prop := Fact (𝔅.VecPr P) - -abbrev IsPred (P : V → Prop) : Prop := Fact (𝔅.Pred P) - -abbrev IsRel (R : V → V → Prop) : Prop := Fact (𝔅.Rel R) - -abbrev IsRel₃ (R : V → V → V → Prop) : Prop := Fact (𝔅.Rel₃ R) - -abbrev IsBoundedConstant (c : V) : Prop := Fact (𝔅.BoundedConstant c) - -abbrev IsBoundedFunction (f : V → V) : Prop := Fact (𝔅.BoundedFunction f) - -abbrev IsBoundedFunction₂ (f : V → V → V) : Prop := Fact (𝔅.BoundedFunction₂ f) - -abbrev IsBoundedFunction₃ (f : V → V → V → V) : Prop := Fact (𝔅.BoundedFunction₃ f) - -variable {𝔅} - -instance : 𝔅.IsRel (· = ·) := ⟨equal' 0 1⟩ - -instance : 𝔅.IsRel (· < ·) := ⟨lessThan' 0 1⟩ - -instance : 𝔅.IsRel (· ≤ ·) := ⟨lessThanOrEq 0 1⟩ - -lemma Pred.comp {P : V → Prop} [hP : 𝔅.IsPred P] {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) : - 𝔅.VecPr fun v ↦ P (f v) := by - simpa using substitution hP.out (f := ![f]) (by simp [hf]) - -lemma Rel.comp {R : V → V → Prop} [hR : 𝔅.IsRel R] {f₁ f₂ : (Fin k → V) → V} (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) : - 𝔅.VecPr fun v ↦ R (f₁ v) (f₂ v) := by - simpa using substitution hR.out (f := ![f₁, f₂]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂]) - -lemma Rel₃.comp {R : V → V → V → Prop} [hR : 𝔅.IsRel₃ R] {f₁ f₂ f₃ : (Fin k → V) → V} - (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) (hf₃ : 𝔅.BoundedVecFunc f₃) : - 𝔅.VecPr fun v ↦ R (f₁ v) (f₂ v) (f₃ v) := by - simpa using substitution hR.out (f := ![f₁, f₂, f₃]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃]) - -@[simp] lemma Constant.comp (c : V) [hc : 𝔅.IsBoundedConstant c] : - 𝔅.BoundedVecFunc fun _ : Fin k → V ↦ c := - BoundedVecFunc.substitution (l := k) hc.out (f := ![]) (by simp) - -lemma Function.comp {F : V → V} [hF : 𝔅.IsBoundedFunction F] {f : (Fin k → V) → V} (hf : 𝔅.BoundedVecFunc f) : - 𝔅.BoundedVecFunc fun v ↦ F (f v) := by - simpa using BoundedVecFunc.substitution hF.out (f := ![f]) (by simp [hf]) - -lemma Function₂.comp {F : V → V → V} [hF : 𝔅.IsBoundedFunction₂ F] {f₁ f₂ : (Fin k → V) → V} - (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) : - 𝔅.BoundedVecFunc fun v ↦ F (f₁ v) (f₂ v) := by - simpa using BoundedVecFunc.substitution hF.out (f := ![f₁, f₂]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂]) - -lemma Function₃.comp {F : V → V → V → V} [hF : 𝔅.IsBoundedFunction₃ F] {f₁ f₂ f₃ : (Fin k → V) → V} - (hf₁ : 𝔅.BoundedVecFunc f₁) (hf₂ : 𝔅.BoundedVecFunc f₂) (hf₃ : 𝔅.BoundedVecFunc f₃) : - 𝔅.BoundedVecFunc fun v ↦ F (f₁ v) (f₂ v) (f₃ v) := by - simpa using BoundedVecFunc.substitution hF.out (f := ![f₁, f₂, f₃]) (by simp [forall_fin_iff_zero_and_forall_succ, hf₁, hf₂, hf₃]) - -variable (𝔅) - -class Arithmetical where - zero : 𝔅.BoundedConstant 0 - one : 𝔅.BoundedConstant 1 - add : 𝔅.BoundedFunction₂ (· + ·) - mul : 𝔅.BoundedFunction₂ (· * ·) - -variable {𝔅} - -section Arithmetical - -variable [𝔅.Arithmetical] - -instance : 𝔅.IsBoundedConstant 0 := ⟨Arithmetical.zero⟩ - -instance : 𝔅.IsBoundedConstant 1 := ⟨Arithmetical.one⟩ - -instance : 𝔅.IsBoundedFunction₂ (· + ·) := ⟨Arithmetical.add⟩ - -instance : 𝔅.IsBoundedFunction₂ (· * ·) := ⟨Arithmetical.mul⟩ - -instance (n : ℕ) : 𝔅.IsBoundedConstant n := ⟨by - induction n - case zero => simp - case succ n ih => - simpa using Function₂.comp ih (by simp)⟩ - -end Arithmetical - -variable (𝔅) - -class Boldface where - const (z : V) : 𝔅.BoundedVecFunc (k := 0) fun _ ↦ z - -variable {𝔅} - -instance [𝔅.Boldface] (z : V) : 𝔅.IsBoundedConstant z := ⟨Boldface.const z⟩ - -end BoundedSystem - -section - -open Lean.Parser.Tactic (config) - -attribute [aesop (rule_sets := [Definability]) norm] - sq - Arith.pow_three - pow_four - -attribute [aesop 5 (rule_sets := [Definability]) safe] - BoundedSystem.Function.comp - BoundedSystem.Function₂.comp - BoundedSystem.Function₃.comp - BoundedSystem.Pred.comp - BoundedSystem.Rel.comp - BoundedSystem.Rel₃.comp - -attribute [aesop 8 (rule_sets := [Definability]) safe] - BoundedSystem.ball_le - BoundedSystem.bex_le - BoundedSystem.ball_lt - BoundedSystem.bex_lt - -attribute [aesop 10 (rule_sets := [Definability]) safe] - BooleanSystem.not - BooleanSystem.imply - BooleanSystem.iff - -attribute [aesop 11 (rule_sets := [Definability]) safe] - BooleanSystem.and - BooleanSystem.or - -macro "definability" : attr => - `(attr|aesop 10 (rule_sets := [$(Lean.mkIdent `Definability):ident]) safe) - -macro "definability" (config)? : tactic => - `(tactic| aesop (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) - -macro "definability?" (config)? : tactic => - `(tactic| aesop? (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) - -end - -end LO.Arith diff --git a/Arithmetization/Definability/BooleanSystem/DeltaZero.lean b/Arithmetization/Definability/BooleanSystem/DeltaZero.lean deleted file mode 100644 index 7214c53..0000000 --- a/Arithmetization/Definability/BooleanSystem/DeltaZero.lean +++ /dev/null @@ -1,104 +0,0 @@ -import Arithmetization.Definability.BooleanSystem.Open - -namespace LO.Arith - -open FirstOrder - -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] - -def DeltaZero.Lightface : BoundedSystem V where - VecPr {k} (R) := Definable ℒₒᵣ (Arith.Hierarchy 𝚺 0) R - verum {k} := ⟨⊤, by simp, by intro v; simp⟩ - and {k P Q} := by - rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ - refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ - not {k P} := by - rintro ⟨p, hp, hP⟩ - exact ⟨~p, by simp [Arith.Hierarchy.pi_zero_iff_sigma_zero, hp], by intro v; simp [hP v]⟩ - equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ - replace {k l P} := by - rintro ⟨p, hp, hP⟩ f - refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ - Polynomial {k} (f) := ∃ t : Semiterm ℒₒᵣ Empty k, f = t.valbm V - polynomial_comp {l k F f} := by - rintro ⟨T, rfl⟩ ht - choose t ht using ht - rcases funext ht - exact ⟨Rew.substs t T, by ext v; simp [Semiterm.val_substs]⟩ - polynomial_replace {k l _} := by - rintro ⟨t, rfl⟩ f; exact ⟨Rew.substs (fun x ↦ #(f x)) t, by ext v; simp [Semiterm.val_substs]⟩ - polynomial_monotone {k _} := by - rintro ⟨t, rfl⟩ v w h - apply Structure.Monotone.term_monotone - · exact h - · simp - polynomial_nth {k i} := ⟨#i, by simp⟩ - ball_poly {k f P} := by - rintro ⟨t, rfl⟩ ⟨p, hp, hP⟩ - exact ⟨“∀ x <⁺ !t ⋯, !p x ⋯”, by simp [hp], by intro v; simp [Semiterm.val_substs, hP.iff]⟩ - lessThan := ⟨“x y | x < y”, by simp, by intro v; simp⟩ - -notation "𝛥₀[" V "]" => DeltaZero.Lightface (V := V) - -notation "𝛥₀" => DeltaZero.Lightface - -instance : 𝛥₀[V].Arithmetical where - zero := ⟨⟨“x | x = 0”, by simp, by intro; simp [Function.Graphᵥ]⟩, 0, ⟨‘0’, by ext v; simp⟩, by intro; simp⟩ - one := ⟨⟨“x | x = 1”, by simp, by intro; simp [Function.Graphᵥ]⟩, 1, ⟨‘1’, by ext v; simp⟩, by intro; simp⟩ - add := ⟨⟨“z x y | z = x + y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 + v 1, ⟨‘#0 + #1’, by ext v; simp⟩, by simp⟩ - mul := ⟨⟨“z x y | z = x * y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 * v 1, ⟨‘#0 * #1’, by ext v; simp⟩, by simp⟩ - -def DeltaZero.Boldface : BoundedSystem V where - VecPr {k} (R) := DefinableWithParam ℒₒᵣ (Arith.Hierarchy 𝚺 0) R - verum {k} := ⟨⊤, by simp, by intro v; simp⟩ - and {k P Q} := by - rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ - refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ - not {k P} := by - rintro ⟨p, hp, hP⟩ - exact ⟨~p, by simp [Arith.Hierarchy.pi_zero_iff_sigma_zero, hp], by intro v; simp [hP v]⟩ - equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ - replace {k l P} := by - rintro ⟨p, hp, hP⟩ f - refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ - Polynomial {k} (f) := ∃ t : Semiterm ℒₒᵣ V k, f = fun v ↦ t.valm V v id - polynomial_comp {l k F f} := by - rintro ⟨T, rfl⟩ ht - choose t ht using ht - rcases funext ht - exact ⟨Rew.substs t T, by ext v; simp [Semiterm.val_substs]⟩ - polynomial_replace {k l _} := by - rintro ⟨t, rfl⟩ f; exact ⟨Rew.substs (fun x ↦ #(f x)) t, by ext v; simp [Semiterm.val_substs]⟩ - polynomial_monotone {k _} := by - rintro ⟨t, rfl⟩ v w h - apply Structure.Monotone.term_monotone - · exact h - · simp - polynomial_nth {k i} := ⟨#i, by simp⟩ - ball_poly {k f P} := by - rintro ⟨t, rfl⟩ ⟨p, hp, hP⟩ - exact ⟨“∀ x <⁺ !t ⋯, !p x ⋯”, by simp [hp], by intro v; simp [Semiterm.val_substs, hP.iff]⟩ - lessThan := ⟨“x y | x < y”, by simp, by intro v; simp⟩ - -notation "𝜟₀" => DeltaZero.Boldface - -notation "𝜟₀[" V "]" => DeltaZero.Boldface (V := V) - -instance : 𝜟₀[V].Arithmetical where - zero := ⟨⟨“x | x = 0”, by simp, by intro; simp [Function.Graphᵥ]⟩, 0, ⟨‘0’, by ext v; simp⟩, by intro; simp⟩ - one := ⟨⟨“x | x = 1”, by simp, by intro; simp [Function.Graphᵥ]⟩, 1, ⟨‘1’, by ext v; simp⟩, by intro; simp⟩ - add := ⟨⟨“z x y | z = x + y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 + v 1, ⟨‘#0 + #1’, by ext v; simp⟩, by simp⟩ - mul := ⟨⟨“z x y | z = x * y”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun v ↦ v 0 * v 1, ⟨‘#0 * #1’, by ext v; simp⟩, by simp⟩ - -instance : 𝜟₀[V].Boldface where - const (z) := ⟨⟨“x | x = &z”, by simp, by intro v; simp [Function.Graphᵥ]⟩, fun _ ↦ z, ⟨&z, by simp⟩, by simp⟩ - -section - -example : 𝛥₀.Rel₃ fun x y z : V ↦ x = 32 → ¬y < z ∧ ∀ w < x + 2 * z, w < x := by definability? - -example (c : V) : 𝜟₀.Rel₃ fun x y z : V ↦ x = 32 → ¬y < z ∧ ∀ w < x + z, c < x := by definability - -end - -namespace LO.Arith diff --git a/Arithmetization/Definability/BooleanSystem/Open.lean b/Arithmetization/Definability/BooleanSystem/Open.lean deleted file mode 100644 index 770c4a3..0000000 --- a/Arithmetization/Definability/BooleanSystem/Open.lean +++ /dev/null @@ -1,37 +0,0 @@ -import Arithmetization.Definability.BooleanSystem.Basic - -namespace LO.Arith - -open FirstOrder - -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] - -def Open.Lightface : BooleanSystem V where - VecPr {k} (R) := Definable ℒₒᵣ Semiformula.Open R - verum {k} := ⟨⊤, by simp, by intro v; simp⟩ - and {k P Q} := by - rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ - refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ - not {k P} := by - rintro ⟨p, hp, hP⟩ - exact ⟨~p, by simp [hp], by intro v; simp [hP v]⟩ - equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ - replace {k l P} := by - rintro ⟨p, hp, hP⟩ f - refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ - -def Open.Boldface : BooleanSystem V where - VecPr {k} (R) := DefinableWithParam ℒₒᵣ Semiformula.Open R - verum {k} := ⟨⊤, by simp, by intro v; simp⟩ - and {k P Q} := by - rintro ⟨p, hp, hP⟩; rintro ⟨q, hq, hQ⟩ - refine ⟨p ⋏ q, by simp [hp, hq], by intro v; simp [hP v, hQ v]⟩ - not {k P} := by - rintro ⟨p, hp, hP⟩ - exact ⟨~p, by simp [hp], by intro v; simp [hP v]⟩ - equal := ⟨“x y | x = y”, by simp, by intro v; simp⟩ - replace {k l P} := by - rintro ⟨p, hp, hP⟩ f - refine ⟨Rew.substs (fun x ↦ #(f x)) |>.hom p, by simp [hp], by intro v; simp [hP.iff]⟩ - -namespace LO.Arith diff --git a/Arithmetization/Definability/BoundedBoldface.lean b/Arithmetization/Definability/BoundedBoldface.lean new file mode 100644 index 0000000..7e05dc3 --- /dev/null +++ b/Arithmetization/Definability/BoundedBoldface.lean @@ -0,0 +1,427 @@ +import Arithmetization.Definability.Boldface +import Arithmetization.Definability.Init + +namespace LO.FirstOrder.Arith + +open LO.Arith + +variable {ξ : Type*} {n : ℕ} + +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐏𝐀⁻] + +variable {ℌ : HierarchySymbol} {Γ Γ' : SigmaPiDelta} + +variable (ℌ) + +class Bounded (f : (Fin k → V) → V) : Prop where + bounded : ∃ t : Semiterm ℒₒᵣ V k, ∀ v : Fin k → V, f v ≤ t.valm V v id + +abbrev Bounded₁ (f : V → V) : Prop := Bounded (k := 1) (fun v ↦ f (v 0)) + +abbrev Bounded₂ (f : V → V → V) : Prop := Bounded (k := 2) (fun v ↦ f (v 0) (v 1)) + +abbrev Bounded₃ (f : V → V → V → V) : Prop := Bounded (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) + +instance (f : (Fin k → V) → V) [h : Bounded f] : Bounded f := by + rcases h with ⟨t, ht⟩ + exact ⟨Semiterm.lMap Language.oringEmb t, by simpa⟩ + +variable {ℌ} + +namespace Bounded + +@[simp] lemma var {k} (i : Fin k) : Bounded fun v : Fin k → V ↦ v i := ⟨#i, by intro _; simp⟩ + +@[simp] lemma const {k} (c : V) : Bounded (fun _ : Fin k → V ↦ c) := ⟨&c, by intro _; simp⟩ + +@[simp] lemma term_retraction (t : Semiterm ℒₒᵣ V n) (e : Fin n → Fin k) : + Bounded fun v : Fin k → V ↦ Semiterm.valm V (fun x ↦ v (e x)) id t := + ⟨Rew.substs (fun x ↦ #(e x)) t, by intro _; simp [Semiterm.val_substs]⟩ + +@[simp] lemma term (t : Semiterm ℒₒᵣ V k) : Bounded fun v : Fin k → V => Semiterm.valm V v id t := + ⟨t, by intro _; simp⟩ + +lemma retraction {f : (Fin k → V) → V} (hf : Bounded f) (e : Fin k → Fin n) : + Bounded fun v ↦ f (fun i ↦ v (e i)) := by + rcases hf with ⟨t, ht⟩ + exact ⟨Rew.substs (fun x ↦ #(e x)) t, by intro; simp [Semiterm.val_substs, ht]⟩ + +lemma comp {k} {f : (Fin l → V) → V} {g : Fin l → (Fin k → V) → V} (hf : Bounded f) (hg : ∀ i, Bounded (g i)) : + Bounded (fun v ↦ f (g · v)) where + bounded := by + rcases hf.bounded with ⟨tf, htf⟩ + choose tg htg using fun i ↦ (hg i).bounded + exact ⟨Rew.substs tg tf, by + intro v; simp [Semiterm.val_substs] + exact le_trans (htf (g · v)) (Structure.Monotone.term_monotone tf (fun i ↦ htg i v) (by simp))⟩ + +end Bounded + +lemma Bounded₁.comp {f : V → V} {k} {g : (Fin k → V) → V} (hf : Bounded₁ f) (hg : Bounded g) : + Bounded (fun v ↦ f (g v)) := Bounded.comp hf (l := 1) (fun _ ↦ hg) + +lemma Bounded₂.comp {f : V → V → V} {k} {g₁ g₂ : (Fin k → V) → V} + (hf : Bounded₂ f) (hg₁ : Bounded g₁) (hg₂ : Bounded g₂) : + Bounded (fun v ↦ f (g₁ v) (g₂ v)) := Bounded.comp hf (g := ![g₁, g₂]) (fun i ↦ by cases i using Fin.cases <;> simp [*]) + +lemma Bounded₃.comp {f : V → V → V → V} {k} {g₁ g₂ g₃ : (Fin k → V) → V} + (hf : Bounded₃ f) (hg₁ : Bounded g₁) (hg₂ : Bounded g₂) (hg₃ : Bounded g₃) : + Bounded (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := Bounded.comp hf (g := ![g₁, g₂, g₃]) + (fun i ↦ by + cases' i using Fin.cases with i <;> simp [*] + cases' i using Fin.cases with i <;> simp [*]) + +namespace Bounded₂ + +instance add : Bounded₂ ((· + ·) : V → V → V) where + bounded := ⟨‘x y | x + y’, by intro _; simp⟩ + +instance mul : Bounded₂ ((· * ·) : V → V → V) where + bounded := ⟨‘x y | x * y’, by intro _; simp⟩ + +instance hAdd : Bounded₂ (HAdd.hAdd : V → V → V) where + bounded := ⟨‘x y | x + y’, by intro _; simp⟩ + +instance hMul : Bounded₂ (HMul.hMul : V → V → V) where + bounded := ⟨‘x y | x * y’, by intro _; simp⟩ + +end Bounded₂ + +def BoldfaceBoundedFunction {k} (f : (Fin k → V) → V) := Bounded f ∧ 𝚺₀.BoldfaceFunction f + +abbrev BoldfaceBoundedFunction₁ (f : V → V) : Prop := BoldfaceBoundedFunction (k := 1) (fun v => f (v 0)) + +abbrev BoldfaceBoundedFunction₂ (f : V → V → V) : Prop := BoldfaceBoundedFunction (k := 2) (fun v => f (v 0) (v 1)) + +abbrev BoldfaceBoundedFunction₃ (f : V → V → V → V) : Prop := BoldfaceBoundedFunction (k := 3) (fun v => f (v 0) (v 1) (v 2)) + +lemma BoldfaceBoundedFunction.bounded {f : (Fin k → V) → V} (h : BoldfaceBoundedFunction f) : Bounded f := h.1 + +lemma BoldfaceBoundedFunction₁.bounded {f : V → V} (h : BoldfaceBoundedFunction₁ f) : Bounded₁ f := h.1 + +lemma BoldfaceBoundedFunction₂.bounded {f : V → V → V} (h : BoldfaceBoundedFunction₂ f) : Bounded₂ f := h.1 + +lemma BoldfaceBoundedFunction₃.bounded {f : V → V → V → V} (h : BoldfaceBoundedFunction₃ f) : Bounded₃ f := h.1 + +lemma BoldfaceBoundedFunction.definable {f : (Fin k → V) → V} (h : BoldfaceBoundedFunction f) : ℌ.BoldfaceFunction f := .of_zero h.2 + +lemma BoldfaceBoundedFunction₁.definable {f : V → V} (h : BoldfaceBoundedFunction₁ f) : ℌ.BoldfaceFunction₁ f := .of_zero h.2 + +lemma BoldfaceBoundedFunction₂.definable {f : V → V → V} (h : BoldfaceBoundedFunction₂ f) : ℌ.BoldfaceFunction₂ f := .of_zero h.2 + +lemma BoldfaceBoundedFunction₃.definable {f : V → V → V → V} (h : BoldfaceBoundedFunction₃ f) : ℌ.BoldfaceFunction₃ f := .of_zero h.2 + +namespace BoldfaceBoundedFunction + +lemma of_polybounded_of_definable (f : (Fin k → V) → V) [hb : Bounded f] [hf : 𝚺₀.BoldfaceFunction f] : + BoldfaceBoundedFunction f := ⟨hb, hf⟩ + +@[simp] lemma of_polybounded_of_definable₁ (f : V → V) [hb : Bounded₁ f] [hf : 𝚺₀.BoldfaceFunction₁ f] : + BoldfaceBoundedFunction₁ f := ⟨hb, hf⟩ + +@[simp] lemma of_polybounded_of_definable₂ (f : V → V → V) [hb : Bounded₂ f] [hf : 𝚺₀.BoldfaceFunction₂ f] : + BoldfaceBoundedFunction₂ f := ⟨hb, hf⟩ + +@[simp] lemma of_polybounded_of_definable₃ (f : V → V → V → V) [hb : Bounded₃ f] [hf : 𝚺₀.BoldfaceFunction₃ f] : + BoldfaceBoundedFunction₃ f := ⟨hb, hf⟩ + +lemma retraction {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (e : Fin k → Fin n) : + BoldfaceBoundedFunction fun v ↦ f (fun i ↦ v (e i)) := ⟨hf.bounded.retraction e, hf.definable.retraction e⟩ + +end BoldfaceBoundedFunction + +namespace HierarchySymbol.Boldface + +variable {P Q : (Fin k → V) → Prop} + +lemma ball_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : + ℌ.Boldface fun v ↦ ∀ x < f v, P v x := by + rcases hf.bounded with ⟨bf, hbf⟩ + rcases hf.definable with ⟨f_graph, hf_graph⟩ + rcases h with ⟨p, hp⟩ + have : ℌ.DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm V v id bf, x = f v ∧ ∀ y < x, P v y) + (HierarchySymbol.Semiformula.bex ‘!!bf + 1’ + (f_graph ⋏ HierarchySymbol.Semiformula.ball (#0) (HierarchySymbol.Semiformula.rew (Rew.substs (#0 :> fun i ↦ #i.succ.succ)) p))) := by + simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).ball #0)).bex ‘!!bf + 1’ + exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) + + +lemma bex_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : + ℌ.Boldface fun v ↦ ∃ x < f v, P v x := by + rcases hf.bounded with ⟨bf, hbf⟩ + rcases hf.definable with ⟨f_graph, hf_graph⟩ + rcases h with ⟨p, hp⟩ + have : ℌ.DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm V v id bf, x = f v ∧ ∃ y < x, P v y) + (HierarchySymbol.Semiformula.bex ‘!!bf + 1’ + (f_graph ⋏ HierarchySymbol.Semiformula.bex (#0) (HierarchySymbol.Semiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by + simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex #0)).bex ‘!!bf + 1’ + exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) + +lemma ball_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : + ℌ.Boldface fun v ↦ ∀ x ≤ f v, P v x := by + rcases hf.bounded with ⟨bf, hbf⟩ + rcases hf.definable with ⟨f_graph, hf_graph⟩ + rcases h with ⟨p, hp⟩ + have : ℌ.DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm V v id bf, x = f v ∧ ∀ y ≤ x, P v y) + (HierarchySymbol.Semiformula.bex ‘!!bf + 1’ + (f_graph ⋏ HierarchySymbol.Semiformula.ball ‘x | x + 1’ (HierarchySymbol.Semiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by + simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).ball ‘x | x + 1’)).bex ‘!!bf + 1’ + exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) + +lemma bex_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : + ℌ.Boldface fun v ↦ ∃ x ≤ f v, P v x := by + rcases hf.bounded with ⟨bf, hbf⟩ + rcases hf.definable with ⟨f_graph, hf_graph⟩ + rcases h with ⟨p, hp⟩ + have : ℌ.DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm V v id bf, x = f v ∧ ∃ y ≤ x, P v y) + (HierarchySymbol.Semiformula.bex ‘!!bf + 1’ + (f_graph ⋏ HierarchySymbol.Semiformula.bex ‘x | x + 1’ (HierarchySymbol.Semiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by + simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex ‘x | x + 1’)).bex ‘!!bf + 1’ + exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) + +lemma ball_lt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : + Γ-[0].Boldface fun v ↦ ∀ x < f v, P v x := ball_lt_boldfaceBoundedFunction hf h + +lemma bex_lt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : + Γ-[0].Boldface fun v ↦ ∃ x < f v, P v x := bex_lt_boldfaceBoundedFunction hf h + +lemma ball_le_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : + Γ-[0].Boldface fun v ↦ ∀ x ≤ f v, P v x := ball_le_boldfaceBoundedFunction hf h + +lemma bex_le_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : + Γ-[0].Boldface fun v ↦ ∃ x ≤ f v, P v x := bex_le_boldfaceBoundedFunction hf h + +lemma bex_vec_le_boldfaceBoundedFunction {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} + (pp : ∀ i, BoldfaceBoundedFunction (p i)) (hP : ℌ.Boldface fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : + ℌ.Boldface fun v ↦ ∃ w ≤ (p · v), P v w := by + induction l generalizing k + case zero => simpa [Matrix.empty_eq (α := V)] using hP + case succ l ih => + simp only [exists_le_vec_iff_exists_le_exists_vec] + apply bex_le_boldfaceBoundedFunction (pp 0) + apply ih + · intro i; apply BoldfaceBoundedFunction.retraction (pp i.succ) + · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) + exact of_iff (retraction hP g) <| by + intro v; simp [g] + apply iff_of_eq; congr + · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · ext i + cases' i using Fin.cases with i + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + · simp; congr 1; ext; simp [Matrix.vecAppend_eq_ite] + +lemma substitution_boldfaceBoundedFunction {f : Fin k → (Fin l → V) → V} + (hP : ℌ.Boldface P) (hf : ∀ i, BoldfaceBoundedFunction (f i)) : + ℌ.Boldface fun z ↦ P (f · z) := by + have : ℌ.Boldface fun v ↦ ∃ w ≤ (f · v), (∀ i, w i = f i v) ∧ P w := by + apply bex_vec_le_boldfaceBoundedFunction hf + apply and + · apply conj; intro i + simpa using retraction (.of_zero (hf i).2) (i.natAdd l :> Fin.castAdd k) + · apply retraction hP + apply of_iff this <| by + intro v; constructor + · intro h; exact ⟨(f · v), by intro i; simp, by simp, h⟩ + · rintro ⟨w, hw, e, h⟩ + rcases funext e + exact h + +end HierarchySymbol.Boldface + +namespace BoldfaceBoundedFunction + +lemma of_iff {f g : (Fin k → V) → V} (H : BoldfaceBoundedFunction f) (h : ∀ v, f v = g v) : BoldfaceBoundedFunction g := by + have : f = g := by funext v; simp [h] + rcases this; exact H + +@[simp] lemma var {k} (i : Fin k) : BoldfaceBoundedFunction (fun v : Fin k → V ↦ v i) := ⟨by simp, by simp⟩ + +@[simp] lemma const {k} (c : V) : BoldfaceBoundedFunction (fun _ : Fin k → V ↦ c) := ⟨by simp, by simp⟩ + +@[simp] lemma term_retraction (t : Semiterm ℒₒᵣ V n) (e : Fin n → Fin k) : + BoldfaceBoundedFunction fun v : Fin k → V ↦ Semiterm.valm V (fun x ↦ v (e x)) id t := ⟨by simp, by simp⟩ + +@[simp] lemma term (t : Semiterm ℒₒᵣ V k) : + BoldfaceBoundedFunction fun v : Fin k → V ↦ Semiterm.valm V v id t := ⟨by simp, by simp⟩ + +end BoldfaceBoundedFunction + +namespace HierarchySymbol.Boldface + +open BoldfaceBoundedFunction + +lemma bcomp₁ {k} {P : V → Prop} {f : (Fin k → V) → V} [hP : Γ-[0].BoldfacePred P] (hf : BoldfaceBoundedFunction f) : + Γ-[0].Boldface fun v ↦ P (f v) := + substitution_boldfaceBoundedFunction (f := ![f]) hP (by simp [*]) + +lemma bcomp₂ {k} {R : V → V → Prop} {f₁ f₂ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : + Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma bcomp₃ {k} {R : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₃ R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) : + Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂, f₃]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma bcomp₄ {k} {R : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₄ R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) (hf₄ : BoldfaceBoundedFunction f₄) : + Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) (f₄ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂, f₃, f₄]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +end HierarchySymbol.Boldface + +lemma HierarchySymbol.BoldfaceFunction.bcomp {k} {F : (Fin l → V) → V} {f : Fin l → (Fin k → V) → V} + (hF : ℌ.BoldfaceFunction F) (hf : ∀ i, BoldfaceBoundedFunction (f i)) : + ℌ.BoldfaceFunction (fun v ↦ F (f · v)) := by + simpa using Boldface.substitution_boldfaceBoundedFunction (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF <| by + intro i + cases' i using Fin.cases with i + · simp + · simpa using BoldfaceBoundedFunction.retraction (hf i) Fin.succ + +lemma HierarchySymbol.BoldfaceFunction₁.bcomp {k} {F : V → V} {f : (Fin k → V) → V} + (hF : ℌ.BoldfaceFunction₁ F) (hf : BoldfaceBoundedFunction f) : + ℌ.BoldfaceFunction (fun v ↦ F (f v)) := + HierarchySymbol.BoldfaceFunction.bcomp (f := ![f]) hF (by simp [*]) + +lemma HierarchySymbol.BoldfaceFunction₂.bcomp {k} {F : V → V → V} {f₁ f₂ : (Fin k → V) → V} + (hF : ℌ.BoldfaceFunction₂ F) + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : + ℌ.BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v)) := + HierarchySymbol.BoldfaceFunction.bcomp (f := ![f₁, f₂]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma HierarchySymbol.BoldfaceFunction₃.bcomp {k} {F : V → V → V → V} {f₁ f₂ f₃ : (Fin k → V) → V} + (hF : ℌ.BoldfaceFunction₃ F) + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) : + ℌ.BoldfaceFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v)) := + HierarchySymbol.BoldfaceFunction.bcomp (f := ![f₁, f₂, f₃]) hF (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma BoldfaceBoundedFunction₁.comp {k} {F : V → V} {f : (Fin k → V) → V} + (hF : BoldfaceBoundedFunction₁ F) (hf : BoldfaceBoundedFunction f) : + BoldfaceBoundedFunction (fun v ↦ F (f v)) := ⟨hF.bounded.comp hf.bounded, hF.definable.bcomp hf⟩ + +lemma BoldfaceBoundedFunction₂.comp {k} {F : V → V → V} {f₁ f₂ : (Fin k → V) → V} + (hF : BoldfaceBoundedFunction₂ F) + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : + BoldfaceBoundedFunction (fun v ↦ F (f₁ v) (f₂ v)) := ⟨hF.bounded.comp hf₁.bounded hf₂.bounded, hF.definable.bcomp hf₁ hf₂⟩ + +lemma BoldfaceBoundedFunction₃.comp {k} {F : V → V → V → V} {f₁ f₂ f₃ : (Fin k → V) → V} + (hF : BoldfaceBoundedFunction₃ F) + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) : + BoldfaceBoundedFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v)) := + ⟨hF.bounded.comp hf₁.bounded hf₂.bounded hf₃.bounded, hF.definable.bcomp hf₁ hf₂ hf₃⟩ + +lemma BoldfaceBoundedFunction.comp₁ {k} {F : V → V} {f : (Fin k → V) → V} + [hFb : Bounded₁ F] [hFd : 𝚺₀.BoldfaceFunction₁ F] (hf : BoldfaceBoundedFunction f) : + BoldfaceBoundedFunction (fun v ↦ F (f v)) := BoldfaceBoundedFunction₁.comp ⟨hFb, hFd⟩ hf + +lemma BoldfaceBoundedFunction.comp₂ {k} {F : V → V → V} {f₁ f₂ : (Fin k → V) → V} + [hFb : Bounded₂ F] [hFd : 𝚺₀.BoldfaceFunction₂ F] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : + BoldfaceBoundedFunction (fun v ↦ F (f₁ v) (f₂ v)) := BoldfaceBoundedFunction₂.comp ⟨hFb, hFd⟩ hf₁ hf₂ + +lemma BoldfaceBoundedFunction.comp₃ {k} {F : V → V → V → V} {f₁ f₂ f₃ : (Fin k → V) → V} + [hFb : Bounded₃ F] [hFd : 𝚺₀.BoldfaceFunction₃ F] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) : + BoldfaceBoundedFunction (fun v ↦ F (f₁ v) (f₂ v) (f₃ v)) := BoldfaceBoundedFunction₃.comp ⟨hFb, hFd⟩ hf₁ hf₂ hf₃ + +section + +-- https://github.com/leanprover-community/mathlib4/blob/77d078e25cc501fae6907bfbcd80821920125266/Mathlib/Tactic/Measurability.lean#L25-L26 +open Lean.Parser.Tactic (config) + +open HierarchySymbol + +attribute [aesop (rule_sets := [Definability]) norm] + sq + Arith.pow_three + pow_four + +attribute [aesop 5 (rule_sets := [Definability]) safe] + BoldfaceFunction.comp₁ + BoldfaceFunction.comp₂ + BoldfaceFunction.comp₃ + BoldfaceBoundedFunction.comp₁ + BoldfaceBoundedFunction.comp₂ + BoldfaceBoundedFunction.comp₃ + +attribute [aesop 6 (rule_sets := [Definability]) safe] + Boldface.comp₁ + Boldface.comp₂ + Boldface.comp₃ + Boldface.comp₄ + Boldface.const + +attribute [aesop 7 (rule_sets := [Definability]) safe] + Boldface.bcomp₁ + Boldface.bcomp₂ + Boldface.bcomp₃ + Boldface.bcomp₄ + +attribute [aesop 8 (rule_sets := [Definability]) safe] + Boldface.ball_lt + Boldface.ball_le + Boldface.ball_lt' + Boldface.ball_le' + Boldface.bex_lt + Boldface.bex_le + +attribute [aesop 9 (rule_sets := [Definability]) safe] + Boldface.ball_lt_zero + Boldface.ball_le_zero + Boldface.bex_lt_zero + Boldface.bex_le_zero + +attribute [aesop 10 (rule_sets := [Definability]) safe] + Boldface.not + Boldface.imp + Boldface.iff + +attribute [aesop 11 (rule_sets := [Definability]) safe] + Boldface.and + Boldface.or + Boldface.all + Boldface.ex + +macro "definability" : attr => + `(attr|aesop 10 (rule_sets := [$(Lean.mkIdent `Definability):ident]) safe) + +macro "definability" (config)? : tactic => + `(tactic| aesop (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) + +macro "definability?" (config)? : tactic => + `(tactic| aesop? (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) + +example (c : V) : BoldfaceBoundedFunction₂ (fun x y : V ↦ c + 2 * x^2) := by definability + +example {ex : V → V} [𝚺₀.BoldfaceFunction₁ ex] (c : V) : + 𝚷₀.BoldfaceRel (fun x y : V ↦ ∃ z < x + c * y, (ex x = x ∧ x < y) ↔ ex x = z ∧ ex (x + 1) = 2 * z) := by + simp [Function.Graph.iff_left ex] + definability? + +example {ex : V → V} [h : 𝚫₁.BoldfaceFunction₁ ex] : + 𝚺₁.BoldfaceRel (fun x y : V ↦ ∃ z, x < y ↔ ex (ex x) = z) := by + definability? + +example {ex : V → V} [h : 𝚺₁.BoldfaceFunction₁ ex] : + 𝚺₁.BoldfaceRel (fun x y : V ↦ ∀ z < ex y, x < y ↔ ex (ex x) = z) := by + definability? + +end + +end LO.FirstOrder.Arith diff --git a/Arithmetization/Definability/Defined.lean b/Arithmetization/Definability/Defined.lean deleted file mode 100644 index 5f2805e..0000000 --- a/Arithmetization/Definability/Defined.lean +++ /dev/null @@ -1,864 +0,0 @@ -import Arithmetization.Definability.HSemiformula -import Arithmetization.Vorspiel.Graph - -namespace LO.FirstOrder.Arith - -end Arith - -def Defined {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semisentence L k) : Prop := - ∀ v, R v ↔ Semiformula.Evalbm M v p - -def DefinedWithParam {k} (R : (Fin k → M) → Prop) [Structure L M] (p : Semiformula L M k) : Prop := - ∀ v, R v ↔ Semiformula.Evalm M v id p - -lemma Defined.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semisentence L k} (h : Defined R p) (v) : - Semiformula.Evalbm M v p ↔ R v := (h v).symm - -lemma DefinedWithParam.iff [Structure L M] {k} {R : (Fin k → M) → Prop} {p : Semiformula L M k} (h : DefinedWithParam R p) (v) : - Semiformula.Evalm M v id p ↔ R v := (h v).symm - -namespace Arith.HierarchySymbol - -variable (ξ : Type*) (n : ℕ) - -open LO.Arith - -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] - -variable {Γ : HierarchySymbol} - -def Defined (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → Γ.Semisentence k → Prop - | 𝚺-[_], p => FirstOrder.Defined R p.val - | 𝚷-[_], p => FirstOrder.Defined R p.val - | 𝚫-[_], p => p.ProperOn M ∧ FirstOrder.Defined R p.val - -def DefinedWithParam (R : (Fin k → M) → Prop) : {Γ : HierarchySymbol} → Γ.Semiformula M k → Prop - | 𝚺-[_], p => FirstOrder.DefinedWithParam R p.val - | 𝚷-[_], p => FirstOrder.DefinedWithParam R p.val - | 𝚫-[_], p => p.ProperWithParamOn M ∧ FirstOrder.DefinedWithParam R p.val - -variable (L Γ) - -class Definable {k} (P : (Fin k → M) → Prop) : Prop where - definable : ∃ p : Γ.Semisentence k, Defined P p - -class DefinableWithParam {k} (P : (Fin k → M) → Prop) : Prop where - definable : ∃ p : Γ.Semiformula M k, DefinedWithParam P p - -abbrev DefinedPred (P : M → Prop) (p : Γ.Semisentence 1) : Prop := - Defined (λ v ↦ P (v 0)) p - -abbrev DefinedRel (R : M → M → Prop) (p : Γ.Semisentence 2) : Prop := - Defined (λ v ↦ R (v 0) (v 1)) p - -abbrev DefinedRel₃ (R : M → M → M → Prop) (p : Γ.Semisentence 3) : Prop := - Defined (λ v ↦ R (v 0) (v 1) (v 2)) p - -abbrev DefinedRel₄ (R : M → M → M → M → Prop) (p : Γ.Semisentence 4) : Prop := - Defined (λ v ↦ R (v 0) (v 1) (v 2) (v 3)) p - -variable {L Γ} - -abbrev DefinedFunction {k} (f : (Fin k → M) → M) (p : Γ.Semisentence (k + 1)) : Prop := - Defined (fun v => v 0 = f (v ·.succ)) p - -variable (L Γ) - -abbrev DefinedFunction₁ (f : M → M) (p : Γ.Semisentence 2) : Prop := - DefinedFunction (fun v => f (v 0)) p - -abbrev DefinedFunction₂ (f : M → M → M) (p : Γ.Semisentence 3) : Prop := - DefinedFunction (fun v => f (v 0) (v 1)) p - -abbrev DefinedFunction₃ (f : M → M → M → M) (p : Γ.Semisentence 4) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2)) p - -abbrev DefinedFunction₄ (f : M → M → M → M → M) (p : Γ.Semisentence 5) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3)) p - -abbrev DefinedFunction₅ (f : M → M → M → M → M → M) (p : Γ.Semisentence 6) : Prop := - DefinedFunction (fun v => f (v 0) (v 1) (v 2) (v 3) (v 4)) p - -abbrev DefinableWithParamPred (P : M → Prop) : Prop := Γ.DefinableWithParam (k := 1) (fun v ↦ P (v 0)) - -abbrev DefinableWithParamRel (P : M → M → Prop) : Prop := Γ.DefinableWithParam (k := 2) (fun v ↦ P (v 0) (v 1)) - -abbrev DefinableWithParamRel₃ (P : M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) - -abbrev DefinableWithParamRel₄ (P : M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) - -abbrev DefinableWithParamRel₅ (P : M → M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) - -abbrev DefinableWithParamRel₆ (P : M → M → M → M → M → M → Prop) : Prop := Γ.DefinableWithParam (k := 6) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) - -abbrev DefinableWithParamFunction (f : (Fin k → M) → M) : Prop := Γ.DefinableWithParam (k := k + 1) (fun v ↦ v 0 = f (v ·.succ)) - -abbrev DefinableWithParamFunction₁ (f : M → M) : Prop := Γ.DefinableWithParamFunction (k := 1) (fun v ↦ f (v 0)) - -abbrev DefinableWithParamFunction₂ (f : M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 2) (fun v ↦ f (v 0) (v 1)) - -abbrev DefinableWithParamFunction₃ (f : M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) - -abbrev DefinableWithParamFunction₄ (f : M → M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) - -abbrev DefinableWithParamFunction₅ (f : M → M → M → M → M → M) : Prop := Γ.DefinableWithParamFunction (k := 5) (fun v ↦ f (v 0) (v 1) (v 2) (v 3) (v 4)) - -variable {L Γ} - -section - -variable {k} {P Q : (Fin k → M) → Prop} - -namespace Defined - -lemma df {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semisentence k} (h : Defined R p) : FirstOrder.Defined R p.val := - match Γ with - | 𝚺-[_] => h - | 𝚷-[_] => h - | 𝚫-[_] => h.2 - -lemma proper {R : (Fin k → M) → Prop} {m} {p : 𝚫-[m].Semisentence k} (h : Defined R p) : p.ProperOn M := h.1 - -lemma of_zero {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : 𝚺₀.Semisentence k} (h : Defined R p) : Defined R (p.ofZero Γ) := - match Γ with - | 𝚺-[m] => by intro _; simp [h.iff] - | 𝚷-[m] => by intro _; simp [h.iff] - | 𝚫-[m] => ⟨by simp, by intro _; simp [h.iff]⟩ - -lemma emb {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semisentence k} (h : Defined R p) : Defined R p.emb := - match Γ with - | 𝚺-[m] => by intro _; simp [h.iff] - | 𝚷-[m] => by intro _; simp [h.iff] - | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ - -lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) - {p : Γ.Semisentence k} (H : Defined Q p) : Defined P p := by - rwa [show P = Q from by funext v; simp [h]] - -lemma to_definable {Γ : HierarchySymbol} (p : Γ.Semisentence k) (hP : Defined P p) : Γ.DefinableWithParam P := ⟨p.rew Rew.emb, by - match Γ with - | 𝚺-[_] => intro; simp [hP.iff] - | 𝚷-[_] => intro; simp [hP.iff] - | 𝚫-[_] => exact ⟨ - fun v ↦ by rcases p; simpa [HierarchySymbol.Semiformula.rew] using hP.proper.rew Rew.emb v, - by intro; simp [hP.df.iff]⟩⟩ - -lemma to_definable₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : - Γ.DefinableWithParam P := Defined.to_definable (p.ofZero Γ) hP.of_zero - -lemma to_definable_oRing (p : Γ.Semisentence k) (hP : Defined P p) : - Γ.DefinableWithParam P := Defined.to_definable p.emb hP.emb - -lemma to_definable_oRing₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : - Γ.DefinableWithParam P := Defined.to_definable₀ p.emb hP.emb - -end Defined - -namespace DefinedFunction - -lemma of_eq {f g : (Fin k → M) → M} (h : ∀ x, f x = g x) - {p : Γ.Semisentence (k + 1)} (H : DefinedFunction f p) : DefinedFunction g p := - Defined.of_iff (by intro; simp [h]) H - -lemma graph_delta {f : (Fin k → M) → M} {p : 𝚺-[m].Semisentence (k + 1)} - (h : DefinedFunction f p) : DefinedFunction f p.graphDelta := - ⟨by cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] - intro e; simp [Empty.eq_elim, h.df.iff] - rw [eq_comm], - by intro v; simp [h.df.iff]⟩ - -end DefinedFunction - -namespace DefinedWithParam - -lemma df {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semiformula M k} (h : DefinedWithParam R p) : FirstOrder.DefinedWithParam R p.val := - match Γ with - | 𝚺-[_] => h - | 𝚷-[_] => h - | 𝚫-[_] => h.2 - -lemma proper {R : (Fin k → M) → Prop} {m} {p : 𝚫-[m].Semiformula M k} (h : DefinedWithParam R p) : p.ProperWithParamOn M := h.1 - -lemma of_zero {R : (Fin k → M) → Prop} {Γ'} {p : Γ'-[0].Semiformula M k} - (h : DefinedWithParam R p) {Γ} : DefinedWithParam R (p.ofZero Γ) := - match Γ with - | 𝚺-[m] => by intro _; simp [h.df.iff] - | 𝚷-[m] => by intro _; simp [h.df.iff] - | 𝚫-[m] => ⟨by simp , by intro _; simp [h.df.iff]⟩ - -lemma of_deltaOne {R : (Fin k → M) → Prop} {Γ m} {p : 𝚫₁.Semiformula M k} - (h : DefinedWithParam R p) : DefinedWithParam R (p.ofDeltaOne Γ m) := - match Γ with - | 𝚺 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma] - | 𝚷 => by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, h.proper.iff'] - | 𝚫 => ⟨by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma, h.proper.iff'], - by intro _; simp [HierarchySymbol.Semiformula.ofDeltaOne, h.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - -lemma emb {R : (Fin k → M) → Prop} {Γ : HierarchySymbol} {p : Γ.Semiformula M k} - (h : DefinedWithParam R p) : DefinedWithParam R p.emb := - match Γ with - | 𝚺-[m] => by intro _; simp [h.iff] - | 𝚷-[m] => by intro _; simp [h.iff] - | 𝚫-[m] => ⟨by simpa using h.proper, by intro _; simp [h.df.iff]⟩ - -lemma of_iff {P Q : (Fin k → M) → Prop} (h : ∀ x, P x ↔ Q x) - {p : Γ.Semiformula M k} (H : DefinedWithParam Q p) : DefinedWithParam P p := by - rwa [show P = Q from by funext v; simp [h]] - -lemma to_definable {p : Γ.Semiformula M k} (h : DefinedWithParam P p) : Γ.DefinableWithParam P := ⟨p, h⟩ - -lemma to_definable₀ {p : Γ'-[0].Semiformula M k} - (h : DefinedWithParam P p) : Γ.DefinableWithParam P := ⟨p.ofZero Γ, h.of_zero⟩ - -lemma to_definable_deltaOne {p : 𝚫₁.Semiformula M k} {Γ m} - (h : DefinedWithParam P p) : Γ-[m + 1].DefinableWithParam P := ⟨p.ofDeltaOne Γ m, h.of_deltaOne⟩ - -variable {Γ : HierarchySymbol} - -lemma retraction {p : Γ.Semiformula M k} (hp : DefinedWithParam P p) (f : Fin k → Fin l) : - DefinedWithParam (fun v ↦ P fun i ↦ v (f i)) (p.rew <| Rew.substs fun x ↦ #(f x)) := - match Γ with - | 𝚺-[_] => by intro; simp [hp.df.iff] - | 𝚷-[_] => by intro; simp [hp.df.iff] - | 𝚫-[_] => ⟨hp.proper.rew _, by intro; simp [hp.df.iff]⟩ - -@[simp] lemma verum : - DefinedWithParam (fun _ ↦ True) (⊤ : Γ.Semiformula M k) := - match Γ with - | 𝚺-[m] => by intro v; simp - | 𝚷-[m] => by intro v; simp - | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ - -@[simp] lemma falsum : - DefinedWithParam (fun _ ↦ False) (⊥ : Γ.Semiformula M k) := - match Γ with -| 𝚺-[m] => by intro v; simp - | 𝚷-[m] => by intro v; simp - | 𝚫-[m] => ⟨by simp, by intro v; simp⟩ - -lemma and {p q : Γ.Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ∧ Q x) (p ⋏ q) := - match Γ with - | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] - | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] - | 𝚫-[m] => ⟨hp.proper.and hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ - -lemma or {p q : Γ.Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ∨ Q x) (p ⋎ q) := - match Γ with - | 𝚺-[m] => by intro v; simp [hp.iff, hq.iff] - | 𝚷-[m] => by intro v; simp [hp.iff, hq.iff] - | 𝚫-[m] => ⟨hp.proper.or hq.proper, by intro v; simp [hp.df.iff, hq.df.iff]⟩ - -lemma negSigma {p : 𝚺-[m].Semiformula M k} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) p.negSigma := by intro v; simp [hp.iff] - -lemma negPi {p : 𝚷-[m].Semiformula M k} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) p.negPi := by intro v; simp [hp.iff] - -lemma not {p : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) : - DefinedWithParam (fun x ↦ ¬P x) (~p) := ⟨hp.proper.neg, by intro v; simp [hp.proper.eval_neg, hp.df.iff]⟩ - -lemma imp {p q : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x → Q x) (p ⟶ q) := (hp.not.or hq).of_iff (by intro x; simp [imp_iff_not_or]) - -lemma iff {p q : 𝚫-[m].Semiformula M k} (hp : DefinedWithParam P p) (hq : DefinedWithParam Q q) : - DefinedWithParam (fun x ↦ P x ↔ Q x) (p ⟷ q) := ((hp.imp hq).and (hq.imp hp)).of_iff <| by intro v; simp [iff_iff_implies_and_implies] - -lemma ball {P : (Fin (k + 1) → M) → Prop} {p : Γ.Semiformula M (k + 1)} - (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ M k) : - DefinedWithParam (fun v ↦ ∀ x < t.valm M v id, P (x :> v)) (HierarchySymbol.Semiformula.ball t p) := - match Γ with - | 𝚺-[m] => by intro v; simp [hp.df.iff] - | 𝚷-[m] => by intro v; simp [hp.df.iff] - | 𝚫-[m] => ⟨hp.proper.ball, by intro v; simp [hp.df.iff]⟩ - -lemma bex {P : (Fin (k + 1) → M) → Prop} {p : Γ.Semiformula M (k + 1)} - (hp : DefinedWithParam P p) (t : Semiterm ℒₒᵣ M k) : - DefinedWithParam (fun v ↦ ∃ x < t.valm M v id, P (x :> v)) (HierarchySymbol.Semiformula.bex t p) := - match Γ with - | 𝚺-[m] => by intro v; simp [hp.df.iff] - | 𝚷-[m] => by intro v; simp [hp.df.iff] - | 𝚫-[m] => ⟨hp.proper.bex, by intro v; simp [hp.df.iff]⟩ - -lemma ex {P : (Fin (k + 1) → M) → Prop} {p : 𝚺-[m + 1].Semiformula M (k + 1)} - (hp : DefinedWithParam P p) : - DefinedWithParam (fun v ↦ ∃ x, P (x :> v)) p.ex := by intro _; simp [hp.df.iff] - -lemma all {P : (Fin (k + 1) → M) → Prop} {p : 𝚷-[m + 1].Semiformula M (k + 1)} - (hp : DefinedWithParam P p) : - DefinedWithParam (fun v ↦ ∀ x, P (x :> v)) p.all := by intro _; simp [hp.df.iff] - -end DefinedWithParam - -namespace DefinableWithParam - -lemma mkPolarity {P : (Fin k → M) → Prop} {Γ : Polarity} - (p : Semiformula ℒₒᵣ M k) (hp : Hierarchy Γ m p) (hP : ∀ v, P v ↔ Semiformula.Evalm M v id p) : Γ-[m].DefinableWithParam P := - match Γ with - | 𝚺 => ⟨.mkSigma p hp, by intro v; simp [hP]⟩ - | 𝚷 => ⟨.mkPi p hp, by intro v; simp [hP]⟩ - -lemma of_iff (Q : (Fin k → M) → Prop) (h : ∀ x, P x ↔ Q x) (H : Γ.DefinableWithParam Q) : Γ.DefinableWithParam P := by - rwa [show P = Q from by funext v; simp [h]] - -lemma of_oRing (h : Γ.DefinableWithParam P) : Γ.DefinableWithParam P := by - rcases h with ⟨p, hP⟩; exact ⟨p.emb, hP.emb⟩ - -lemma of_delta (h : 𝚫-[m].DefinableWithParam P) {Γ} : Γ-[m].DefinableWithParam P := by - rcases h with ⟨p, h⟩ - match Γ with - | 𝚺 => exact ⟨p.sigma, by intro v; simp [HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ - | 𝚷 => exact ⟨p.pi, by intro v; simp [←h.proper v, HierarchySymbol.Semiformula.val_sigma, h.df.iff]⟩ - | 𝚫 => exact ⟨p, h⟩ - -instance [𝚫-[m].DefinableWithParam P] (Γ) : Γ-[m].DefinableWithParam P := of_delta inferInstance - -lemma of_sigma_of_pi (hσ : 𝚺-[m].DefinableWithParam P) (hπ : 𝚷-[m].DefinableWithParam P) : 𝚫-[m].DefinableWithParam P := by - rcases hσ with ⟨p, hp⟩; rcases hπ with ⟨q, hq⟩ - exact ⟨.mkDelta p q, by intro v; simp [hp.df.iff, hq.df.iff], by intro v; simp [hp.df.iff]⟩ - -lemma of_zero (h : Γ'-[0].DefinableWithParam P) : Γ.DefinableWithParam P := by - rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable₀ - -lemma of_deltaOne (h : 𝚫₁.DefinableWithParam P) (Γ m) : Γ-[m + 1].DefinableWithParam P := by - rcases h with ⟨⟨p, hp⟩⟩; exact hp.to_definable_deltaOne - -instance [𝚺₀.DefinableWithParam P] (Γ : HierarchySymbol) : Γ.DefinableWithParam P := DefinableWithParam.of_zero (Γ' := 𝚺) (Γ := Γ) inferInstance - -variable {Γ : HierarchySymbol} - -lemma retraction (h : Γ.DefinableWithParam P) (f : Fin k → Fin n) : - Γ.DefinableWithParam fun v ↦ P (fun i ↦ v (f i)) := by - rcases h with ⟨p, h⟩ - exact ⟨p.rew (Rew.substs (fun i ↦ #(f i))), - match Γ with - | 𝚺-[_] => by intro; simp [h.df.iff] - | 𝚷-[_] => by intro; simp [h.df.iff] - | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ - -lemma retractiont (h : Γ.DefinableWithParam P) (f : Fin k → Semiterm ℒₒᵣ M n) : - Γ.DefinableWithParam fun v ↦ P (fun i ↦ Semiterm.valm M v id (f i)) := by - rcases h with ⟨p, h⟩ - exact ⟨p.rew (Rew.substs f), - match Γ with - | 𝚺-[_] => by intro; simp [h.df.iff] - | 𝚷-[_] => by intro; simp [h.df.iff] - | 𝚫-[_] => ⟨h.proper.rew _, by intro; simp [h.df.iff]⟩⟩ - -lemma const {P : Prop} : Γ.DefinableWithParam (fun _ : Fin k → M ↦ P) := of_zero (by - by_cases hP : P - · exact ⟨.mkSigma ⊤ (by simp), by intro; simp[hP]⟩ - · exact ⟨.mkSigma ⊥ (by simp), by intro; simp[hP]⟩) - -lemma and (h₁ : Γ.DefinableWithParam P) (h₂ : Γ.DefinableWithParam Q) : - Γ.DefinableWithParam (fun v ↦ P v ∧ Q v) := by - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁ ⋏ p₂, h₁.and h₂⟩ - -lemma or (h₁ : Γ.DefinableWithParam P) (h₂ : Γ.DefinableWithParam Q) : - Γ.DefinableWithParam (fun v ↦ P v ∨ Q v) := by - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁ ⋎ p₂, h₁.or h₂⟩ - -lemma not {Γ : SigmaPiDelta} (h : Γ.alt-[m].DefinableWithParam P) : - Γ-[m].DefinableWithParam (fun v ↦ ¬P v) := by - match Γ with - | 𝚺 => rcases h with ⟨p, h⟩; exact ⟨p.negPi, h.negPi⟩ - | 𝚷 => rcases h with ⟨p, h⟩; exact ⟨p.negSigma, h.negSigma⟩ - | 𝚫 => rcases h with ⟨p, h⟩; exact ⟨p.negDelta, h.not⟩ - -lemma imp {Γ : SigmaPiDelta} (h₁ : Γ.alt-[m].DefinableWithParam P) (h₂ : Γ-[m].DefinableWithParam Q) : - Γ-[m].DefinableWithParam (fun v ↦ P v → Q v) := by - match Γ with - | 𝚺 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁.negPi.or p₂, (h₁.negPi.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ - | 𝚷 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩ - exact ⟨p₁.negSigma.or p₂, (h₁.negSigma.or h₂).of_iff (fun x ↦ by simp [imp_iff_not_or])⟩ - | 𝚫 => - rcases h₁ with ⟨p₁, h₁⟩; rcases h₂ with ⟨p₂, h₂⟩; exact ⟨p₁ ⟶ p₂, h₁.imp h₂⟩ - -lemma iff (h₁ : 𝚫-[m].DefinableWithParam P) (h₂ : 𝚫-[m].DefinableWithParam Q) {Γ} : - Γ-[m].DefinableWithParam (fun v ↦ P v ↔ Q v) := - .of_delta (by rcases h₁ with ⟨p, hp⟩; rcases h₂ with ⟨q, hq⟩; exact ⟨p ⟷ q, hp.iff hq⟩) - -lemma all {P : (Fin k → M) → M → Prop} (h : 𝚷-[s + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - 𝚷-[s + 1].DefinableWithParam (fun v ↦ ∀ x, P v x) := by - rcases h with ⟨p, hp⟩ - exact ⟨.mkPi (∀' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ - -lemma ex {P : (Fin k → M) → M → Prop} (h : 𝚺-[s + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - 𝚺-[s + 1].DefinableWithParam (fun v ↦ ∃ x, P v x) := by - rcases h with ⟨p, hp⟩ - exact ⟨.mkSigma (∃' p.val) (by simp), by intro v; simp [hp.df.iff]⟩ - -lemma comp₁ {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : 𝚺-[m + 1].DefinableWithParamFunction f) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamPred P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f v)) := by - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact ⟨(pf ⋏ (p.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact ⟨(pf.negSigma ⋎ (p.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩ - exact of_sigma_of_pi - ⟨(pf ⋏ (p.sigma.rew (Rew.substs ![#0]))).ex, by intro v; simp [hp.df.iff, hpf.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ - ⟨(pf.negSigma ⋎ (p.pi.rew (Rew.substs ![#0]))).all, by intro v; simp [hp.df.iff, hpf.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₁_infer {k} {P : M → Prop} {f : (Fin k → M) → M} (hf : 𝚺-[m + 1].DefinableWithParamFunction f) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamPred P] : Γ-[m + 1].DefinableWithParam fun v ↦ P (f v) := - comp₁ hf inferInstance - -lemma comp₂ {k} {P : M → M → Prop} {f g : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel P) : Γ-[m + 1].DefinableWithParam fun v ↦ P (f v) (g v) := by - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.rew (Rew.substs ![#0, #1]))).ex.ex, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.rew (Rew.substs ![#0, #1]))).all.all, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩; rcases hf with ⟨pf, hpf⟩; rcases hg with ⟨pg, hpg⟩ - exact of_sigma_of_pi - ⟨(pf.rew (Rew.substs $ #0 :> (#·.succ.succ)) ⋏ pg.rew (Rew.substs $ #1 :> (#·.succ.succ)) ⋏ (p.sigma.rew (Rew.substs ![#0, #1]))).ex.ex, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - ⟨((pf.rew (Rew.substs $ #0 :> (#·.succ.succ))).negSigma - ⋎ (pg.rew (Rew.substs $ #1 :> (#·.succ.succ))).negSigma ⋎ (p.pi.rew (Rew.substs ![#0, #1]))).all.all, by - intro v; simp [hp.df.iff, hpf.df.iff, hpg.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₂_infer {k} {P : M → M → Prop} {f g : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f v) (g v)) := - comp₂ hf hg inferInstance - -lemma comp₃ {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₃ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2]))).ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2]))).all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₃_infer {k} {P : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₃ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v)) := - comp₃ hf₁ hf₂ hf₃ inferInstance - -lemma comp₄ {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₄ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩; rcases hf₄ with ⟨pf₄, hpf₄⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3]))).ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3]))).all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₄_infer {k} {P : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₄ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := - comp₄ hf₁ hf₂ hf₃ hf₄ inferInstance - -lemma comp₅ {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₅ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4]))).ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4]))).all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₅_infer {k} {P : M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₅ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v)) := - comp₅ hf₁ hf₂ hf₃ hf₄ hf₅ inferInstance - -lemma comp₆ {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) (hf₆ : 𝚺-[m + 1].DefinableWithParamFunction f₆) - {Γ : SigmaPiDelta} (hP : Γ-[m + 1].DefinableWithParamRel₆ P) : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := by - rcases hf₁ with ⟨pf₁, hpf₁⟩; rcases hf₂ with ⟨pf₂, hpf₂⟩; rcases hf₃ with ⟨pf₃, hpf₃⟩ - rcases hf₄ with ⟨pf₄, hpf₄⟩; rcases hf₅ with ⟨pf₅, hpf₅⟩; rcases hf₆ with ⟨pf₆, hpf₆⟩ - match Γ with - | 𝚺 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff]⟩ - | 𝚷 => - rcases hP with ⟨p, hp⟩ - exact - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or]⟩ - | 𝚫 => - rcases hP with ⟨p, hp⟩ - exact of_sigma_of_pi - ⟨( pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ)) - ⋏ (p.sigma.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).ex.ex.ex.ex.ex.ex, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, HierarchySymbol.Semiformula.val_sigma]⟩ - ⟨( (pf₁.rew (Rew.substs $ #0 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₂.rew (Rew.substs $ #1 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₃.rew (Rew.substs $ #2 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₄.rew (Rew.substs $ #3 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₅.rew (Rew.substs $ #4 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (pf₆.rew (Rew.substs $ #5 :> (#·.succ.succ.succ.succ.succ.succ))).negSigma - ⋎ (p.pi.rew (Rew.substs ![#0, #1, #2, #3, #4, #5]))).all.all.all.all.all.all, by - intro v; simp [hp.df.iff, hpf₁.df.iff, hpf₂.df.iff, hpf₃.df.iff, hpf₄.df.iff, hpf₅.df.iff, hpf₆.df.iff, ←imp_iff_not_or, hp.proper.iff']⟩ - -lemma comp₆_infer {k} {P : M → M → M → M → M → M → Prop} {f₁ f₂ f₃ f₄ f₅ f₆ : (Fin k → M) → M} - (hf₁ : 𝚺-[m + 1].DefinableWithParamFunction f₁) (hf₂ : 𝚺-[m + 1].DefinableWithParamFunction f₂) - (hf₃ : 𝚺-[m + 1].DefinableWithParamFunction f₃) (hf₄ : 𝚺-[m + 1].DefinableWithParamFunction f₄) - (hf₅ : 𝚺-[m + 1].DefinableWithParamFunction f₅) (hf₆ : 𝚺-[m + 1].DefinableWithParamFunction f₆) - {Γ : SigmaPiDelta} [Γ-[m + 1].DefinableWithParamRel₆ P] : Γ-[m + 1].DefinableWithParam (fun v ↦ P (f₁ v) (f₂ v) (f₃ v) (f₄ v) (f₅ v) (f₆ v)) := - comp₆ hf₁ hf₂ hf₃ hf₄ hf₅ hf₆ inferInstance - -end DefinableWithParam - -lemma DefinableWithParamPred.of_iff {P : M → Prop} (Q : M → Prop) (h : ∀ x, P x ↔ Q x) (H : Γ.DefinableWithParamPred Q) : Γ.DefinableWithParamPred P := by - rwa [show P = Q from by funext v; simp [h]] - -instance DefinableWithParamFunction₁.graph {f : M → M} [h : Γ.DefinableWithParamFunction₁ f] : - Γ.DefinableWithParamRel (Function.Graph f) := h - -instance DefinableWithParamFunction₂.graph {f : M → M → M} [h : Γ.DefinableWithParamFunction₂ f] : - Γ.DefinableWithParamRel₃ (Function.Graph₂ f) := h - -instance DefinableWithParamFunction₃.graph {f : M → M → M → M} [h : Γ.DefinableWithParamFunction₃ f] : - Γ.DefinableWithParamRel₄ (Function.Graph₃ f) := h - -namespace DefinableWithParamFunction - -lemma graph_delta {k} {f : (Fin k → M) → M} - (h : 𝚺-[m].DefinableWithParamFunction f) : 𝚫-[m].DefinableWithParamFunction f := by - rcases h with ⟨p, h⟩ - exact ⟨p.graphDelta, by - cases' m with m <;> simp [HierarchySymbol.Semiformula.graphDelta] - intro e; simp [Empty.eq_elim, h.df.iff] - exact eq_comm, by - intro v; simp [h.df.iff]⟩ - -instance {k} {f : (Fin k → M) → M} [h : 𝚺-[m].DefinableWithParamFunction f] : 𝚫-[m].DefinableWithParamFunction f := - DefinableWithParamFunction.graph_delta h - -instance {k} {f : (Fin k → M) → M} [𝚺₀.DefinableWithParamFunction f] : Γ.DefinableWithParamFunction f := inferInstance - -lemma of_sigmaOne {k} {f : (Fin k → M) → M} - (h : 𝚺₁.DefinableWithParamFunction f) (Γ m) : Γ-[m + 1].DefinableWithParamFunction f := DefinableWithParam.of_deltaOne (graph_delta h) Γ m - -@[simp] lemma var {k} (i : Fin k) : Γ.DefinableWithParamFunction (fun v : Fin k → M ↦ v i) := - .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!#i.succ” (by simp), by intro _; simp⟩ - -@[simp] lemma const {k} (c : M) : Γ.DefinableWithParamFunction (fun _ : Fin k → M ↦ c) := - .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | #0 = &c” (by simp), by intro v; simp⟩ - -@[simp] lemma term_retraction (t : Semiterm ℒₒᵣ M n) (e : Fin n → Fin k) : - Γ.DefinableWithParamFunction fun v : Fin k → M ↦ Semiterm.valm M (fun x ↦ v (e x)) id t := - .of_zero (Γ' := 𝚺) - ⟨.mkSigma “x | x = !!(Rew.substs (fun x ↦ #(e x).succ) t)” (by simp), by intro v; simp [Semiterm.val_substs]⟩ - -@[simp] lemma term (t : Semiterm ℒₒᵣ M k) : - Γ.DefinableWithParamFunction fun v : Fin k → M ↦ Semiterm.valm M v id t := - .of_zero (Γ' := 𝚺) ⟨.mkSigma “x | x = !!(Rew.bShift t)” (by simp), by intro v; simp [Semiterm.val_bShift']⟩ - -lemma of_eq {f : (Fin k → M) → M} (g) (h : ∀ v, f v = g v) (H : Γ.DefinableWithParamFunction f) : Γ.DefinableWithParamFunction g := by - rwa [show g = f from by funext v; simp [h]] - -lemma retraction {n k} {f : (Fin k → M) → M} (hf : Γ.DefinableWithParamFunction f) (e : Fin k → Fin n) : - Γ.DefinableWithParamFunction fun v ↦ f (fun i ↦ v (e i)) := by - have := DefinableWithParam.retraction (n := n + 1) hf (0 :> fun i ↦ (e i).succ); simp at this - exact this.of_iff _ (by intro x; simp) - -lemma retractiont {f : (Fin k → M) → M} (hf : Γ.DefinableWithParamFunction f) (t : Fin k → Semiterm ℒₒᵣ M n) : - Γ.DefinableWithParamFunction fun v ↦ f (fun i ↦ Semiterm.valm M v id (t i)) := by - have := DefinableWithParam.retractiont (n := n + 1) hf (#0 :> fun i ↦ Rew.bShift (t i)); simp at this - exact this.of_iff _ (by intro x; simp [Semiterm.val_bShift']) - -lemma rel {f : (Fin k → M) → M} (h : Γ.DefinableWithParamFunction f) : - Γ.DefinableWithParam (fun v ↦ v 0 = f (v ·.succ)) := h - -end DefinableWithParamFunction - -lemma DefinableWithParamFunction₁.comp {Γ} {k} {f : M → M} {g : (Fin k → M) → M} - (hf : Γ-[m + 1].DefinableWithParamFunction₁ f) (hg : 𝚺-[m + 1].DefinableWithParamFunction g) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g v)) := by - simpa using DefinableWithParam.comp₂ (P := Function.Graph f) (DefinableWithParamFunction.var 0) (hg.retraction Fin.succ) hf - -lemma DefinableWithParamFunction₂.comp {Γ} {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} - (hf : Γ-[m + 1].DefinableWithParamFunction₂ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v)) := by - simpa using DefinableWithParam.comp₃ (P := Function.Graph₂ f) (DefinableWithParamFunction.var 0) (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) hf - -lemma DefinableWithParamFunction₃.comp {Γ} {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} - (hf : Γ-[m + 1].DefinableWithParamFunction₃ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) - (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := by - simpa using DefinableWithParam.comp₄ (P := Function.Graph₃ f) (DefinableWithParamFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) hf - -lemma DefinableWithParamFunction₄.comp {Γ} {k} {f : M → M → M → M → M} {g₁ g₂ g₃ g₄ : (Fin k → M) → M} - (hf : Γ-[m + 1].DefinableWithParamFunction₄ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) - (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) - (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := by - simpa using DefinableWithParam.comp₅ (P := Function.Graph₄ f) (DefinableWithParamFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) (hg₄.retraction Fin.succ) hf - -lemma DefinableWithParamFunction₅.comp {Γ} {k} {f : M → M → M → M → M → M} {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} - (hf : Γ-[m + 1].DefinableWithParamFunction₅ f) (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) - (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) - (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) (hg₅ : 𝚺-[m + 1].DefinableWithParamFunction g₅) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := by - simpa using DefinableWithParam.comp₆ (P := Function.Graph₅ f) (DefinableWithParamFunction.var 0) - (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) - (hg₄.retraction Fin.succ) (hg₅.retraction Fin.succ) hf - -lemma DefinableWithParamFunction.comp₁_infer {Γ} {k} {f : M → M} [Γ-[m + 1].DefinableWithParamFunction₁ f] - {g : (Fin k → M) → M} (hg : 𝚺-[m + 1].DefinableWithParamFunction g) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g v)) := - DefinableWithParamFunction₁.comp inferInstance hg - -lemma DefinableWithParamFunction.comp₂_infer {Γ} {k} {f : M → M → M} [Γ-[m + 1].DefinableWithParamFunction₂ f] - {g₁ g₂ : (Fin k → M) → M} (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v)) := - DefinableWithParamFunction₂.comp inferInstance hg₁ hg₂ - -lemma DefinableWithParamFunction.comp₃_infer {Γ} {k} {f : M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₃ f] - {g₁ g₂ g₃ : (Fin k → M) → M} - (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := - DefinableWithParamFunction₃.comp inferInstance hg₁ hg₂ hg₃ - -lemma DefinableWithParamFunction.comp₄_infer {Γ} {k} {f : M → M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₄ f] - {g₁ g₂ g₃ g₄ : (Fin k → M) → M} - (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) - (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v)) := - DefinableWithParamFunction₄.comp inferInstance hg₁ hg₂ hg₃ hg₄ - -lemma DefinableWithParamFunction.comp₅_infer {Γ} {k} {f : M → M → M → M → M → M} [Γ-[m + 1].DefinableWithParamFunction₅ f] - {g₁ g₂ g₃ g₄ g₅ : (Fin k → M) → M} - (hg₁ : 𝚺-[m + 1].DefinableWithParamFunction g₁) (hg₂ : 𝚺-[m + 1].DefinableWithParamFunction g₂) - (hg₃ : 𝚺-[m + 1].DefinableWithParamFunction g₃) (hg₄ : 𝚺-[m + 1].DefinableWithParamFunction g₄) - (hg₅ : 𝚺-[m + 1].DefinableWithParamFunction g₅) : - Γ-[m + 1].DefinableWithParamFunction (fun v ↦ f (g₁ v) (g₂ v) (g₃ v) (g₄ v) (g₅ v)) := - DefinableWithParamFunction₅.comp inferInstance hg₁ hg₂ hg₃ hg₄ hg₅ - -namespace DefinableWithParamRel - -@[simp] instance eq : Γ.DefinableWithParamRel (Eq : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1” (by simp)) (by intro _; simp) - -@[simp] instance lt : Γ.DefinableWithParamRel (LT.lt : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 < #1” (by simp)) (by intro _; simp) - -@[simp] instance le : Γ.DefinableWithParamRel (LE.le : M → M → Prop) := - Defined.to_definable_oRing₀ (.mkSigma “#0 ≤ #1” (by simp)) (by intro _; simp) - -end DefinableWithParamRel - -namespace DefinableWithParamFunction₂ - -@[simp] instance add : Γ.DefinableWithParamFunction₂ ((· + ·) : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) - -@[simp] instance mul : Γ.DefinableWithParamFunction₂ ((· * ·) : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) - -@[simp] instance hAdd : Γ.DefinableWithParamFunction₂ (HAdd.hAdd : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 + #2” (by simp)) (by intro _; simp) - -@[simp] instance hMul : Γ.DefinableWithParamFunction₂ (HMul.hMul : M → M → M) := - Defined.to_definable_oRing₀ (.mkSigma “#0 = #1 * #2” (by simp)) (by intro _; simp) - -end DefinableWithParamFunction₂ - -namespace DefinableWithParam - -lemma ball_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x < f v, P v x) := by - rcases hf with ⟨bf, hbf⟩ - rcases h with ⟨p, hp⟩ - match Γ with - | 𝚺 => exact - ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚷 => exact - ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚫 => - exact .of_sigma_of_pi - ⟨ .mkSigma (∃' (bf.val ⋏ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ - ⟨ .mkPi (∀' (bf.val ⟶ (∀[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ - -lemma bex_lt {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x < f v, P v x) := by - rcases hf with ⟨bf, hbf⟩ - rcases h with ⟨p, hp⟩ - match Γ with - | 𝚺 => exact - ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚷 => exact - ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff] ⟩ - | 𝚫 => - exact .of_sigma_of_pi - ⟨ .mkSigma (∃' (bf.val ⋏ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.sigma.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, HierarchySymbol.Semiformula.val_sigma] ⟩ - ⟨ .mkPi (∀' (bf.val ⟶ (∃[“#0 < #1”] Rew.substs (#0 :> (#·.succ.succ)) |>.hom p.pi.val))) (by simp), - by intro v; simp [hbf.df.iff, hp.df.iff, hp.proper.iff'] ⟩ - -lemma ball_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x ≤ f v, P v x) := by - have : Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ x < f v + 1, P v x) := ball_lt (DefinableWithParamFunction₂.comp (by simp) hf (by simp)) h - exact this.of_iff <| by intro v; simp [lt_succ_iff_le] - -lemma bex_le {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x ≤ f v, P v x) := by - have : Γ-[m + 1].DefinableWithParam (fun v ↦ ∃ x < f v + 1, P v x) := bex_lt (DefinableWithParamFunction₂.comp (by simp) hf (by simp)) h - exact this.of_iff <| by intro v; simp [lt_succ_iff_le] - -lemma ball_lt' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ {x}, x < f v → P v x) := ball_lt hf h - -lemma ball_le' {Γ} {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : 𝚺-[m + 1].DefinableWithParamFunction f) (h : Γ-[m + 1].DefinableWithParam (fun w ↦ P (w ·.succ) (w 0))) : - Γ-[m + 1].DefinableWithParam (fun v ↦ ∀ {x}, x ≤ f v → P v x) := ball_le hf h - -end DefinableWithParam - -end - -end Arith.HierarchySymbol - -end LO.FirstOrder diff --git a/Arithmetization/Definability/HSemiformula.lean b/Arithmetization/Definability/Hierarchy.lean similarity index 100% rename from Arithmetization/Definability/HSemiformula.lean rename to Arithmetization/Definability/Hierarchy.lean diff --git a/Arithmetization/Definability/PolyBounded.lean b/Arithmetization/Definability/PolyBounded.lean deleted file mode 100644 index 23c3e63..0000000 --- a/Arithmetization/Definability/PolyBounded.lean +++ /dev/null @@ -1,421 +0,0 @@ -import Arithmetization.Definability.Hierarchy - -namespace LO.FirstOrder.Arith - -open LO.Arith - -variable {L : Language} [L.ORing] {ξ : Type*} {n : ℕ} - -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] [Structure L M] [Structure.ORing L M] [Structure.Monotone L M] - -variable {Γ : HierarchySymbol} - -variable (L Γ) - -class Bounded (f : (Fin k → M) → M) : Prop where - bounded : ∃ t : Semiterm L M k, ∀ v : Fin k → M, f v ≤ t.valm M v id - -abbrev Bounded₁ (f : M → M) : Prop := Bounded L (k := 1) (fun v ↦ f (v 0)) - -abbrev Bounded₂ (f : M → M → M) : Prop := Bounded L (k := 2) (fun v ↦ f (v 0) (v 1)) - -abbrev Bounded₃ (f : M → M → M → M) : Prop := Bounded L (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) - -instance (f : (Fin k → M) → M) [h : Bounded ℒₒᵣ f] : Bounded L f := by - rcases h with ⟨t, ht⟩ - exact ⟨Semiterm.lMap Language.oringEmb t, by simpa⟩ - -variable {L Γ} - -namespace Bounded - -@[simp] lemma var {k} (i : Fin k) : Bounded L fun v : Fin k → M ↦ v i := ⟨#i, by intro _; simp⟩ - -@[simp] lemma const {k} (c : M) : Bounded L (fun _ : Fin k → M ↦ c) := ⟨&c, by intro _; simp⟩ - -@[simp] lemma term_retraction (t : Semiterm L M n) (e : Fin n → Fin k) : - Bounded L fun v : Fin k → M ↦ Semiterm.valm M (fun x ↦ v (e x)) id t := - ⟨Rew.substs (fun x ↦ #(e x)) t, by intro _; simp [Semiterm.val_substs]⟩ - -@[simp] lemma term (t : Semiterm L M k) : Bounded L fun v : Fin k → M => Semiterm.valm M v id t := - ⟨t, by intro _; simp⟩ - -lemma retraction {f : (Fin k → M) → M} (hf : Bounded L f) (e : Fin k → Fin n) : - Bounded L fun v ↦ f (fun i ↦ v (e i)) := by - rcases hf with ⟨t, ht⟩ - exact ⟨Rew.substs (fun x ↦ #(e x)) t, by intro; simp [Semiterm.val_substs, ht]⟩ - -lemma comp {k} {f : (Fin l → M) → M} {g : Fin l → (Fin k → M) → M} (hf : Bounded L f) (hg : ∀ i, Bounded L (g i)) : - Bounded L (fun v ↦ f (g · v)) where - bounded := by - rcases hf.bounded with ⟨tf, htf⟩ - choose tg htg using fun i ↦ (hg i).bounded - exact ⟨Rew.substs tg tf, by - intro v; simp [Semiterm.val_substs] - exact le_trans (htf (g · v)) (Structure.Monotone.term_monotone tf (fun i ↦ htg i v) (by simp))⟩ - -end Bounded - -lemma Bounded₁.comp {f : M → M} {k} {g : (Fin k → M) → M} (hf : Bounded₁ L f) (hg : Bounded L g) : - Bounded L (fun v ↦ f (g v)) := Bounded.comp hf (l := 1) (fun _ ↦ hg) - -lemma Bounded₂.comp {f : M → M → M} {k} {g₁ g₂ : (Fin k → M) → M} - (hf : Bounded₂ L f) (hg₁ : Bounded L g₁) (hg₂ : Bounded L g₂) : - Bounded L (fun v ↦ f (g₁ v) (g₂ v)) := Bounded.comp hf (g := ![g₁, g₂]) (fun i ↦ by cases i using Fin.cases <;> simp [*]) - -lemma Bounded₃.comp {f : M → M → M → M} {k} {g₁ g₂ g₃ : (Fin k → M) → M} - (hf : Bounded₃ L f) (hg₁ : Bounded L g₁) (hg₂ : Bounded L g₂) (hg₃ : Bounded L g₃) : - Bounded L (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := Bounded.comp hf (g := ![g₁, g₂, g₃]) - (fun i ↦ by - cases' i using Fin.cases with i <;> simp [*] - cases' i using Fin.cases with i <;> simp [*]) - -namespace Bounded₂ - -instance add : Bounded₂ L ((· + ·) : M → M → M) where - bounded := ⟨‘x y | x + y’, by intro _; simp⟩ - -instance mul : Bounded₂ L ((· * ·) : M → M → M) where - bounded := ⟨‘x y | x * y’, by intro _; simp⟩ - -instance hAdd : Bounded₂ L (HAdd.hAdd : M → M → M) where - bounded := ⟨‘x y | x + y’, by intro _; simp⟩ - -instance hMul : Bounded₂ L (HMul.hMul : M → M → M) where - bounded := ⟨‘x y | x * y’, by intro _; simp⟩ - -end Bounded₂ - -variable (L Γ) - -def DefinableBoundedFunction {k} (f : (Fin k → M) → M) := Bounded L f ∧ DefinableFunction L Γ f - -abbrev DefinableBoundedFunction₁ (f : M → M) : Prop := DefinableBoundedFunction L Γ (k := 1) (fun v => f (v 0)) - -abbrev DefinableBoundedFunction₂ (f : M → M → M) : Prop := DefinableBoundedFunction L Γ (k := 2) (fun v => f (v 0) (v 1)) - -abbrev DefinableBoundedFunction₃ (f : M → M → M → M) : Prop := DefinableBoundedFunction L Γ (k := 3) (fun v => f (v 0) (v 1) (v 2)) - -variable {L Γ} - -lemma DefinableBoundedFunction.bounded {f : (Fin k → M) → M} (h : DefinableBoundedFunction L Γ f) : Bounded L f := h.1 - -lemma DefinableBoundedFunction₁.bounded {f : M → M} (h : DefinableBoundedFunction₁ L Γ f) : Bounded₁ L f := h.1 - -lemma DefinableBoundedFunction₂.bounded {f : M → M → M} (h : DefinableBoundedFunction₂ L Γ f) : Bounded₂ L f := h.1 - -lemma DefinableBoundedFunction₃.bounded {f : M → M → M → M} (h : DefinableBoundedFunction₃ L Γ f) : Bounded₃ L f := h.1 - -lemma DefinableBoundedFunction.definable {f : (Fin k → M) → M} (h : DefinableBoundedFunction L Γ f) : DefinableFunction L Γ f := h.2 - -lemma DefinableBoundedFunction₁.definable {f : M → M} (h : DefinableBoundedFunction₁ L Γ f) : DefinableFunction₁ L Γ f := h.2 - -lemma DefinableBoundedFunction₂.definable {f : M → M → M} (h : DefinableBoundedFunction₂ L Γ f) : DefinableFunction₂ L Γ f := h.2 - -lemma DefinableBoundedFunction₃.definable {f : M → M → M → M} (h : DefinableBoundedFunction₃ L Γ f) : DefinableFunction₃ L Γ f := h.2 - -namespace DefinableBoundedFunction - -lemma of_polybounded_of_definable (f : (Fin k → M) → M) [hb : Bounded L f] [hf : DefinableFunction L Γ f] : - DefinableBoundedFunction L Γ f := ⟨hb, hf⟩ - -@[simp] lemma of_polybounded_of_definable₁ (f : M → M) [hb : Bounded₁ L f] [hf : DefinableFunction₁ L Γ f] : - DefinableBoundedFunction₁ L Γ f := ⟨hb, hf⟩ - -@[simp] lemma of_polybounded_of_definable₂ (f : M → M → M) [hb : Bounded₂ L f] [hf : DefinableFunction₂ L Γ f] : - DefinableBoundedFunction₂ L Γ f := ⟨hb, hf⟩ - -@[simp] lemma of_polybounded_of_definable₃ (f : M → M → M → M) [hb : Bounded₃ L f] [hf : DefinableFunction₃ L Γ f] : - DefinableBoundedFunction₃ L Γ f := ⟨hb, hf⟩ - -lemma retraction {f : (Fin k → M) → M} (hf : DefinableBoundedFunction L Γ f) (e : Fin k → Fin n) : - DefinableBoundedFunction L Γ fun v ↦ f (fun i ↦ v (e i)) := ⟨hf.bounded.retraction e, hf.definable.retraction e⟩ - -lemma of_zero {Γ' Γ} {f : (Fin k → M) → M} (h : DefinableBoundedFunction L (Γ', 0) f) : - DefinableBoundedFunction L (Γ, 0) f := by - rcases h with ⟨hb, h⟩ - exact ⟨hb, .of_zero h _⟩ - -lemma of_delta {f : (Fin k → M) → M} (h : DefinableBoundedFunction L (𝚫, m) f) {Γ} : DefinableBoundedFunction L (Γ, m) f := - ⟨h.bounded, h.definable.of_delta⟩ - -end DefinableBoundedFunction - -namespace Definable - -variable {P Q : (Fin k → M) → Prop} - -lemma ball_lt₀ {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableBoundedFunction L Γ f) (h : Definable L Γ (fun w ↦ P (w ·.succ) (w 0))) : - Definable L Γ (fun v ↦ ∀ x < f v, P v x) := by - rcases hf.bounded with ⟨bf, hbf⟩ - rcases hf.definable with ⟨f_graph, hf_graph⟩ - rcases h with ⟨p, hp⟩ - have : DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm M v id bf, x = f v ∧ ∀ y < x, P v y) - (HSemiformula.bex ‘!!bf + 1’ - (f_graph ⋏ HSemiformula.ball (#0) (HSemiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by - simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).ball #0)).bex ‘!!bf + 1’ - exact .of_iff _ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) ⟨_, this⟩ - -lemma bex_lt₀ {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableBoundedFunction L Γ f) (h : Definable L Γ (fun w ↦ P (w ·.succ) (w 0))) : - Definable L Γ (fun v ↦ ∃ x < f v, P v x) := by - rcases hf.bounded with ⟨bf, hbf⟩ - rcases hf.definable with ⟨f_graph, hf_graph⟩ - rcases h with ⟨p, hp⟩ - have : DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm M v id bf, x = f v ∧ ∃ y < x, P v y) - (HSemiformula.bex ‘!!bf + 1’ - (f_graph ⋏ HSemiformula.bex (#0) (HSemiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by - simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex #0)).bex ‘!!bf + 1’ - exact .of_iff _ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) ⟨_, this⟩ - -lemma ball_le₀ {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableBoundedFunction L Γ f) (h : Definable L Γ (fun w ↦ P (w ·.succ) (w 0))) : - Definable L Γ (fun v ↦ ∀ x ≤ f v, P v x) := by - rcases hf.bounded with ⟨bf, hbf⟩ - rcases hf.definable with ⟨f_graph, hf_graph⟩ - rcases h with ⟨p, hp⟩ - have : DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm M v id bf, x = f v ∧ ∀ y ≤ x, P v y) - (HSemiformula.bex ‘!!bf + 1’ - (f_graph ⋏ HSemiformula.ball ‘x | x + 1’ (HSemiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by - simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).ball ‘x | x + 1’)).bex ‘!!bf + 1’ - exact .of_iff _ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) ⟨_, this⟩ - -lemma bex_le₀ {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableBoundedFunction L Γ f) (h : Definable L Γ (fun w ↦ P (w ·.succ) (w 0))) : - Definable L Γ (fun v ↦ ∃ x ≤ f v, P v x) := by - rcases hf.bounded with ⟨bf, hbf⟩ - rcases hf.definable with ⟨f_graph, hf_graph⟩ - rcases h with ⟨p, hp⟩ - have : DefinedWithParam (fun v ↦ ∃ x ≤ Semiterm.valm M v id bf, x = f v ∧ ∃ y ≤ x, P v y) - (HSemiformula.bex ‘!!bf + 1’ - (f_graph ⋏ HSemiformula.bex ‘x | x + 1’ (HSemiformula.rew (Rew.substs (#0 :> fun i => #i.succ.succ)) p))) := by - simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex ‘x | x + 1’)).bex ‘!!bf + 1’ - exact .of_iff _ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) ⟨_, this⟩ - -end Definable - -namespace DefinableBoundedFunction - -lemma of_iff {g : (Fin k → M) → M} (f) (h : ∀ v, f v = g v) (H : DefinableBoundedFunction L Γ f) : DefinableBoundedFunction L Γ g := by - have : f = g := by funext v; simp [h] - rcases this; exact H - -@[simp] lemma var {k} (i : Fin k) : DefinableBoundedFunction L Γ (fun v : Fin k → M ↦ v i) := ⟨by simp, by simp⟩ - -@[simp] lemma const {k} (c : M) : DefinableBoundedFunction L Γ (fun _ : Fin k → M ↦ c) := ⟨by simp, by simp⟩ - -@[simp] lemma term_retraction (t : Semiterm L M n) (e : Fin n → Fin k) : - DefinableBoundedFunction L Γ fun v : Fin k → M ↦ Semiterm.valm M (fun x ↦ v (e x)) id t := ⟨by simp, by simp⟩ - -@[simp] lemma term (t : Semiterm L M k) : - DefinableBoundedFunction L Γ fun v : Fin k → M ↦ Semiterm.valm M v id t := ⟨by simp, by simp⟩ - -end DefinableBoundedFunction - -namespace Definable - -lemma bcomp₁ {k} {P : M → Prop} {f : (Fin k → M) → M} [hP : DefinablePred L Γ P] (hf : DefinableBoundedFunction L Γ f) : - Definable L Γ (fun v ↦ P (f v)) := by - rcases hf.bounded with ⟨bf, hbf⟩ - have : Definable L Γ fun v ↦ ∃ z ≤ Semiterm.valm M v id bf, z = f v ∧ P z := - bex_le₀ (by simp) (and hf.definable <| hP.retraction (fun _ ↦ 0)) - exact this.of_iff _ (by - intro v; constructor - · intro h; exact ⟨f v, hbf v, rfl, h⟩ - · rintro ⟨_, _, rfl, h⟩; exact h) - -lemma bcomp₂ {k} {R : M → M → Prop} {f₁ f₂ : (Fin k → M) → M} - [hR : DefinableRel L Γ R] (hf₁ : DefinableBoundedFunction L Γ f₁) (hf₂ : DefinableBoundedFunction L Γ f₂) : - Definable L Γ (fun v ↦ R (f₁ v) (f₂ v)) := by - rcases hf₁.bounded with ⟨bf₁, hbf₁⟩ - rcases hf₂.bounded with ⟨bf₂, hbf₂⟩ - have : Definable L Γ (fun v ↦ - ∃ z₁ ≤ Semiterm.valm M v id bf₁, ∃ z₂ ≤ Semiterm.valm M v id bf₂, z₁ = f₁ v ∧ z₂ = f₂ v ∧ R z₁ z₂) := - bex_le₀ (DefinableBoundedFunction.term _) <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) - <| and (hf₁.definable.rel.retraction _) - <| and (by simpa using hf₂.definable.rel.retraction (0 :> (·.succ.succ))) - <| by simpa using hR.retraction (n := k + 2) ![1, 0] - exact this.of_iff _ (by - intro v; constructor - · intro h; exact ⟨f₁ v, hbf₁ v, f₂ v, hbf₂ v, rfl, rfl, h⟩ - · rintro ⟨_, _, _, _, rfl, rfl, h⟩; exact h) - -lemma bcomp₃ {k} {R : M → M → M → Prop} {f₁ f₂ f₃ : (Fin k → M) → M} - [hR : DefinableRel₃ L Γ R] (hf₁ : DefinableBoundedFunction L Γ f₁) (hf₂ : DefinableBoundedFunction L Γ f₂) (hf₃ : DefinableBoundedFunction L Γ f₃) : - Definable L Γ (fun v ↦ R (f₁ v) (f₂ v) (f₃ v)) := by - rcases hf₁.bounded with ⟨bf₁, hbf₁⟩ - rcases hf₂.bounded with ⟨bf₂, hbf₂⟩ - rcases hf₃.bounded with ⟨bf₃, hbf₃⟩ - have : Definable L Γ (fun v ↦ - ∃ z₁ ≤ Semiterm.valm M v id bf₁, ∃ z₂ ≤ Semiterm.valm M v id bf₂, ∃ z₃ ≤ Semiterm.valm M v id bf₃, - z₁ = f₁ v ∧ z₂ = f₂ v ∧ z₃ = f₃ v ∧ R z₁ z₂ z₃) := - bex_le₀ (DefinableBoundedFunction.term _) <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) - <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) - <| and (by simpa using hf₁.definable.rel.retraction (n := k + 3) (2 :> (·.succ.succ.succ))) - <| and (by simpa using hf₂.definable.rel.retraction (n := k + 3) (1 :> (·.succ.succ.succ))) - <| and (by simpa using hf₃.definable.rel.retraction (n := k + 3) (0 :> (·.succ.succ.succ))) - <| by simpa using hR.retraction (n := k + 3) ![2, 1, 0] - exact this.of_iff _ (by - intro v; constructor - · intro h; exact ⟨f₁ v, hbf₁ v, f₂ v, hbf₂ v, f₃ v, hbf₃ v, rfl, rfl, rfl, h⟩ - · rintro ⟨_, _, _, _, _, _, rfl, rfl, rfl, h⟩; exact h) - -lemma bcomp₄ {k} {R : M → M → M → M → Prop} {f₁ f₂ f₃ f₄ : (Fin k → M) → M} - [hR : DefinableRel₄ L Γ R] (hf₁ : DefinableBoundedFunction L Γ f₁) (hf₂ : DefinableBoundedFunction L Γ f₂) (hf₃ : DefinableBoundedFunction L Γ f₃) (hf₄ : DefinableBoundedFunction L Γ f₄) : - Definable L Γ (fun v ↦ R (f₁ v) (f₂ v) (f₃ v) (f₄ v)) := by - rcases hf₁.bounded with ⟨bf₁, hbf₁⟩ - rcases hf₂.bounded with ⟨bf₂, hbf₂⟩ - rcases hf₃.bounded with ⟨bf₃, hbf₃⟩ - rcases hf₄.bounded with ⟨bf₄, hbf₄⟩ - have : Definable L Γ (fun v ↦ - ∃ z₁ ≤ Semiterm.valm M v id bf₁, ∃ z₂ ≤ Semiterm.valm M v id bf₂, ∃ z₃ ≤ Semiterm.valm M v id bf₃, ∃ z₄ ≤ Semiterm.valm M v id bf₄, - z₁ = f₁ v ∧ z₂ = f₂ v ∧ z₃ = f₃ v ∧ z₄ = f₄ v ∧ R z₁ z₂ z₃ z₄) := - bex_le₀ (DefinableBoundedFunction.term _) <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) - <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) <| bex_le₀ (DefinableBoundedFunction.term_retraction _ _) - <| and (by simpa using hf₁.definable.rel.retraction (n := k + 4) (3 :> (·.succ.succ.succ.succ))) - <| and (by simpa using hf₂.definable.rel.retraction (n := k + 4) (2 :> (·.succ.succ.succ.succ))) - <| and (by simpa using hf₃.definable.rel.retraction (n := k + 4) (1 :> (·.succ.succ.succ.succ))) - <| and (by simpa using hf₄.definable.rel.retraction (n := k + 4) (0 :> (·.succ.succ.succ.succ))) - <| by simpa using hR.retraction (n := k + 4) ![3, 2, 1, 0] - exact this.of_iff _ (by - intro v; constructor - · intro h; exact ⟨f₁ v, hbf₁ v, f₂ v, hbf₂ v, f₃ v, hbf₃ v, f₄ v, hbf₄ v, rfl, rfl, rfl, rfl, h⟩ - · rintro ⟨_, _, _, _, _, _, _, _, rfl, rfl, rfl, rfl, h⟩; exact h) - -end Definable - -lemma DefinableFunction₁.bcomp {k} {f : M → M} {g : (Fin k → M) → M} - (hf : DefinableFunction₁ L Γ f) (hg : DefinableBoundedFunction L Γ g) : - DefinableFunction L Γ (fun v ↦ f (g v)) := by - have := Definable.bcomp₂ (k := k + 1) (R := Function.Graph f) (DefinableBoundedFunction.var 0) (hg.retraction Fin.succ) - simpa using this - -lemma DefinableFunction₂.bcomp {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} - (hf : DefinableFunction₂ L Γ f) (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) : - DefinableFunction L Γ (fun v ↦ f (g₁ v) (g₂ v)) := by - have := Definable.bcomp₃ (k := k + 1) (R := Function.Graph₂ f) (DefinableBoundedFunction.var 0) (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) - simpa using this - -lemma DefinableFunction₃.bcomp {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} - (hf : DefinableFunction₃ L Γ f) (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) (hg₃ : DefinableBoundedFunction L Γ g₃) : - DefinableFunction L Γ (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := by - have := Definable.bcomp₄ (k := k + 1) (R := Function.Graph₃ f) (DefinableBoundedFunction.var 0) (hg₁.retraction Fin.succ) (hg₂.retraction Fin.succ) (hg₃.retraction Fin.succ) - simpa using this - -lemma DefinableBoundedFunction₁.comp {k} {f : M → M} {g : (Fin k → M) → M} (hf : DefinableBoundedFunction₁ L Γ f) (hg : DefinableBoundedFunction L Γ g) : - DefinableBoundedFunction L Γ (fun v ↦ f (g v)) := ⟨hf.bounded.comp hg.bounded, hf.definable.bcomp hg⟩ - -lemma DefinableBoundedFunction₂.comp {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} - (hf : DefinableBoundedFunction₂ L Γ f) (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) : - DefinableBoundedFunction L Γ (fun v ↦ f (g₁ v) (g₂ v)) := ⟨hf.bounded.comp hg₁.bounded hg₂.bounded, hf.definable.bcomp hg₁ hg₂⟩ - -lemma DefinableBoundedFunction₃.comp {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} - (hf : DefinableBoundedFunction₃ L Γ f) (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) (hg₃ : DefinableBoundedFunction L Γ g₃) : - DefinableBoundedFunction L Γ (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := ⟨hf.bounded.comp hg₁.bounded hg₂.bounded hg₃.bounded, hf.definable.bcomp hg₁ hg₂ hg₃⟩ - -lemma DefinableBoundedFunction.comp₁ {k} {f : M → M} {g : (Fin k → M) → M} - [hfb : Bounded₁ L f] [hfd : DefinableFunction₁ L Γ f] (hg : DefinableBoundedFunction L Γ g) : - DefinableBoundedFunction L Γ (fun v ↦ f (g v)) := DefinableBoundedFunction₁.comp ⟨hfb, hfd⟩ hg - -lemma DefinableBoundedFunction.comp₂ {k} {f : M → M → M} {g₁ g₂ : (Fin k → M) → M} - [hfb : Bounded₂ L f] [hfd : DefinableFunction₂ L Γ f] (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) : - DefinableBoundedFunction L Γ (fun v ↦ f (g₁ v) (g₂ v)) := DefinableBoundedFunction₂.comp ⟨hfb, hfd⟩ hg₁ hg₂ - -lemma DefinableBoundedFunction.comp₃ {k} {f : M → M → M → M} {g₁ g₂ g₃ : (Fin k → M) → M} - [hfb : Bounded₃ L f] [hfd : DefinableFunction₃ L Γ f] (hg₁ : DefinableBoundedFunction L Γ g₁) (hg₂ : DefinableBoundedFunction L Γ g₂) (hg₃ : DefinableBoundedFunction L Γ g₃) : - DefinableBoundedFunction L Γ (fun v ↦ f (g₁ v) (g₂ v) (g₃ v)) := DefinableBoundedFunction₃.comp ⟨hfb, hfd⟩ hg₁ hg₂ hg₃ - -section - --- https://github.com/leanprover-community/mathlib4/blob/77d078e25cc501fae6907bfbcd80821920125266/Mathlib/Tactic/Measurability.lean#L25-L26 -open Lean.Parser.Tactic (config) - -open Definable - -lemma DefinablePred.infer {R : M → Prop} [DefinablePred L Γ R] : DefinablePred L Γ R := inferInstance -lemma DefinableRel.infer {R : M → M → Prop} [DefinableRel L Γ R] : DefinableRel L Γ R := inferInstance -lemma DefinableRel₃.infer {R : M → M → M → Prop} [DefinableRel₃ L Γ R] : DefinableRel₃ L Γ R := inferInstance -lemma DefinableFunction₁.infer {f : M → M} [DefinableFunction₁ L Γ f] : DefinableFunction₁ L Γ f := inferInstance -lemma DefinableFunction₂.infer {f : M → M → M} [DefinableFunction₂ L Γ f] : DefinableFunction₂ L Γ f := inferInstance -lemma DefinableFunction₃.infer {f : M → M → M → M} [DefinableFunction₃ L Γ f] : DefinableFunction₃ L Γ f := inferInstance - -attribute [aesop (rule_sets := [Definability]) norm] - sq - Arith.pow_three - pow_four - Definable.const - -attribute [aesop 5 (rule_sets := [Definability]) safe] - DefinableFunction.comp₁_infer - DefinableFunction.comp₂_infer - DefinableFunction.comp₃_infer - DefinableBoundedFunction.comp₁ - DefinableBoundedFunction.comp₂ - DefinableBoundedFunction.comp₃ - -attribute [aesop 6 (rule_sets := [Definability]) safe] - Definable.comp₁_infer - Definable.comp₂_infer - Definable.comp₃_infer - Definable.comp₄_infer - Definable.const - -attribute [aesop 7 (rule_sets := [Definability]) safe] - Definable.bcomp₁ - Definable.bcomp₂ - Definable.bcomp₃ - Definable.bcomp₄ - -attribute [aesop 8 (rule_sets := [Definability]) safe] - Definable.ball_lt - Definable.ball_le - Definable.ball_lt' - Definable.ball_le' - Definable.bex_lt - Definable.bex_le - -attribute [aesop 9 (rule_sets := [Definability]) safe] - Definable.ball_lt₀ - Definable.ball_le₀ - Definable.bex_lt₀ - Definable.bex_le₀ - -attribute [aesop 10 (rule_sets := [Definability]) safe] - Definable.not - Definable.imp - Definable.iff - -attribute [aesop 11 (rule_sets := [Definability]) safe] - Definable.and - Definable.or - Definable.all - Definable.ex - -macro "definability" : attr => - `(attr|aesop 10 (rule_sets := [$(Lean.mkIdent `Definability):ident]) safe) - -macro "definability" (config)? : tactic => - `(tactic| aesop (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) - -macro "definability?" (config)? : tactic => - `(tactic| aesop? (config := { terminal := true }) (rule_sets := [$(Lean.mkIdent `Definability):ident])) - -example (c : M) : DefinableBoundedFunction₂ L (𝚺, 0) (fun x y : M ↦ c + 2 * x^2) := by definability - -example {ex : M → M} [DefinableFunction₁ L 𝚺₀ ex] (c : M) : - DefinableRel L 𝚷₀ (fun x y : M ↦ ∃ z < x + c * y, (ex x = x ∧ x < y) ↔ ex x = z ∧ ex (x + 1) = 2 * z) := by - simp [Function.Graph.iff_left ex] - definability? - -example {ex : M → M} [h : DefinableFunction₁ L (𝚫, 1) ex] (c : M) : - DefinableRel L (𝚺, 1) (fun x y : M ↦ ∃ z, x < y ↔ ex (ex x) = z) := by - definability? - -example {ex : M → M} [h : DefinableFunction₁ L (𝚺, 1) ex] (c : M) : - DefinableRel L (𝚺, 1) (fun x y : M ↦ ∀ z < ex y, x < y ↔ ex (ex x) = z) := by - definability? - -end - -end LO.FirstOrder.Arith diff --git a/Arithmetization/Definability/SigmaPiDeltaSystem.lean b/Arithmetization/Definability/SigmaPiDeltaSystem.lean deleted file mode 100644 index 8472fb7..0000000 --- a/Arithmetization/Definability/SigmaPiDeltaSystem.lean +++ /dev/null @@ -1,299 +0,0 @@ -import Arithmetization.Vorspiel.Lemmata -import Logic.FirstOrder.Arith.StrictHierarchy - -/-! - -# Arithmetical Formula Sorted by Arithmetical Hierarchy - -This file defines the $\Sigma_n / \Pi_n / \Delta_n$ formulas of arithmetic of first-order logic. - -- `𝚺-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚺-[m]`. -- `𝚷-[m].Semiformula ξ n` is a `Semiformula ℒₒᵣ ξ n` which is `𝚷-[m]`. -- `𝚫-[m].Semiformula ξ n` is a pair of `𝚺-[m].Semiformula ξ n` and `𝚷-[m].Semiformula ξ n`. -- `ProperOn` : `p.ProperOn M` iff `p`'s two element `p.sigma` and `p.pi` are equivalent on model `M`. - --/ - -namespace LO.Arith - -class SigmaPiDeltaLike (Ω : Type*) [SigmaSymbol Ω] [PiSymbol Ω] [DeltaSymbol Ω] where - alt : Ω → Ω - -variable {V : Type*} - -structure SigmaPiDeltaSystem (V : Type*) where - VecPr : SigmaPiDelta → {k : ℕ} → ((Fin k → V) → Prop) → Prop - vecPr_delta_iff_sigma_and_pi {k} {P : (Fin k → V) → Prop} : VecPr 𝚫 P ↔ VecPr 𝚺 P ∧ VecPr 𝚷 P - verum (Γ k) : VecPr Γ fun _ : Fin k → V ↦ ⊤ - and {k} {P Q : (Fin k → V) → Prop} : VecPr Γ P → VecPr Γ Q → VecPr Γ fun x ↦ P x ∧ Q x - not {k} {P : (Fin k → V) → Prop} : VecPr Γ.alt P → VecPr Γ fun x ↦ ¬P x - all {k} {P : (Fin k → V) → V → Prop} : VecPr 𝚷 (fun x ↦ P (x ·.succ) (x 0)) → VecPr 𝚷 fun x ↦ ∀ z, P x z - retraction {k l} {P : (Fin k → V) → Prop} (hP : VecPr Γ P) (f : Fin k → Fin l) : VecPr Γ fun v ↦ P fun i ↦ v (f i) - equal (Γ) : VecPr Γ fun v : Fin 2 → V ↦ v 0 = v 1 - -abbrev SigmaPiDeltaSystem.VecFunc (𝔖 : SigmaPiDeltaSystem V) - (Γ : SigmaPiDelta) (f : (Fin k → V) → V) : Prop := 𝔖.VecPr Γ fun v ↦ v 0 = f (v ·.succ) - -namespace SigmaPiDeltaSystem - -variable {𝔖 : SigmaPiDeltaSystem V} {Γ : SigmaPiDelta} {k} {P Q : (Fin k → V) → Prop} - -/- -variable (𝔖 Γ) - -abbrev Pred (P : V → Prop) : Prop := 𝔖.VecPr Γ (k := 1) (fun v ↦ P (v 0)) - -abbrev Rel (P : V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 2) (fun v ↦ P (v 0) (v 1)) - -abbrev Rel₃ (P : V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 3) (fun v ↦ P (v 0) (v 1) (v 2)) - -abbrev Rel₄ (P : V → V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 4) (fun v ↦ P (v 0) (v 1) (v 2) (v 3)) - -abbrev Rel₅ (P : V → V → V → V → V → Prop) : Prop := 𝔖.VecPr Γ (k := 5) (fun v ↦ P (v 0) (v 1) (v 2) (v 3) (v 4)) - -abbrev Function (f : V → V) : Prop := 𝔖.VecFunc Γ (k := 1) (fun v ↦ f (v 0)) - -abbrev Function₂ (f : V → V → V) : Prop := 𝔖.VecFunc Γ (k := 2) (fun v ↦ f (v 0) (v 1)) - -abbrev Function₃ (f : V → V → V → V) : Prop := 𝔖.VecFunc Γ (k := 3) (fun v ↦ f (v 0) (v 1) (v 2)) - -abbrev Function₄ (f : V → V → V → V → V) : Prop := 𝔖.VecFunc Γ (k := 4) (fun v ↦ f (v 0) (v 1) (v 2) (v 3)) - --/ - -lemma of_iff (hP : 𝔖.VecPr Γ P) (h : ∀ x, P x ↔ Q x) : 𝔖.VecPr Γ Q := by - have : P = Q := funext <| by simpa - rcases this - exact hP - -lemma of_sigma_of_pi (hσ : 𝔖.VecPr 𝚺 P) (hπ : 𝔖.VecPr 𝚷 P) : 𝔖.VecPr Γ P := - match Γ with - | 𝚺 => hσ - | 𝚷 => hπ - | 𝚫 => (vecPr_delta_iff_sigma_and_pi _).mpr ⟨hσ, hπ⟩ - -lemma of_delta (h : 𝔖.VecPr 𝚫 P) {Γ} : 𝔖.VecPr Γ P := - of_sigma_of_pi - ((vecPr_delta_iff_sigma_and_pi _).mp h |>.1) - ((vecPr_delta_iff_sigma_and_pi _).mp h |>.2) - -lemma not' (h : 𝔖.VecPr Γ P) : 𝔖.VecPr Γ.alt fun x ↦ ¬P x := 𝔖.not (by simpa) - -lemma of_not (h : 𝔖.VecPr Γ.alt (fun x ↦ ¬P x)) : 𝔖.VecPr Γ P := by simpa using not' h - -lemma falsum (Γ : SigmaPiDelta) (k : ℕ) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ ⊥ := - of_sigma_of_pi (by simpa using not' (𝔖.verum 𝚷 k)) (by simpa using not' (𝔖.verum 𝚺 k)) - -@[simp] lemma constant (Γ : SigmaPiDelta) (k : ℕ) (P : Prop) : 𝔖.VecPr Γ fun _ : Fin k → V ↦ P := by - by_cases h : P <;> simp [h] - · apply verum - · apply falsum - -lemma or (hP : 𝔖.VecPr Γ P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x ∨ Q x := - of_not <| by - simp only [not_or]; apply and - · apply not' hP - · apply not' hQ - -lemma imply (hP : 𝔖.VecPr Γ.alt P) (hQ : 𝔖.VecPr Γ Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x → Q x := by - simp only [imp_iff_not_or]; apply or - · apply 𝔖.not hP - · exact hQ - -lemma ex {k} {P : (Fin k → V) → V → Prop} (h : 𝔖.VecPr 𝚺 fun x ↦ P (x ·.succ) (x 0)) : 𝔖.VecPr 𝚺 fun x ↦ ∃ z, P x z := of_not <| by - simpa using 𝔖.all (by apply not' h) - -lemma iff (hP : 𝔖.VecPr 𝚫 P) (hQ : 𝔖.VecPr 𝚫 Q) : 𝔖.VecPr Γ fun x : Fin k → V ↦ P x ↔ Q x := of_delta <| by - simp only [iff_iff_implies_and_implies] - apply and - · exact imply hP hQ - · exact imply hQ hP - -lemma equal' (Γ) (i j : Fin k) : 𝔖.VecPr Γ fun v ↦ v i = v j := by - simpa using 𝔖.retraction (𝔖.equal Γ) ![i, j] - -lemma VecFunc.of_sigma {f : (Fin k → V) → V} (h : 𝔖.VecFunc 𝚺 f) {Γ} : 𝔖.VecFunc Γ f := by - apply of_sigma_of_pi - · exact h - · have : 𝔖.VecPr 𝚷 fun v ↦ ∀ y, y = f (v ·.succ) → v 0 = y := 𝔖.all <| imply - (by simpa using 𝔖.retraction h (0 :> (·.succ.succ))) - (by simpa using equal' 𝚷 1 0) - exact of_iff this (fun v ↦ by simp) - -lemma conj {k l} {P : Fin l → (Fin k → V) → Prop} - (h : ∀ i, 𝔖.VecPr Γ fun w : Fin k → V ↦ P i w) : - 𝔖.VecPr Γ fun v : Fin k → V ↦ ∀ i, P i v := by - induction l - case zero => simp - case succ l ih => - suffices 𝔖.VecPr Γ fun v : Fin k → V ↦ P 0 v ∧ ∀ i : Fin l, P i.succ v by - apply of_iff this; intro x - constructor - · rintro ⟨h0, hs⟩ - intro i; cases' i using Fin.cases with i - · exact h0 - · exact hs i - · intro h - exact ⟨h 0, fun i ↦ h i.succ⟩ - apply 𝔖.and (h 0); apply ih - intro i; exact h i.succ - -lemma exVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} - (h : 𝔖.VecPr 𝚺 fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : - 𝔖.VecPr 𝚺 fun v : Fin k → V ↦ ∃ ys : Fin l → V, P v ys := by - induction l generalizing k - case zero => simpa [Matrix.empty_eq] using h - case succ l ih => - suffices 𝔖.VecPr 𝚺 fun v : Fin k → V ↦ ∃ y, ∃ ys : Fin l → V, P v (y :> ys) by - apply of_iff this; intro x - constructor - · rintro ⟨y, ys, h⟩; exact ⟨_, h⟩ - · rintro ⟨ys, h⟩; exact ⟨ys 0, (ys ·.succ), by simpa using h⟩ - apply ex; apply ih - let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (𝔖.retraction h g) (by - intro v; simp [g] - apply iff_of_eq; congr - · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · ext i - cases' i using Fin.cases with i - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · congr 1; ext; simp [Matrix.vecAppend_eq_ite]) - -lemma allVec {k l} {P : (Fin k → V) → (Fin l → V) → Prop} - (h : 𝔖.VecPr 𝚷 fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : - 𝔖.VecPr 𝚷 fun v : Fin k → V ↦ ∀ ys : Fin l → V, P v ys := by - induction l generalizing k - case zero => simpa [Matrix.empty_eq] using h - case succ l ih => - suffices 𝔖.VecPr 𝚷 fun v : Fin k → V ↦ ∀ y, ∀ ys : Fin l → V, P v (y :> ys) by - apply of_iff this; intro x - constructor - · intro h ys; simpa using h (ys 0) (ys ·.succ) - · intro h y ys; apply h - apply all; apply ih - let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) - exact of_iff (𝔖.retraction h g) (by - intro v; simp [g] - apply iff_of_eq; congr - · ext i; congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · ext i - cases' i using Fin.cases with i - · congr 1; ext; simp [Matrix.vecAppend_eq_ite] - · congr 1; ext; simp [Matrix.vecAppend_eq_ite]) - -private lemma substitution_sigma {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr 𝚺 P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : - 𝔖.VecPr 𝚺 fun z ↦ P (fun i ↦ f i z) := by - have : 𝔖.VecPr 𝚺 fun z ↦ ∃ ys : Fin k → V, (∀ i, ys i = f i z) ∧ P ys := by - apply exVec; apply and - · apply conj; intro i - simpa using 𝔖.retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) - · exact 𝔖.retraction hP (Fin.natAdd l) - exact of_iff this <| by - intro v - constructor - · rintro ⟨ys, hys, hP⟩ - have : ys = fun i ↦ f i v := funext hys - rcases this; exact hP - · intro hP - exact ⟨(f · v), by simp, hP⟩ - -private lemma substitution_pi {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr 𝚷 P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : - 𝔖.VecPr 𝚷 fun z ↦ P (fun i ↦ f i z) := by - have : 𝔖.VecPr 𝚷 fun z ↦ ∀ ys : Fin k → V, (∀ i, ys i = f i z) → P ys := by - apply allVec; apply imply - · apply conj; intro i - simpa using 𝔖.retraction (VecFunc.of_sigma (hf i)) (i.natAdd l :> fun i ↦ i.castAdd k) - · exact 𝔖.retraction hP (Fin.natAdd l) - exact of_iff this <| by - intro v - constructor - · intro h; apply h _ (by simp) - · intro h ys e - have : ys = (f · v) := funext e - rcases this; exact h - -lemma substitution {f : Fin k → (Fin l → V) → V} (hP : 𝔖.VecPr Γ P) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : 𝔖.VecPr Γ fun z ↦ P (fun i ↦ f i z) := - match Γ with - | 𝚺 => substitution_sigma hP hf - | 𝚷 => substitution_pi hP hf - | 𝚫 => of_sigma_of_pi (substitution_sigma (of_delta hP) hf) (substitution_pi (of_delta hP) hf) - -namespace VecFunc - -variable {F : (Fin k → V) → V} - -lemma nth (Γ) (i : Fin k) : 𝔖.VecFunc Γ fun w ↦ w i := 𝔖.equal' Γ 0 i.succ - -lemma substitution {f : Fin k → (Fin l → V) → V} (hF : 𝔖.VecFunc Γ F) (hf : ∀ i, 𝔖.VecFunc 𝚺 (f i)) : - 𝔖.VecFunc Γ fun z ↦ F (fun i ↦ f i z) := by - simp only [VecFunc, Nat.succ_eq_add_one] - simpa using 𝔖.substitution (f := (· 0) :> fun i w ↦ f i (w ·.succ)) hF - (by intro i - cases' i using Fin.cases with i - · simpa using nth 𝚺 0 - · simpa using 𝔖.retraction (hf i) (0 :> (·.succ.succ))) - -end VecFunc - -variable [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] - -class Arithmetical (𝔖 : SigmaPiDeltaSystem V) where - zero (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 0 - one (Γ) : 𝔖.VecFunc Γ fun _ : Fin 1 → V ↦ 1 - add (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 + v 1 - mul (Γ) : 𝔖.VecFunc Γ fun v : Fin 2 → V ↦ v 0 * v 1 - lt (Γ) : 𝔖.VecPr Γ fun v : Fin 2 → V ↦ v 0 < v 1 - ball {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∀ x ≤ v 0, P (v ·.succ) x - bex' {Γ} {P : (Fin k → V) → V → Prop} : 𝔖.VecPr Γ (fun v ↦ P (v ·.succ) (v 0)) → 𝔖.VecPr Γ fun v ↦ ∃ x ≤ v 0, P (v ·.succ) x - - - -end SigmaPiDeltaSystem - -/- -class EmbeddingType (V : outParam Type*) (β : Type*) where - embedding : β ↪ V - -namespace EmbeddingType - -instance : EmbeddingType V V := ⟨Function.Embedding.refl V⟩ - -instance (p : V → Prop) : EmbeddingType V (Subtype p) := ⟨Function.Embedding.subtype p⟩ - -end EmbeddingType - -namespace SigmaPiDeltaSystem - -class Class {V : Type*} (𝔖 : SigmaPiDeltaSystem V) (α : Type*) [EmbeddingType V α] where - delta : 𝔖.VecPr 𝚫 fun x : Fin 1 → V ↦ x 0 ∈ Set.range (EmbeddingType.embedding : α ↪ V) - -section Class - -instance (𝔖 : SigmaPiDeltaSystem V) : 𝔖.Class V where - delta := VecPr.of_iff (𝔖.verum' 𝚫 1) <| by intro v; simp; exact ⟨v 0, by rfl⟩ - -variable {𝔖 : SigmaPiDeltaSystem V} - -variable {α β γ δ ε ζ : Type*} - [EmbeddingType V α] [EmbeddingType V β] [EmbeddingType V γ] [EmbeddingType V δ] [EmbeddingType V ε] [EmbeddingType V ζ] - [𝔖.Class α] [𝔖.Class β] [𝔖.Class γ] [𝔖.Class δ] [𝔖.Class ε] [𝔖.Class ζ] - -def Pr₁ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → Prop) : Prop := 𝔖.VecPr Γ fun x : Fin 1 → V ↦ ∃ a : α, x 0 = EmbeddingType.embedding a ∧ P a -def Pr₂ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → Prop) : Prop := - 𝔖.VecPr Γ fun x : Fin 2 → V ↦ ∃ (a : α) (b : β), x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ P a b -def Pr₃ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → γ → Prop) : Prop := - 𝔖.VecPr Γ fun x : Fin 3 → V ↦ ∃ (a : α) (b : β) (c : γ), - x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ x 2 = EmbeddingType.embedding c ∧ P a b c -def Pr₄ (𝔖 : SigmaPiDeltaSystem V) (Γ : SigmaPiDelta) (P : α → β → γ → δ → Prop) : Prop := - 𝔖.VecPr Γ fun x : Fin 4 → V ↦ ∃ (a : α) (b : β) (c : γ) (d : δ), - x 0 = EmbeddingType.embedding a ∧ x 1 = EmbeddingType.embedding b ∧ x 2 = EmbeddingType.embedding c ∧ P a b c d - - -end Class - -end SigmaPiDeltaSystem - --/ - -end LO diff --git a/Arithmetization/Vorspiel/Lemmata.lean b/Arithmetization/Vorspiel/Lemmata.lean index 60531a1..7490428 100644 --- a/Arithmetization/Vorspiel/Lemmata.lean +++ b/Arithmetization/Vorspiel/Lemmata.lean @@ -78,7 +78,7 @@ namespace Arith noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] +variable {M : Type*} [ORingStruc M] [M ⊧ₘ* 𝐏𝐀⁻] variable {a b c : M} diff --git a/Arithmetization/Vorspiel/Vorspiel.lean b/Arithmetization/Vorspiel/Vorspiel.lean index 15b76d0..7795e15 100644 --- a/Arithmetization/Vorspiel/Vorspiel.lean +++ b/Arithmetization/Vorspiel/Vorspiel.lean @@ -187,7 +187,7 @@ section model variable {T : Theory ℒₒᵣ} [𝐄𝐐 ≼ T] -variable (M : Type*) [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* T] +variable (M : Type*) [ORingStruc M] [M ⊧ₘ* T] lemma oring_sound {σ : Sentence ℒₒᵣ} (h : T ⊢! σ) : M ⊧ₘ σ := (consequence_iff' (T := T)).mp (LO.Sound.sound h) M @@ -251,7 +251,7 @@ end end Hierarchy -variable (M : Type*) [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] +variable (M : Type*) [ORingStruc M] [M ⊧ₘ* 𝐏𝐀⁻] lemma nat_extention_sigmaOne {σ : Sentence ℒₒᵣ} (hσ : Hierarchy 𝚺 1 σ) : ℕ ⊧ₘ σ → M ⊧ₘ σ := fun h ↦ by diff --git a/lake-manifest.json b/lake-manifest.json index a7c93de..597868a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -75,7 +75,7 @@ "type": "git", "subDir": null, "scope": "", - "rev": "107bf1f502b48cba1e9eb10e4f5faa5d871ebec0", + "rev": "352b7bd19ddb940b29fb5abb91e94a53317fc5ba", "name": "logic", "manifestFile": "lake-manifest.json", "inputRev": "master", From 7c84d331ae9a0961f51e7d2c7a8c9e82f5566e7a Mon Sep 17 00:00:00 2001 From: palalansouki Date: Fri, 16 Aug 2024 03:42:09 +0900 Subject: [PATCH 06/10] refactor `Basic` --- Arithmetization/Basic/IOpen.lean | 376 +++++++++--------- Arithmetization/Basic/Ind.lean | 167 ++++---- Arithmetization/Basic/PeanoMinus.lean | 88 ++-- .../Definability/Absoluteness.lean | 30 +- Arithmetization/Definability/Boldface.lean | 43 +- .../Definability/BoundedBoldface.lean | 77 ++-- Arithmetization/ISigmaOne/Bit.lean | 22 +- Arithmetization/ISigmaOne/HFS/Basic.lean | 30 +- Arithmetization/ISigmaOne/HFS/Fixpoint.lean | 4 +- Arithmetization/ISigmaOne/HFS/PRF.lean | 4 +- Arithmetization/ISigmaOne/HFS/Seq.lean | 18 +- .../ISigmaOne/HFS/Supplemental.lean | 2 +- Arithmetization/ISigmaOne/HFS/Vec.lean | 18 +- Arithmetization/ISigmaOne/Ind.lean | 8 +- .../ISigmaOne/Metamath/Formula/Basic.lean | 20 +- .../ISigmaOne/Metamath/Formula/Functions.lean | 8 +- .../ISigmaOne/Metamath/Formula/Iteration.lean | 12 +- .../ISigmaOne/Metamath/Proof/Derivation.lean | 24 +- .../ISigmaOne/Metamath/Proof/Theory.lean | 4 +- .../ISigmaOne/Metamath/Term/Basic.lean | 6 +- .../ISigmaOne/Metamath/Term/Functions.lean | 8 +- .../ISigmaZero/Exponential/Exp.lean | 38 +- .../ISigmaZero/Exponential/Log.lean | 26 +- .../ISigmaZero/Exponential/PPow2.lean | 8 +- .../ISigmaZero/Exponential/Pow2.lean | 10 +- Arithmetization/OmegaOne/Basic.lean | 4 +- Arithmetization/OmegaOne/Nuon.lean | 32 +- 27 files changed, 567 insertions(+), 520 deletions(-) diff --git a/Arithmetization/Basic/IOpen.lean b/Arithmetization/Basic/IOpen.lean index 32a5064..d3c9fe7 100644 --- a/Arithmetization/Basic/IOpen.lean +++ b/Arithmetization/Basic/IOpen.lean @@ -7,25 +7,25 @@ open FirstOrder FirstOrder.Arith noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] +variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐏𝐀⁻] section IOpen -variable [M ⊧ₘ* 𝐈open] +variable [V ⊧ₘ* 𝐈open] @[elab_as_elim] -lemma open_induction {P : M → Prop} - (hP : ∃ p : Semiformula ℒₒᵣ M 1, p.Open ∧ ∀ x, P x ↔ Semiformula.Evalm M ![x] id p) +lemma open_induction {P : V → Prop} + (hP : ∃ p : Semiformula ℒₒᵣ V 1, p.Open ∧ ∀ x, P x ↔ Semiformula.Evalm V ![x] id p) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := - induction ℒₒᵣ (C := Semiformula.Open) + induction (C := Semiformula.Open) (by rcases hP with ⟨p, hp, hhp⟩ - haveI : Inhabited M := Classical.inhabited_of_nonempty' + haveI : Inhabited V := Classical.inhabited_of_nonempty' exact ⟨p.fvEnumInv, (Rew.rewriteMap p.fvEnum).hom p, by simp[hp], by intro x; simp [Semiformula.eval_rewriteMap, hhp] exact Semiformula.eval_iff_of_funEqOn p (by intro z hz; simp [Semiformula.fvEnumInv_fvEnum hz])⟩) zero succ -lemma open_leastNumber {P : M → Prop} - (hP : ∃ p : Semiformula ℒₒᵣ M 1, p.Open ∧ ∀ x, P x ↔ Semiformula.Evalm M ![x] id p) +lemma open_leastNumber {P : V → Prop} + (hP : ∃ p : Semiformula ℒₒᵣ V 1, p.Open ∧ ∀ x, P x ↔ Semiformula.Evalm V ![x] id p) (zero : P 0) {a} (counterex : ¬P a) : ∃ x, P x ∧ ¬P (x + 1) := by by_contra A have : ∀ x, P x := by @@ -38,7 +38,7 @@ lemma open_leastNumber {P : M → Prop} have : P a := this a contradiction -lemma div_exists_unique_pos (a : M) {b} (pos : 0 < b) : ∃! u, b * u ≤ a ∧ a < b * (u + 1) := by +lemma div_exists_unique_pos (a : V) {b} (pos : 0 < b) : ∃! u, b * u ≤ a ∧ a < b * (u + 1) := by have : ∃ u, b * u ≤ a ∧ a < b * (u + 1) := by have : a < b * (a + 1) → ∃ u, b * u ≤ a ∧ a < b * (u + 1) := by simpa using open_leastNumber (P := fun u ↦ b * u ≤ a) ⟨“x | &b * x ≤ &a”, by simp, by intro x; simp⟩ @@ -61,7 +61,7 @@ lemma div_exists_unique_pos (a : M) {b} (pos : 0 < b) : ∃! u, b * u ≤ a ∧ exact LT.lt.false this) /- -lemma mod (a : M) {b} (pos : 0 < b) : ∃! u, ∃ v < b, a = b * u + v := by +lemma mod (a : V) {b} (pos : 0 < b) : ∃! u, ∃ v < b, a = b * u + v := by have : ∃! u, b * u ≤ a ∧ a < b * (u + 1) := by have : ∃ u, b * u ≤ a ∧ a < b * (u + 1) := by have : a < b * (a + 1) → ∃ u, b * u ≤ a ∧ a < b * (u + 1) := by @@ -104,18 +104,18 @@ lemma mod (a : M) {b} (pos : 0 < b) : ∃! u, ∃ v < b, a = b * u + v := by section div -lemma div_exists_unique (a b : M) : ∃! u, (0 < b → b * u ≤ a ∧ a < b * (u + 1)) ∧ (b = 0 → u = 0) := by +lemma div_exists_unique (a b : V) : ∃! u, (0 < b → b * u ≤ a ∧ a < b * (u + 1)) ∧ (b = 0 → u = 0) := by have : 0 ≤ b := by exact zero_le b rcases this with (rfl | pos) <;> simp [*] · simpa [pos_iff_ne_zero.mp pos] using div_exists_unique_pos a pos -scoped instance : Div M := ⟨fun a b ↦ Classical.choose! (div_exists_unique a b)⟩ +scoped instance : Div V := ⟨fun a b ↦ Classical.choose! (div_exists_unique a b)⟩ -lemma mul_div_le_pos (a : M) (h : 0 < b) : b * (a / b) ≤ a := ((Classical.choose!_spec (div_exists_unique a b)).1 h).1 +lemma mul_div_le_pos (a : V) (h : 0 < b) : b * (a / b) ≤ a := ((Classical.choose!_spec (div_exists_unique a b)).1 h).1 -lemma lt_mul_div_succ (a : M) (h : 0 < b) : a < b * (a / b + 1) := ((Classical.choose!_spec (div_exists_unique a b)).1 h).2 +lemma lt_mul_div_succ (a : V) (h : 0 < b) : a < b * (a / b + 1) := ((Classical.choose!_spec (div_exists_unique a b)).1 h).2 -lemma eq_mul_div_add_of_pos (a : M) {b} (hb : 0 < b) : ∃ r < b, a = b * (a / b) + r := by +lemma eq_mul_div_add_of_pos (a : V) {b} (hb : 0 < b) : ∃ r < b, a = b * (a / b) + r := by let r := a - b * (a / b) have e : a = b * (a / b) + r := by simp [add_tsub_self_of_le (mul_div_le_pos a hb)] exact ⟨r, by @@ -127,38 +127,38 @@ lemma eq_mul_div_add_of_pos (a : M) {b} (hb : 0 < b) : ∃ r < b, a = b * (a / b _ = a := e.symm simp at this, e⟩ -@[simp] lemma div_spec_zero (a : M) : a / 0 = 0 := (Classical.choose!_spec (div_exists_unique a 0)).2 (by simp) +@[simp] lemma div_spec_zero (a : V) : a / 0 = 0 := (Classical.choose!_spec (div_exists_unique a 0)).2 (by simp) -lemma div_graph {a b c : M} : c = a / b ↔ ((0 < b → b * c ≤ a ∧ a < b * (c + 1)) ∧ (b = 0 → c = 0)) := +lemma div_graph {a b c : V} : c = a / b ↔ ((0 < b → b * c ≤ a ∧ a < b * (c + 1)) ∧ (b = 0 → c = 0)) := Classical.choose!_eq_iff _ -def _root_.LO.FirstOrder.Arith.divDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.divDef : 𝚺₀.Semisentence 3 := .mkSigma “c a b | (0 < b → b * c ≤ a ∧ a < b * (c + 1)) ∧ (b = 0 → c = 0)” (by simp[Hierarchy.pi_zero_iff_sigma_zero]) -lemma div_defined : 𝚺₀-Function₂ ((· / ·) : M → M → M) via divDef := by +lemma div_defined : 𝚺₀-Function₂ ((· / ·) : V → V → V) via divDef := by intro v; simp[div_graph, divDef, Matrix.vecHead, Matrix.vecTail] @[simp] lemma div_defined_iff (v) : - Semiformula.Evalbm M v divDef.val ↔ v 0 = v 1 / v 2 := div_defined.df.iff v + Semiformula.Evalbm V v divDef.val ↔ v 0 = v 1 / v 2 := div_defined.df.iff v -lemma div_spec_of_pos' (a : M) (h : 0 < b) : ∃ v < b, a = (a / b) * b + v := by +lemma div_spec_of_pos' (a : V) (h : 0 < b) : ∃ v < b, a = (a / b) * b + v := by simpa [mul_comm] using eq_mul_div_add_of_pos a h -lemma div_eq_of {b : M} (hb : b * c ≤ a) (ha : a < b * (c + 1)) : a / b = c := by +lemma div_eq_of {b : V} (hb : b * c ≤ a) (ha : a < b * (c + 1)) : a / b = c := by have pos : 0 < b := pos_of_mul_pos_left (pos_of_gt ha) (by simp) exact (div_exists_unique_pos a pos).unique ⟨mul_div_le_pos a pos, lt_mul_div_succ a pos⟩ ⟨hb, ha⟩ -lemma div_mul_add (a b : M) {r} (hr : r < b) : (a * b + r) / b = a := +lemma div_mul_add (a b : V) {r} (hr : r < b) : (a * b + r) / b = a := div_eq_of (by simp [mul_comm]) (by simp [mul_comm b a, mul_add, hr]) -lemma div_mul_add' (a b : M) {r} (hr : r < b) : (b * a + r) / b = a := by simpa [mul_comm] using div_mul_add a b hr +lemma div_mul_add' (a b : V) {r} (hr : r < b) : (b * a + r) / b = a := by simpa [mul_comm] using div_mul_add a b hr -@[simp] lemma zero_div (a : M) : 0 / a = 0 := by +@[simp] lemma zero_div (a : V) : 0 / a = 0 := by rcases zero_le a with (rfl | pos) · simp · exact div_eq_of (by simp) (by simpa) -lemma div_mul (a b c : M) : a / (b * c) = a / b / c := by +lemma div_mul (a b c : V) : a / (b * c) = a / b / c := by rcases zero_le b with (rfl | hb) · simp rcases zero_le c with (rfl | hc) @@ -171,87 +171,87 @@ lemma div_mul (a b c : M) : a / (b * c) = a / b / c := by a < b * (a / b + 1) := lt_mul_div_succ a hb _ ≤ b * c * (a / b / c + 1) := by simp [mul_assoc]; exact mul_le_mul_left (lt_iff_succ_le.mp <| lt_mul_div_succ (a / b) hc)) -@[simp] lemma mul_div_le (a b : M) : b * (a / b) ≤ a := by +@[simp] lemma mul_div_le (a b : V) : b * (a / b) ≤ a := by have : 0 ≤ b := by exact zero_le b rcases this with (rfl | pos) <;> simp [*] rcases eq_mul_div_add_of_pos a pos with ⟨v, _, e⟩ simpa [← e] using show b * (a / b) ≤ b * (a / b) + v from le_self_add -@[simp] lemma div_le (a b : M) : a / b ≤ a := by +@[simp] lemma div_le (a b : V) : a / b ≤ a := by have : 0 ≤ b := zero_le b rcases this with (rfl | pos) <;> simp [*] have : 1 * (a / b) ≤ b * (a / b) := mul_le_mul_of_nonneg_right (le_iff_lt_succ.mpr (by simp[pos])) (by simp) simpa using le_trans this (mul_div_le a b) -instance div_polybounded : Bounded₂ ℒₒᵣ ((· / ·) : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance div_polybounded : Bounded₂ ((· / ·) : V → V → V) := ⟨#0, λ _ ↦ by simp⟩ -instance div_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ ((· / ·) : M → M → M) := div_defined.to_definable _ +instance div_definable : 𝚺₀-Function₂ ((· / ·) : V → V → V) := div_defined.to_definable _ -@[simp] lemma div_mul_le (a b : M) : a / b * b ≤ a := by rw [mul_comm]; exact mul_div_le _ _ +@[simp] lemma div_mul_le (a b : V) : a / b * b ≤ a := by rw [mul_comm]; exact mul_div_le _ _ -lemma lt_mul_div (a : M) {b} (pos : 0 < b) : a < b * (a / b + 1) := by +lemma lt_mul_div (a : V) {b} (pos : 0 < b) : a < b * (a / b + 1) := by rcases eq_mul_div_add_of_pos a pos with ⟨v, hv, e⟩ calc a = b * (a / b) + v := e _ < b * (a / b + 1) := by simp [mul_add, hv] -@[simp] lemma div_one (a : M) : a / 1 = a := +@[simp] lemma div_one (a : V) : a / 1 = a := le_antisymm (by simp) (le_iff_lt_succ.mpr $ by simpa using lt_mul_div a one_pos) -lemma div_add_mul_self (a c : M) {b} (pos : 0 < b) : (a + c * b) / b = a / b + c := by +lemma div_add_mul_self (a c : V) {b} (pos : 0 < b) : (a + c * b) / b = a / b + c := by rcases div_spec_of_pos' a pos with ⟨r, hr, ex⟩ simpa [add_mul, add_right_comm, ← ex] using div_mul_add (a / b + c) _ hr -lemma div_add_mul_self' (a c : M) {b} (pos : 0 < b) : (a + b * c) / b = a / b + c := by +lemma div_add_mul_self' (a c : V) {b} (pos : 0 < b) : (a + b * c) / b = a / b + c := by simpa [mul_comm] using div_add_mul_self a c pos -lemma div_mul_add_self (a c : M) {b} (pos : 0 < b) : (a * b + c) / b = a + c / b := by +lemma div_mul_add_self (a c : V) {b} (pos : 0 < b) : (a * b + c) / b = a + c / b := by simp [div_add_mul_self, pos, add_comm] -lemma div_mul_add_self' (a c : M) {b} (pos : 0 < b) : (b * a + c) / b = a + c / b := by +lemma div_mul_add_self' (a c : V) {b} (pos : 0 < b) : (b * a + c) / b = a + c / b := by simp [mul_comm b a, div_mul_add_self, pos] -@[simp] lemma div_mul_left (a : M) {b} (pos : 0 < b) : (a * b) / b = a := by +@[simp] lemma div_mul_left (a : V) {b} (pos : 0 < b) : (a * b) / b = a := by simpa using div_mul_add a _ pos -@[simp] lemma div_mul_right (a : M) {b} (pos : 0 < b) : (b * a) / b = a := by +@[simp] lemma div_mul_right (a : V) {b} (pos : 0 < b) : (b * a) / b = a := by simpa [mul_comm] using div_mul_add a _ pos -@[simp] lemma div_eq_zero_of_lt (b : M) {a} (h : a < b) : a / b = 0 := by +@[simp] lemma div_eq_zero_of_lt (b : V) {a} (h : a < b) : a / b = 0 := by simpa using div_mul_add 0 b h -@[simp] lemma div_sq (a : M) : a^2 / a = a := by +@[simp] lemma div_sq (a : V) : a^2 / a = a := by rcases zero_le a with (rfl | pos) · simp · simp [sq, pos] -@[simp] lemma div_self {a : M} (hx : 0 < a) : a / a = 1 := by +@[simp] lemma div_self {a : V} (hx : 0 < a) : a / a = 1 := by simpa using div_mul_left 1 hx -@[simp] lemma div_mul' (a : M) {b} (pos : 0 < b) : (b * a) / b = a := by simp [mul_comm, pos] +@[simp] lemma div_mul' (a : V) {b} (pos : 0 < b) : (b * a) / b = a := by simp [mul_comm, pos] -@[simp] lemma div_add_self_left {a} (pos : 0 < a) (b : M) : (a + b) / a = 1 + b / a := by +@[simp] lemma div_add_self_left {a} (pos : 0 < a) (b : V) : (a + b) / a = 1 + b / a := by simpa using div_mul_add_self 1 b pos -@[simp] lemma div_add_self_right (a : M) {b} (pos : 0 < b) : (a + b) / b = a / b + 1 := by +@[simp] lemma div_add_self_right (a : V) {b} (pos : 0 < b) : (a + b) / b = a / b + 1 := by simpa using div_add_mul_self a 1 pos -lemma mul_div_self_of_dvd {a b : M} : a * (b / a) = b ↔ a ∣ b := by +lemma mul_div_self_of_dvd {a b : V} : a * (b / a) = b ↔ a ∣ b := by rcases zero_le a with (rfl | pos) · simp[eq_comm] · constructor · intro e; rw [←e]; simp · rintro ⟨r, rfl⟩; simp [pos] -lemma div_lt_of_pos_of_one_lt {a b : M} (ha : 0 < a) (hb : 1 < b) : a / b < a := by +lemma div_lt_of_pos_of_one_lt {a b : V} (ha : 0 < a) (hb : 1 < b) : a / b < a := by rcases zero_le (a / b) with (e | lt) · simp [←e, ha] · exact lt_of_lt_of_le (lt_mul_of_one_lt_left lt hb) (mul_div_le a b) -lemma le_two_mul_div_two_add_one (a : M) : a ≤ 2 * (a / 2) + 1 := by +lemma le_two_mul_div_two_add_one (a : V) : a ≤ 2 * (a / 2) + 1 := by have : a < 2 * (a / 2 + 1) := lt_mul_div_succ a (show 0 < 2 from by simp) exact le_iff_lt_succ.mpr (by simpa [add_assoc, one_add_one_eq_two, mul_add] using this) -lemma div_monotone {a b : M} (h : a ≤ b) (c : M) : a / c ≤ b / c := by +lemma div_monotone {a b : V} (h : a ≤ b) (c : V) : a / c ≤ b / c := by rcases zero_le c with (rfl | pos) · simp by_contra A @@ -263,7 +263,7 @@ lemma div_monotone {a b : M} (h : a ≤ b) (c : M) : a / c ≤ b / c := by _ ≤ a := mul_div_le a c simp_all -lemma div_lt_of_lt_mul {a b c : M} (h : a < b * c) : a / c < b := by +lemma div_lt_of_lt_mul {a b c : V} (h : a < b * c) : a / c < b := by by_contra hb simp at hb have : a < a := calc @@ -272,133 +272,133 @@ lemma div_lt_of_lt_mul {a b c : M} (h : a < b * c) : a / c < b := by _ ≤ a := by simp simp_all -lemma div_cancel_left {c} (pos : 0 < c) (a b : M) : (c * a) / (c * b) = a / b := by simp [div_mul, pos] +lemma div_cancel_left {c} (pos : 0 < c) (a b : V) : (c * a) / (c * b) = a / b := by simp [div_mul, pos] -lemma div_cancel_right {c} (pos : 0 < c) (a b : M) : (a * c) / (b * c) = a / b := by simp [mul_comm _ c, div_cancel_left pos] +lemma div_cancel_right {c} (pos : 0 < c) (a b : V) : (a * c) / (b * c) = a / b := by simp [mul_comm _ c, div_cancel_left pos] -@[simp] lemma two_mul_add_one_div_two (a : M) : (2 * a + 1) / 2 = a := by simp [div_mul_add_self', one_lt_two] +@[simp] lemma two_mul_add_one_div_two (a : V) : (2 * a + 1) / 2 = a := by simp [div_mul_add_self', one_lt_two] end div section mod -def rem (a b : M) : M := a - b * (a / b) +def rem (a b : V) : V := a - b * (a / b) -scoped instance : Mod M := ⟨rem⟩ +scoped instance : Mod V := ⟨rem⟩ -lemma mod_def (a b : M) : a % b = a - b * (a / b) := rfl +lemma mod_def (a b : V) : a % b = a - b * (a / b) := rfl -def _root_.LO.FirstOrder.Arith.remDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.remDef : 𝚺₀.Semisentence 3 := .mkSigma “c a b | ∃ d <⁺ a, !divDef.val d a b ∧ !subDef.val c a (b * d)” (by simp) -lemma rem_graph (a b c : M) : a = b % c ↔ ∃ x ≤ b, (x = b / c ∧ a = b - c * x) := by +lemma rem_graph (a b c : V) : a = b % c ↔ ∃ x ≤ b, (x = b / c ∧ a = b - c * x) := by simp [mod_def]; constructor · rintro rfl; exact ⟨b / c, by simp, rfl, by rfl⟩ · rintro ⟨_, _, rfl, rfl⟩; simp -lemma rem_defined : 𝚺₀-Function₂ ((· % ·) : M → M → M) via remDef := by +lemma rem_defined : 𝚺₀-Function₂ ((· % ·) : V → V → V) via remDef := by intro v; simp [Matrix.vecHead, Matrix.vecTail, remDef, rem_graph, Semiformula.eval_substs, le_iff_lt_succ] @[simp] lemma rem_defined_iff (v) : - Semiformula.Evalbm M v remDef.val ↔ v 0 = v 1 % v 2 := rem_defined.df.iff v + Semiformula.Evalbm V v remDef.val ↔ v 0 = v 1 % v 2 := rem_defined.df.iff v -instance rem_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ ((· % ·) : M → M → M) := rem_defined.to_definable _ +instance rem_definable : 𝚺₀-Function₂ ((· % ·) : V → V → V) := rem_defined.to_definable _ -lemma div_add_mod (a b : M) : b * (a / b) + (a % b) = a := +lemma div_add_mod (a b : V) : b * (a / b) + (a % b) = a := add_tsub_self_of_le (mul_div_le a b) -@[simp] lemma mod_zero (a : M) : a % 0 = a := by simp [mod_def] +@[simp] lemma mod_zero (a : V) : a % 0 = a := by simp [mod_def] -@[simp] lemma zero_mod (a : M) : 0 % a = 0 := by simp [mod_def] +@[simp] lemma zero_mod (a : V) : 0 % a = 0 := by simp [mod_def] -@[simp] lemma mod_self (a : M) : a % a = 0 := by +@[simp] lemma mod_self (a : V) : a % a = 0 := by rcases zero_le a with (rfl | h) · simp · simp [mod_def, h] -lemma mod_mul_add_of_lt (a b : M) {r} (hr : r < b) : (a * b + r) % b = r := by +lemma mod_mul_add_of_lt (a b : V) {r} (hr : r < b) : (a * b + r) % b = r := by simp [mod_def, div_mul_add a b hr, mul_comm] -@[simp] lemma mod_mul_add (a c : M) (pos : 0 < b) : (a * b + c) % b = c % b := by +@[simp] lemma mod_mul_add (a c : V) (pos : 0 < b) : (a * b + c) % b = c % b := by simp [mod_def, div_mul_add_self, pos, mul_add, ←sub_sub, show b * a = a * b from mul_comm _ _] -@[simp] lemma mod_add_mul (a b : M) (pos : 0 < c) : (a + b * c) % c = a % c := by +@[simp] lemma mod_add_mul (a b : V) (pos : 0 < c) : (a + b * c) % c = a % c := by simp [add_comm a (b * c), pos] -@[simp] lemma mod_add_mul' (a b : M) (pos : 0 < c) : (a + c * b) % c = a % c := by +@[simp] lemma mod_add_mul' (a b : V) (pos : 0 < c) : (a + c * b) % c = a % c := by simp [mul_comm c b, pos] -@[simp] lemma mod_mul_add' (a c : M) (pos : 0 < b) : (b * a + c) % b = c % b := by +@[simp] lemma mod_mul_add' (a c : V) (pos : 0 < b) : (b * a + c) % b = c % b := by simp [mul_comm b a, pos] -@[simp] lemma mod_mul_self_left (a b : M) : (a * b) % b = 0 := by +@[simp] lemma mod_mul_self_left (a b : V) : (a * b) % b = 0 := by rcases zero_le b with (rfl | h) · simp · simpa using mod_mul_add_of_lt a b h -@[simp] lemma mod_mul_self_right (a b : M) : (b * a) % b = 0 := by +@[simp] lemma mod_mul_self_right (a b : V) : (b * a) % b = 0 := by simp [mul_comm] -@[simp] lemma mod_eq_self_of_lt {a b : M} (h : a < b) : a % b = a := by +@[simp] lemma mod_eq_self_of_lt {a b : V} (h : a < b) : a % b = a := by simpa using mod_mul_add_of_lt 0 b h -@[simp] lemma mod_lt (a : M) {b} (pos : 0 < b) : a % b < b := by +@[simp] lemma mod_lt (a : V) {b} (pos : 0 < b) : a % b < b := by rcases div_spec_of_pos' a pos with ⟨r, hr, ha⟩ have : ((a / b) * b + r) % b = r := mod_mul_add_of_lt _ _ hr have : a % b = r := by simpa [←ha] using this simp [this, hr] -@[simp] lemma mod_le (a b : M) : a % b ≤ a := by +@[simp] lemma mod_le (a b : V) : a % b ≤ a := by simp [mod_def] -instance mod_polybounded : Bounded₂ ℒₒᵣ ((· % ·) : M → M → M) := ⟨#0, by intro v; simp⟩ +instance mod_polybounded : Bounded₂ ((· % ·) : V → V → V) := ⟨#0, by intro v; simp⟩ -lemma mod_eq_zero_iff_dvd {a b : M} : b % a = 0 ↔ a ∣ b := by +lemma mod_eq_zero_iff_dvd {a b : V} : b % a = 0 ↔ a ∣ b := by simp [mod_def] constructor · intro H; exact mul_div_self_of_dvd.mp (le_antisymm (mul_div_le b a) H) · intro H; simp [mul_div_self_of_dvd.mpr H] -@[simp] lemma mod_add_remove_right {a b : M} (pos : 0 < b) : (a + b) % b = a % b := by +@[simp] lemma mod_add_remove_right {a b : V} (pos : 0 < b) : (a + b) % b = a % b := by simpa using mod_add_mul a 1 pos -lemma mod_add_remove_right_of_dvd {a b m : M} (h : m ∣ b) (pos : 0 < m) : (a + b) % m = a % m := by +lemma mod_add_remove_right_of_dvd {a b m : V} (h : m ∣ b) (pos : 0 < m) : (a + b) % m = a % m := by rcases h with ⟨b, rfl⟩; simp [pos] -@[simp] lemma mod_add_remove_left {a b : M} (pos : 0 < a) : (a + b) % a = b % a := by +@[simp] lemma mod_add_remove_left {a b : V} (pos : 0 < a) : (a + b) % a = b % a := by simpa using mod_mul_add 1 b pos -lemma mod_add_remove_left_of_dvd {a b m : M} (h : m ∣ a) (pos : 0 < m) : (a + b) % m = b % m := by +lemma mod_add_remove_left_of_dvd {a b m : V} (h : m ∣ a) (pos : 0 < m) : (a + b) % m = b % m := by rcases h with ⟨b, rfl⟩; simp [pos] -lemma mod_add {a b m : M} (pos : 0 < m) : (a + b) % m = (a % m + b % m) % m := calc +lemma mod_add {a b m : V} (pos : 0 < m) : (a + b) % m = (a % m + b % m) % m := calc (a + b) % m = ((m * (a / m) + a % m) + (m * (b / m) + b % m)) % m := by simp [div_add_mod] _ = (m * (a / m) + m * (b / m) + (a % m) + (b % m)) % m := by simp [←add_assoc, add_right_comm] _ = (a % m + b % m) % m := by simp [add_assoc, pos] -lemma mod_mul {a b m : M} (pos : 0 < m) : (a * b) % m = ((a % m) * (b % m)) % m := calc +lemma mod_mul {a b m : V} (pos : 0 < m) : (a * b) % m = ((a % m) * (b % m)) % m := calc (a * b) % m = ((m * (a / m) + (a % m)) * (m * (b / m) + b % m)) % m := by simp [div_add_mod] _ = ((a % m) * (b % m)) % m := by simp [add_mul, mul_add, pos, mul_left_comm _ m, add_assoc, mul_assoc] -@[simp] lemma mod_div (a b : M) : a % b / b = 0 := by +@[simp] lemma mod_div (a b : V) : a % b / b = 0 := by rcases zero_le b with (rfl | pos) · simp · exact div_eq_zero_of_lt b (by simp [pos]) -@[simp] lemma mod_one (a : M) : a % 1 = 0 := lt_one_iff_eq_zero.mp <| mod_lt a (by simp) +@[simp] lemma mod_one (a : V) : a % 1 = 0 := lt_one_iff_eq_zero.mp <| mod_lt a (by simp) -lemma mod_two (a : M) : a % 2 = 0 ∨ a % 2 = 1 := +lemma mod_two (a : V) : a % 2 = 0 ∨ a % 2 = 1 := le_one_iff_eq_zero_or_one.mp <| lt_two_iff_le_one.mp <| mod_lt a (b := 2) (by simp) -@[simp] lemma mod_two_not_zero_iff {a : M} : ¬a % 2 = 0 ↔ a % 2 = 1 := by +@[simp] lemma mod_two_not_zero_iff {a : V} : ¬a % 2 = 0 ↔ a % 2 = 1 := by rcases mod_two a with (h | h) <;> simp [*] -@[simp] lemma mod_two_not_one_iff {a : M} : ¬a % 2 = 1 ↔ a % 2 = 0 := by +@[simp] lemma mod_two_not_one_iff {a : V} : ¬a % 2 = 1 ↔ a % 2 = 0 := by rcases mod_two a with (h | h) <;> simp [*] end mod -lemma two_dvd_mul {a b : M} : 2 ∣ a * b → 2 ∣ a ∨ 2 ∣ b := by +lemma two_dvd_mul {a b : V} : 2 ∣ a * b → 2 ∣ a ∨ 2 ∣ b := by intro H; by_contra A simp [not_or] at A have ha : a % 2 = 1 := by @@ -412,20 +412,20 @@ lemma two_dvd_mul {a b : M} : 2 ∣ a * b → 2 ∣ a ∨ 2 ∣ b := by have : ¬2 ∣ a * b := by simp [←mod_eq_zero_iff_dvd, this] contradiction -lemma even_or_odd (a : M) : ∃ x, a = 2 * x ∨ a = 2 * x + 1 := +lemma even_or_odd (a : V) : ∃ x, a = 2 * x ∨ a = 2 * x + 1 := ⟨a / 2, by have : 2 * (a / 2) + (a % 2) = a := div_add_mod a 2 rcases mod_two a with (e | e) <;> { simp[e] at this; simp [this] }⟩ -lemma even_or_odd' (a : M) : a = 2 * (a / 2) ∨ a = 2 * (a / 2) + 1 := by +lemma even_or_odd' (a : V) : a = 2 * (a / 2) ∨ a = 2 * (a / 2) + 1 := by have : 2 * (a / 2) + (a % 2) = a := div_add_mod a 2 rcases mod_two a with (e | e) <;> simp [e] at this <;> simp [*] -lemma two_prime : Prime (2 : M) := ⟨by simp, by simp, by intro a b h; exact two_dvd_mul h⟩ +lemma two_prime : Prime (2 : V) := ⟨by simp, by simp, by intro a b h; exact two_dvd_mul h⟩ section sqrt -lemma sqrt_exists_unique (a : M) : ∃! x, x * x ≤ a ∧ a < (x + 1) * (x + 1) := by +lemma sqrt_exists_unique (a : V) : ∃! x, x * x ≤ a ∧ a < (x + 1) * (x + 1) := by have : ∃ x, x * x ≤ a ∧ a < (x + 1) * (x + 1) := by have : a < (a + 1) * (a + 1) → ∃ x, x * x ≤ a ∧ a < (x + 1) * (x + 1) := by simpa using open_leastNumber (P := λ x ↦ x * x ≤ a) ⟨“x | x * x ≤ &a”, by simp, by simp⟩ @@ -446,66 +446,66 @@ lemma sqrt_exists_unique (a : M) : ∃! x, x * x ≤ a ∧ a < (x + 1) * (x + 1) _ ≤ a := hy.1 simp at this) -def sqrt (a : M) : M := Classical.choose! (sqrt_exists_unique a) +def sqrt (a : V) : V := Classical.choose! (sqrt_exists_unique a) prefix:75 "√" => sqrt -@[simp] lemma sqrt_spec_le (a : M) : √a * √a ≤ a := (Classical.choose!_spec (sqrt_exists_unique a)).1 +@[simp] lemma sqrt_spec_le (a : V) : √a * √a ≤ a := (Classical.choose!_spec (sqrt_exists_unique a)).1 -@[simp] lemma sqrt_spec_lt (a : M) : a < (√a + 1) * (√a + 1) := (Classical.choose!_spec (sqrt_exists_unique a)).2 +@[simp] lemma sqrt_spec_lt (a : V) : a < (√a + 1) * (√a + 1) := (Classical.choose!_spec (sqrt_exists_unique a)).2 -lemma sqrt_graph {a b : M} : b = √a ↔ b * b ≤ a ∧ a < (b + 1) * (b + 1) := Classical.choose!_eq_iff _ +lemma sqrt_graph {a b : V} : b = √a ↔ b * b ≤ a ∧ a < (b + 1) * (b + 1) := Classical.choose!_eq_iff _ -def _root_.LO.FirstOrder.Arith.sqrtDef : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.sqrtDef : 𝚺₀.Semisentence 2 := .mkSigma “b a | b * b ≤ a ∧ a < (b + 1) * (b + 1)” (by simp[Hierarchy.pi_zero_iff_sigma_zero]) -lemma sqrt_defined : 𝚺₀-Function₁ (λ a : M ↦ √a) via sqrtDef := by +lemma sqrt_defined : 𝚺₀-Function₁ (λ a : V ↦ √a) via sqrtDef := by intro v; simp[sqrt_graph, sqrtDef, Matrix.vecHead, Matrix.vecTail] @[simp] lemma sqrt_defined_iff (v) : - Semiformula.Evalbm M v sqrtDef.val ↔ v 0 = √(v 1) := sqrt_defined.df.iff v + Semiformula.Evalbm V v sqrtDef.val ↔ v 0 = √(v 1) := sqrt_defined.df.iff v -instance sqrt_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ ((√·) : M → M) := Defined.to_definable _ sqrt_defined +instance sqrt_definable : 𝚺₀-Function₁ ((√·) : V → V) := sqrt_defined.to_definable -lemma eq_sqrt (x a : M) : x * x ≤ a ∧ a < (x + 1) * (x + 1) → x = √a := Classical.choose_uniq (sqrt_exists_unique a) +lemma eq_sqrt (x a : V) : x * x ≤ a ∧ a < (x + 1) * (x + 1) → x = √a := Classical.choose_uniq (sqrt_exists_unique a) -lemma sqrt_eq_of_le_of_lt {x a : M} (le : x * x ≤ a) (lt : a < (x + 1) * (x + 1)) : √a = x := +lemma sqrt_eq_of_le_of_lt {x a : V} (le : x * x ≤ a) (lt : a < (x + 1) * (x + 1)) : √a = x := Eq.symm <| eq_sqrt x a ⟨le, lt⟩ -lemma sqrt_eq_of_le_of_le {x a : M} (le : x * x ≤ a) (h : a ≤ x * x + 2 * x) : √a = x := +lemma sqrt_eq_of_le_of_le {x a : V} (le : x * x ≤ a) (h : a ≤ x * x + 2 * x) : √a = x := sqrt_eq_of_le_of_lt le (by simp [add_mul_self_eq]; exact le_iff_lt_succ.mp h) -@[simp] lemma sq_sqrt_le (a : M) : (√a) ^ 2 ≤ a := by simp [sq] +@[simp] lemma sq_sqrt_le (a : V) : (√a) ^ 2 ≤ a := by simp [sq] -@[simp] lemma sqrt_lt_sq (a : M) : a < (√a + 1) ^ 2 := by simp [sq] +@[simp] lemma sqrt_lt_sq (a : V) : a < (√a + 1) ^ 2 := by simp [sq] -@[simp] lemma sqrt_mul_self (a : M) : √(a * a) = a := +@[simp] lemma sqrt_mul_self (a : V) : √(a * a) = a := Eq.symm <| eq_sqrt a (a * a) (by simp; exact mul_self_lt_mul_self (by simp) (by simp)) -@[simp] lemma sqrt_sq (a : M) : √(a^2) = a := by simp [sq] +@[simp] lemma sqrt_sq (a : V) : √(a^2) = a := by simp [sq] -@[simp] lemma sqrt_zero : √(0 : M) = 0 := by simpa using sqrt_mul_self (0 : M) +@[simp] lemma sqrt_zero : √(0 : V) = 0 := by simpa using sqrt_mul_self (0 : V) -@[simp] lemma sqrt_one : √(1 : M) = 1 := by simpa using sqrt_mul_self (1 : M) +@[simp] lemma sqrt_one : √(1 : V) = 1 := by simpa using sqrt_mul_self (1 : V) -lemma sqrt_two : √(2 : M) = 1 := +lemma sqrt_two : √(2 : V) = 1 := Eq.symm <| eq_sqrt 1 2 (by simp [one_le_two, one_add_one_eq_two, one_lt_two]) -lemma sqrt_three : √(3 : M) = 1 := +lemma sqrt_three : √(3 : V) = 1 := Eq.symm <| eq_sqrt 1 3 <| by simp [one_add_one_eq_two, two_mul_two_eq_four, ←three_add_one_eq_four] -@[simp] lemma sqrt_four : √(4 : M) = 2 := by +@[simp] lemma sqrt_four : √(4 : V) = 2 := by simp [←two_mul_two_eq_four] -@[simp] lemma two_ne_square (a : M) : 2 ≠ a^2 := by +@[simp] lemma two_ne_square (a : V) : 2 ≠ a^2 := by intro h rcases show a = √2 from by rw [h]; simp with rfl simp [sqrt_two] at h -@[simp] lemma sqrt_le_add (a : M) : a ≤ √a * √a + 2 * √a := +@[simp] lemma sqrt_le_add (a : V) : a ≤ √a * √a + 2 * √a := le_iff_lt_succ.mpr (by have := sqrt_spec_lt a; rw [add_mul_self_eq] at this; simpa using this) -@[simp] lemma sqrt_le_self (a : M) : √a ≤ a := by +@[simp] lemma sqrt_le_self (a : V) : √a ≤ a := by by_contra A have : a < a := calc a ≤ a^2 := le_sq a @@ -513,15 +513,15 @@ lemma sqrt_three : √(3 : M) = 1 := _ ≤ a := sq_sqrt_le a simp_all -instance : Bounded₁ ℒₒᵣ ((√·) : M → M) := ⟨#0, by intro v; simp⟩ +instance : Bounded₁ ((√·) : V → V) := ⟨#0, by intro v; simp⟩ -lemma sqrt_lt_self_of_one_lt {a : M} (h : 1 < a) : √a < a := by +lemma sqrt_lt_self_of_one_lt {a : V} (h : 1 < a) : √a < a := by by_contra A have : a * a ≤ √a * √a := mul_self_le_mul_self (by simp) (by simpa using A) have : a * a ≤ a := le_trans this (sqrt_spec_le a) exact not_lt.mpr this (lt_mul_self h) -lemma sqrt_le_of_le_sq {a b : M} : a ≤ b^2 → √a ≤ b := by +lemma sqrt_le_of_le_sq {a b : V} : a ≤ b^2 → √a ≤ b := by intro h; by_contra A have : a < a := calc a ≤ b^2 := h @@ -529,7 +529,7 @@ lemma sqrt_le_of_le_sq {a b : M} : a ≤ b^2 → √a ≤ b := by _ ≤ a := by simp simp_all -lemma sq_lt_of_lt_sqrt {a b : M} : a < √b → a^2 < b := by +lemma sq_lt_of_lt_sqrt {a b : V} : a < √b → a^2 < b := by intro h; by_contra A exact not_le.mpr h (sqrt_le_of_le_sq $ show b ≤ a^2 from by simpa using A) @@ -540,7 +540,7 @@ section pair open Classical -- https://github.com/leanprover-community/mathlib4/blob/b075cdd0e6ad8b5a3295e7484b2ae59e9b2ec2a7/Mathlib/Data/Nat/Pairing.lean#L37 -def pair (a b : M) : M := if a < b then b * b + a else a * a + a + b +def pair (a b : V) : V := if a < b then b * b + a else a * a + a + b --notation "⟪" a ", " b "⟫" => pair a b @@ -556,38 +556,38 @@ def pairUnexpander : Lean.PrettyPrinter.Unexpander | `($_ $term $term2) => `(⟪$term, $term2⟫) | _ => throw () -lemma pair_graph {a b c : M} : +lemma pair_graph {a b c : V} : c = ⟪a, b⟫ ↔ (a < b ∧ c = b * b + a) ∨ (b ≤ a ∧ c = a * a + a + b) := by simp [pair] by_cases h : a < b · simp [h, show ¬b ≤ a from by simpa using h] · simp [h, show b ≤ a from by simpa using h] -def _root_.LO.FirstOrder.Arith.pairDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.pairDef : 𝚺₀.Semisentence 3 := .mkSigma “c a b | (a < b ∧ c = b * b + a) ∨ (b ≤ a ∧ c = a * a + a + b)” (by simp) -lemma pair_defined : 𝚺₀-Function₂ (λ a b : M ↦ ⟪a, b⟫) via pairDef := by +lemma pair_defined : 𝚺₀-Function₂ (λ a b : V ↦ ⟪a, b⟫) via pairDef := by intro v; simp [pair_graph, pairDef] @[simp] lemma pair_defined_iff (v) : - Semiformula.Evalbm M v pairDef.val ↔ v 0 = ⟪v 1, v 2⟫ := pair_defined.df.iff v + Semiformula.Evalbm V v pairDef.val ↔ v 0 = ⟪v 1, v 2⟫ := pair_defined.df.iff v -instance pair_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ (pair : M → M → M) := Defined.to_definable _ pair_defined +instance pair_definable : 𝚺₀-Function₂ (pair : V → V → V) := pair_defined.to_definable -instance : Bounded₂ ℒₒᵣ (pair : M → M → M) := +instance : Bounded₂ (pair : V → V → V) := ⟨‘x y | (y * y + x) + (x * x + x + y)’, by intro v; simp [pair]; split_ifs <;> try simp [pair, *]⟩ -def unpair (a : M) : M × M := if a - √a * √a < √a then (a - √a * √a, √a) else (√a, a - √a * √a - √a) +def unpair (a : V) : V × V := if a - √a * √a < √a then (a - √a * √a, √a) else (√a, a - √a * √a - √a) -abbrev pi₁ (a : M) : M := (unpair a).1 +abbrev pi₁ (a : V) : V := (unpair a).1 -abbrev pi₂ (a : M) : M := (unpair a).2 +abbrev pi₂ (a : V) : V := (unpair a).2 prefix: 80 "π₁" => pi₁ prefix: 80 "π₂" => pi₂ -@[simp] lemma pair_unpair (a : M) : ⟪π₁ a, π₂ a⟫ = a := by +@[simp] lemma pair_unpair (a : V) : ⟪π₁ a, π₂ a⟫ = a := by simp [pi₁, pi₂, unpair] split_ifs with h · simp [pair, h] @@ -600,7 +600,7 @@ prefix: 80 "π₂" => pi₂ rw [add_tsub_self_of_le, add_tsub_self_of_le] <;> simp [this] _ = a := add_tsub_self_of_le (by simp) -@[simp] lemma unpair_pair (a b : M) : unpair ⟪a, b⟫ = (a, b) := by +@[simp] lemma unpair_pair (a b : V) : unpair ⟪a, b⟫ = (a, b) := by simp [pair]; split_ifs with h · have : √(b * b + a) = b := sqrt_eq_of_le_of_le (by simp) (by simp; exact le_trans (le_of_lt h) (by simp)) simp [unpair, this, show ¬b ≤ a from by simpa using h] @@ -608,60 +608,60 @@ prefix: 80 "π₂" => pi₂ sqrt_eq_of_le_of_le (by simp [add_assoc]) (by simp [add_assoc, two_mul, show b ≤ a from by simpa using h]) simp [unpair, this, add_assoc] -@[simp] lemma pi₁_pair (a b : M) : π₁ ⟪a, b⟫ = a := by simp [pi₁] +@[simp] lemma pi₁_pair (a b : V) : π₁ ⟪a, b⟫ = a := by simp [pi₁] -@[simp] lemma pi₂_pair (a b : M) : π₂ ⟪a, b⟫ = b := by simp [pi₂] +@[simp] lemma pi₂_pair (a b : V) : π₂ ⟪a, b⟫ = b := by simp [pi₂] -def pairEquiv : M × M ≃ M := ⟨Function.uncurry pair, unpair, fun ⟨a, b⟩ => unpair_pair a b, pair_unpair⟩ +def pairEquiv : V × V ≃ V := ⟨Function.uncurry pair, unpair, fun ⟨a, b⟩ => unpair_pair a b, pair_unpair⟩ -@[simp] lemma pi₁_le_self (a : M) : π₁ a ≤ a := by simp [pi₁, unpair]; split_ifs <;> simp +@[simp] lemma pi₁_le_self (a : V) : π₁ a ≤ a := by simp [pi₁, unpair]; split_ifs <;> simp -@[simp] lemma pi₂_le_self (a : M) : π₂ a ≤ a := by simp [pi₂, unpair]; split_ifs <;> simp [add_assoc] +@[simp] lemma pi₂_le_self (a : V) : π₂ a ≤ a := by simp [pi₂, unpair]; split_ifs <;> simp [add_assoc] -@[simp] lemma le_pair_left (a b : M) : a ≤ ⟪a, b⟫ := by simpa using pi₁_le_self ⟪a, b⟫ +@[simp] lemma le_pair_left (a b : V) : a ≤ ⟪a, b⟫ := by simpa using pi₁_le_self ⟪a, b⟫ -@[simp] lemma le_pair_right (a b : M) : b ≤ ⟪a, b⟫ := by simpa using pi₂_le_self ⟪a, b⟫ +@[simp] lemma le_pair_right (a b : V) : b ≤ ⟪a, b⟫ := by simpa using pi₂_le_self ⟪a, b⟫ -@[simp] lemma lt_pair_left_of_pos {a} (pos : 0 < a) (b : M) : a < ⟪a, b⟫ := by +@[simp] lemma lt_pair_left_of_pos {a} (pos : 0 < a) (b : V) : a < ⟪a, b⟫ := by simp [pair]; split_ifs · simp; exact pos_iff_ne_zero.mp <| pos_of_gt (by assumption) · calc a < a * a + a := lt_add_of_pos_left a (by simpa using (pos_iff_ne_zero.mp pos)) _ ≤ a * a + a + b := by simp -instance : Bounded₁ ℒₒᵣ (pi₁ : M → M) := ⟨#0, by intro v; simp⟩ +instance : Bounded₁ (pi₁ : V → V) := ⟨#0, by intro v; simp⟩ -instance : Bounded₁ ℒₒᵣ (pi₂ : M → M) := ⟨#0, by intro v; simp⟩ +instance : Bounded₁ (pi₂ : V → V) := ⟨#0, by intro v; simp⟩ -def _root_.LO.FirstOrder.Arith.pi₁Def : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.pi₁Def : 𝚺₀.Semisentence 2 := .mkSigma “x p | ∃ y <⁺ p, !pairDef p x y” (by simp) -def _root_.LO.FirstOrder.Arith.pi₂Def : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.pi₂Def : 𝚺₀.Semisentence 2 := .mkSigma “y p | ∃ x <⁺ p, !pairDef p x y” (by simp) -lemma pi₁_defined : 𝚺₀-Function₁ (pi₁ : M → M) via pi₁Def := by +lemma pi₁_defined : 𝚺₀-Function₁ (pi₁ : V → V) via pi₁Def := by intro v; simp [pi₁Def] constructor · intro h; exact ⟨π₂ v 1, by simp [←le_iff_lt_succ], by simp [h]⟩ · rintro ⟨a, _, e⟩; simp [e] @[simp] lemma pi₁_defined_iff (v) : - Semiformula.Evalbm M v pi₁Def.val ↔ v 0 = π₁ (v 1) := pi₁_defined.df.iff v + Semiformula.Evalbm V v pi₁Def.val ↔ v 0 = π₁ (v 1) := pi₁_defined.df.iff v -instance pi₁_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (pi₁ : M → M) := Defined.to_definable₀ _ pi₁_defined +instance pi₁_definable : 𝚺₀-Function₁ (pi₁ : V → V) := pi₁_defined.to_definable₀ -lemma pi₂_defined : 𝚺₀-Function₁ (pi₂ : M → M) via pi₂Def := by +lemma pi₂_defined : 𝚺₀-Function₁ (pi₂ : V → V) via pi₂Def := by intro v; simp [pi₂Def] constructor · intro h; exact ⟨π₁ v 1, by simp [←le_iff_lt_succ], by simp [h]⟩ · rintro ⟨a, _, e⟩; simp [e] @[simp] lemma pi₂_defined_iff (v) : - Semiformula.Evalbm M v pi₂Def.val ↔ v 0 = π₂ (v 1) := pi₂_defined.df.iff v + Semiformula.Evalbm V v pi₂Def.val ↔ v 0 = π₂ (v 1) := pi₂_defined.df.iff v -instance pi₂_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (pi₂ : M → M) := Defined.to_definable₀ _ pi₂_defined +instance pi₂_definable : 𝚺₀-Function₁ (pi₂ : V → V) := pi₂_defined.to_definable₀ -lemma pair_lt_pair_left {a₁ a₂ : M} (h : a₁ < a₂) (b) : ⟪a₁, b⟫ < ⟪a₂, b⟫ := by +lemma pair_lt_pair_left {a₁ a₂ : V} (h : a₁ < a₂) (b) : ⟪a₁, b⟫ < ⟪a₂, b⟫ := by by_cases h₁ : a₁ < b <;> simp [pair, h₁] · by_cases h₂ : a₂ < b <;> simp [pair, h₂, h] calc @@ -671,12 +671,12 @@ lemma pair_lt_pair_left {a₁ a₂ : M} (h : a₁ < a₂) (b) : ⟪a₁, b⟫ < · simp[show ¬a₂ < b from by simp; exact le_trans (by simpa using h₁) (le_of_lt h)] apply _root_.add_lt_add (by simpa [←sq] using h) h -lemma pair_le_pair_left {a₁ a₂ : M} (h : a₁ ≤ a₂) (b) : ⟪a₁, b⟫ ≤ ⟪a₂, b⟫ := by +lemma pair_le_pair_left {a₁ a₂ : V} (h : a₁ ≤ a₂) (b) : ⟪a₁, b⟫ ≤ ⟪a₂, b⟫ := by rcases h with (rfl | lt) · simp · exact le_of_lt (pair_lt_pair_left lt b) -lemma pair_lt_pair_right (a : M) {b₁ b₂} (h : b₁ < b₂) : ⟪a, b₁⟫ < ⟪a, b₂⟫ := by +lemma pair_lt_pair_right (a : V) {b₁ b₂} (h : b₁ < b₂) : ⟪a, b₁⟫ < ⟪a, b₂⟫ := by by_cases h₁ : a < b₁ <;> simp [pair, h₁] · simpa [lt_trans h₁ h, ←sq] using h · by_cases h₂ : a < b₂ <;> simp [h₂, h] @@ -685,50 +685,50 @@ lemma pair_lt_pair_right (a : M) {b₁ b₂} (h : b₁ < b₂) : ⟪a, b₁⟫ < _ ≤ b₂ * b₂ + b₁ := by simpa [←sq, succ_le_iff_lt] using h₂ _ ≤ b₂ * b₂ + a := by simpa using h₁ -lemma pair_le_pair_right (a : M) {b₁ b₂} (h : b₁ ≤ b₂) : ⟪a, b₁⟫ ≤ ⟪a, b₂⟫ := by +lemma pair_le_pair_right (a : V) {b₁ b₂} (h : b₁ ≤ b₂) : ⟪a, b₁⟫ ≤ ⟪a, b₂⟫ := by rcases h with (rfl | lt) · simp · exact le_of_lt (pair_lt_pair_right a lt) -lemma pair_le_pair {a₁ a₂ b₁ b₂ : M} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : ⟪a₁, b₁⟫ ≤ ⟪a₂, b₂⟫ := +lemma pair_le_pair {a₁ a₂ b₁ b₂ : V} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : ⟪a₁, b₁⟫ ≤ ⟪a₂, b₂⟫ := calc ⟪a₁, b₁⟫ ≤ ⟪a₂, b₁⟫ := pair_le_pair_left ha b₁ _ ≤ ⟪a₂, b₂⟫ := pair_le_pair_right a₂ hb -lemma pair_lt_pair {a₁ a₂ b₁ b₂ : M} (ha : a₁ < a₂) (hb : b₁ < b₂) : ⟪a₁, b₁⟫ < ⟪a₂, b₂⟫ := +lemma pair_lt_pair {a₁ a₂ b₁ b₂ : V} (ha : a₁ < a₂) (hb : b₁ < b₂) : ⟪a₁, b₁⟫ < ⟪a₂, b₂⟫ := calc ⟪a₁, b₁⟫ < ⟪a₂, b₁⟫ := pair_lt_pair_left ha b₁ _ < ⟪a₂, b₂⟫ := pair_lt_pair_right a₂ hb -@[simp] lemma pair_polybound (a b : M) : ⟪a, b⟫ ≤ (a + b + 1)^2 := by +@[simp] lemma pair_polybound (a b : V) : ⟪a, b⟫ ≤ (a + b + 1)^2 := by by_cases h : a < b <;> simp [pair, h, sq, add_mul_self_eq, two_mul] · simp [←add_assoc, add_right_comm _ a]; simp [add_right_comm _ (b * b)] · simp [←add_assoc, add_right_comm _ b]; simp [add_right_comm _ a]; simp [add_assoc] -@[simp] lemma pair_ext_iff {a₁ a₂ b₁ b₂ : M} : ⟪a₁, b₁⟫ = ⟪a₂, b₂⟫ ↔ a₁ = a₂ ∧ b₁ = b₂ := +@[simp] lemma pair_ext_iff {a₁ a₂ b₁ b₂ : V} : ⟪a₁, b₁⟫ = ⟪a₂, b₂⟫ ↔ a₁ = a₂ ∧ b₁ = b₂ := ⟨fun e ↦ ⟨by simpa using congr_arg (π₁ ·) e, by simpa using congr_arg (π₂ ·) e⟩, by rintro ⟨rfl, rfl⟩; simp⟩ section -def _root_.LO.FirstOrder.Arith.pair₃Def : 𝚺₀-Semisentence 4 := +def _root_.LO.FirstOrder.Arith.pair₃Def : 𝚺₀.Semisentence 4 := .mkSigma “p a b c | ∃ bc <⁺ p, !pairDef p a bc ∧ !pairDef bc b c” (by simp) -def _root_.LO.FirstOrder.Arith.pair₄Def : 𝚺₀-Semisentence 5 := +def _root_.LO.FirstOrder.Arith.pair₄Def : 𝚺₀.Semisentence 5 := .mkSigma “p a b c d | ∃ bcd <⁺ p, ∃ cd <⁺ bcd, !pairDef p a bcd ∧ !pairDef bcd b cd ∧ !pairDef cd c d” (by simp) -def _root_.LO.FirstOrder.Arith.pair₅Def : 𝚺₀-Semisentence 6 := +def _root_.LO.FirstOrder.Arith.pair₅Def : 𝚺₀.Semisentence 6 := .mkSigma “p a b c d e | ∃ bcde <⁺ p, ∃ cde <⁺ bcde, ∃ de <⁺ cde, !pairDef p a bcde ∧ !pairDef bcde b cde ∧ !pairDef cde c de ∧ !pairDef de d e” (by simp) -def _root_.LO.FirstOrder.Arith.pair₆Def : 𝚺₀-Semisentence 7 := +def _root_.LO.FirstOrder.Arith.pair₆Def : 𝚺₀.Semisentence 7 := .mkSigma “p a b c d e f | ∃ bcdef <⁺ p, !pair₅Def bcdef b c d e f ∧ !pairDef p a bcdef” (by simp) -lemma pair₃_defined : 𝚺₀-Function₃ ((⟪·, ·, ·⟫) : M → M → M → M) via pair₃Def := by +lemma pair₃_defined : 𝚺₀-Function₃ ((⟪·, ·, ·⟫) : V → V → V → V) via pair₃Def := by intro v; simp [pair₃Def]; rintro h; simp [h] @[simp] lemma eval_pair₃Def (v) : - Semiformula.Evalbm M v pair₃Def.val ↔ v 0 = ⟪v 1, v 2, v 3⟫ := pair₃_defined.df.iff v + Semiformula.Evalbm V v pair₃Def.val ↔ v 0 = ⟪v 1, v 2, v 3⟫ := pair₃_defined.df.iff v -lemma pair₄_defined : 𝚺₀-Function₄ ((⟪·, ·, ·, ·⟫) : M → M → M → M → M) via pair₄Def := by +lemma pair₄_defined : 𝚺₀-Function₄ ((⟪·, ·, ·, ·⟫) : V → V → V → V → V) via pair₄Def := by intro v; simp [pair₄Def] constructor · intro h; simp only [Fin.isValue, h, pair_ext_iff, true_and] @@ -736,9 +736,9 @@ lemma pair₄_defined : 𝚺₀-Function₄ ((⟪·, ·, ·, ·⟫) : M → M · rintro ⟨_, _, _, _, h, rfl, rfl⟩; exact h @[simp] lemma eval_pair₄Def (v) : - Semiformula.Evalbm M v pair₄Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4⟫ := pair₄_defined.df.iff v + Semiformula.Evalbm V v pair₄Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4⟫ := pair₄_defined.df.iff v -lemma pair₅_defined : DefinedFunction (fun v : Fin 5 → M ↦ (⟪v 0, v 1, v 2, v 3, v 4⟫)) pair₅Def := by +lemma pair₅_defined : 𝚺₀.DefinedFunction (fun v : Fin 5 → V ↦ (⟪v 0, v 1, v 2, v 3, v 4⟫)) pair₅Def := by intro v; simp [pair₅Def] constructor · intro h; simp only [Fin.isValue, h, pair_ext_iff, true_and] @@ -746,9 +746,9 @@ lemma pair₅_defined : DefinedFunction (fun v : Fin 5 → M ↦ (⟪v 0, v 1, v · rintro ⟨_, _, _, _, _, _, h, rfl, rfl, rfl⟩; exact h @[simp] lemma eval_pair₅Def (v) : - Semiformula.Evalbm M v pair₅Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4, v 5⟫ := pair₅_defined.df.iff v + Semiformula.Evalbm V v pair₅Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4, v 5⟫ := pair₅_defined.df.iff v -lemma pair₆_defined : DefinedFunction (fun v : Fin 6 → M ↦ (⟪v 0, v 1, v 2, v 3, v 4, v 5⟫)) pair₆Def := by +lemma pair₆_defined : 𝚺₀.DefinedFunction (fun v : Fin 6 → V ↦ (⟪v 0, v 1, v 2, v 3, v 4, v 5⟫)) pair₆Def := by intro v; simp [pair₆Def] constructor · intro h; simp only [Fin.isValue, h, pair_ext_iff, true_and] @@ -756,35 +756,35 @@ lemma pair₆_defined : DefinedFunction (fun v : Fin 6 → M ↦ (⟪v 0, v 1, v · rintro ⟨_, _, rfl, h⟩; exact h @[simp] lemma eval_pair₆Def (v) : - Semiformula.Evalbm M v pair₆Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4, v 5, v 6⟫ := pair₆_defined.df.iff v + Semiformula.Evalbm V v pair₆Def.val ↔ v 0 = ⟪v 1, v 2, v 3, v 4, v 5, v 6⟫ := pair₆_defined.df.iff v end -def npair : {n : ℕ} → (v : Fin n → M) → M +def npair : {n : ℕ} → (v : Fin n → V) → V | 0, _ => 0 | _ + 1, v => ⟪v 0, npair (v ·.succ)⟫ -@[simp] lemma npair_zero (v : Fin 0 → M) : npair v = 0 := by simp [npair] +@[simp] lemma npair_zero (v : Fin 0 → V) : npair v = 0 := by simp [npair] -lemma npair_succ (x) (v : Fin n → M) : npair (x :> v) = ⟪x, npair v⟫ := by simp [npair] +lemma npair_succ (x) (v : Fin n → V) : npair (x :> v) = ⟪x, npair v⟫ := by simp [npair] -def unNpair : {n : ℕ} → Fin n → M → M +def unNpair : {n : ℕ} → Fin n → V → V | 0, i, _ => i.elim0 | _ + 1, i, x => Fin.cases (π₁ x) (fun i ↦ unNpair i (π₂ x)) i -@[simp] lemma unNpair_npair {n} (i : Fin n) (v : Fin n → M) : unNpair i (npair v) = v i := by +@[simp] lemma unNpair_npair {n} (i : Fin n) (v : Fin n → V) : unNpair i (npair v) = v i := by induction' n with n ih <;> simp [npair, unNpair, *] · exact i.elim0 · cases i using Fin.cases <;> simp section -def _root_.LO.FirstOrder.Arith.unNpairDef : {n : ℕ} → (i : Fin n) → 𝚺₀-Semisentence 2 +def _root_.LO.FirstOrder.Arith.unNpairDef : {n : ℕ} → (i : Fin n) → 𝚺₀.Semisentence 2 | 0, i => i.elim0 | n + 1, i => Fin.cases pi₁Def (fun i ↦ .mkSigma “z v | ∃ r <⁺ v, !pi₂Def r v ∧ !(unNpairDef i) z r” (by simp)) i -lemma unNpair_defined {n} (i : Fin n) : 𝚺₀-Function₁ (unNpair i : M → M) via unNpairDef i := by +lemma unNpair_defined {n} (i : Fin n) : 𝚺₀-Function₁ (unNpair i : V → V) via unNpairDef i := by induction' n with n ih · exact i.elim0 · intro v @@ -796,10 +796,10 @@ lemma unNpair_defined {n} (i : Fin n) : 𝚺₀-Function₁ (unNpair i : M → M · rintro ⟨x, _, rfl, h⟩; exact h @[simp] lemma eval_unNpairDef {n} (i : Fin n) (v) : - Semiformula.Evalbm M v (unNpairDef i).val ↔ v 0 = unNpair i (v 1) := (unNpair_defined i).df.iff v + Semiformula.Evalbm V v (unNpairDef i).val ↔ v 0 = unNpair i (v 1) := (unNpair_defined i).df.iff v -@[definability, simp] instance unNpair_definable {n} (i : Fin n) (Γ) : Γ-Function₁ (unNpair i : M → M) := - Defined.to_definable₀ _ (unNpair_defined i) +@[definability, simp] instance unNpair_definable {n} (i : Fin n) (Γ) : Γ-Function₁ (unNpair i : V → V) := + (unNpair_defined i).to_definable₀ end @@ -809,19 +809,17 @@ end IOpen section polynomial_induction -variable [M ⊧ₘ* 𝐈open] -variable {L : Language} [L.ORing] [Structure L M] [Structure.ORing L M] [Structure.Monotone L M] +variable [V ⊧ₘ* 𝐈open] @[elab_as_elim] -lemma hierarchy_polynomial_induction (Γ m) [M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy Γ m)] - {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma hierarchy_polynomial_induction (Γ m) [V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy Γ m)] + {P : V → Prop} (hP : Γ-[m]-Predicate P) (zero : P 0) (even : ∀ x > 0, P x → P (2 * x)) (odd : ∀ x, P x → P (2 * x + 1)) : ∀ x, P x := by intro x; induction x using order_induction_h · exact Γ · exact m · exact hP case inst => exact inferInstance - case inst => exact inferInstance case ind x IH => rcases zero_le x with (rfl | pos) · exact zero @@ -832,19 +830,19 @@ lemma hierarchy_polynomial_induction (Γ m) [M ⊧ₘ* Theory.indScheme L (Arith end polynomial_induction -@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_sigma₀ [M ⊧ₘ* 𝐈𝚺₀] {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₀ P) +@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_sigma₀ [V ⊧ₘ* 𝐈𝚺₀] {P : V → Prop} (hP : 𝚺₀-Predicate P) (zero : P 0) (even : ∀ x > 0, P x → P (2 * x)) (odd : ∀ x, P x → P (2 * x + 1)) : ∀ x, P x := hierarchy_polynomial_induction 𝚺 0 hP zero even odd -@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_sigma₁ [M ⊧ₘ* 𝐈𝚺₁] {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₁ P) +@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_sigma₁ [V ⊧ₘ* 𝐈𝚺₁] {P : V → Prop} (hP : 𝚺₁-Predicate P) (zero : P 0) (even : ∀ x > 0, P x → P (2 * x)) (odd : ∀ x, P x → P (2 * x + 1)) : ∀ x, P x := hierarchy_polynomial_induction 𝚺 1 hP zero even odd -@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_pi₁ [M ⊧ₘ* 𝐈𝚷₁] {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚷₁ P) +@[elab_as_elim] lemma hierarchy_polynomial_induction_oRing_pi₁ [V ⊧ₘ* 𝐈𝚷₁] {P : V → Prop} (hP : 𝚷₁-Predicate P) (zero : P 0) (even : ∀ x > 0, P x → P (2 * x)) (odd : ∀ x, P x → P (2 * x + 1)) : ∀ x, P x := hierarchy_polynomial_induction 𝚷 1 hP zero even odd -lemma nat_cast_pair (n m : ℕ) : (⟪n, m⟫ : ℕ) = ⟪(n : M), (m : M)⟫ := by simp [pair] +lemma nat_cast_pair (n m : ℕ) : (⟪n, m⟫ : ℕ) = ⟪(n : V), (m : V)⟫ := by simp [pair] lemma nat_pair_eq (m n : ℕ) : ⟪n, m⟫ = Nat.pair n m := by simp [Arith.pair, Nat.pair]; congr diff --git a/Arithmetization/Basic/Ind.lean b/Arithmetization/Basic/Ind.lean index 97dee44..1c37fd6 100644 --- a/Arithmetization/Basic/Ind.lean +++ b/Arithmetization/Basic/Ind.lean @@ -4,17 +4,17 @@ namespace LO.FirstOrder.Arith namespace Theory -variable {L : Language} [L.ORing] {C C' : Semiformula L ℕ 1 → Prop} +variable {C C' : Semiformula ℒₒᵣ ℕ 1 → Prop} -lemma mem_indScheme_of_mem {p : Semiformula L ℕ 1} (hp : C p) : - ∀ᶠ* succInd p ∈ indScheme L C := by +lemma mem_indScheme_of_mem {p : Semiformula ℒₒᵣ ℕ 1} (hp : C p) : + ∀ᶠ* succInd p ∈ indScheme ℒₒᵣ C := by simp [indScheme]; exact ⟨p, hp, rfl⟩ -lemma mem_iOpen_of_qfree {p : Semiformula L ℕ 1} (hp : p.Open) : - ∀ᶠ* succInd p ∈ indScheme L Semiformula.Open := by +lemma mem_iOpen_of_qfree {p : Semiformula ℒₒᵣ ℕ 1} (hp : p.Open) : + ∀ᶠ* succInd p ∈ indScheme ℒₒᵣ Semiformula.Open := by exact ⟨p, hp, rfl⟩ -lemma indScheme_subset (h : ∀ {p : Semiformula L ℕ 1}, C p → C' p) : indScheme L C ⊆ indScheme L C' := by +lemma indScheme_subset (h : ∀ {p : Semiformula ℒₒᵣ ℕ 1}, C p → C' p) : indScheme ℒₒᵣ C ⊆ indScheme ℒₒᵣ C' := by intro _; simp [indScheme]; rintro p hp rfl; exact ⟨p, h hp, rfl⟩ lemma iSigma_subset_mono {s₁ s₂} (h : s₁ ≤ s₂) : 𝐈𝚺 s₁ ⊆ 𝐈𝚺 s₂ := @@ -30,63 +30,61 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] section -variable [M ⊧ₘ* 𝐏𝐀⁻] {L : Language} [L.ORing] [Structure L M] [Structure.ORing L M] [Structure.Monotone L M] +variable [V ⊧ₘ* 𝐏𝐀⁻] section IndScheme -variable {C : Semiformula L ℕ 1 → Prop} [M ⊧ₘ* Theory.indScheme L C] +variable {C : Semiformula ℒₒᵣ ℕ 1 → Prop} [V ⊧ₘ* Theory.indScheme ℒₒᵣ C] -private lemma induction_eval {p : Semiformula L ℕ 1} (hp : C p) (v) : - Semiformula.Evalm M ![0] v p → - (∀ x, Semiformula.Evalm M ![x] v p → Semiformula.Evalm M ![x + 1] v p) → - ∀ x, Semiformula.Evalm M ![x] v p := by - have : M ⊧ₘ (∀ᶠ* succInd p) := - ModelsTheory.models (T := Theory.indScheme _ C) M (by simpa using Theory.mem_indScheme_of_mem hp) +private lemma induction_eval {p : Semiformula ℒₒᵣ ℕ 1} (hp : C p) (v) : + Semiformula.Evalm V ![0] v p → + (∀ x, Semiformula.Evalm V ![x] v p → Semiformula.Evalm V ![x + 1] v p) → + ∀ x, Semiformula.Evalm V ![x] v p := by + have : V ⊧ₘ (∀ᶠ* succInd p) := + ModelsTheory.models (T := Theory.indScheme _ C) V (by simpa using Theory.mem_indScheme_of_mem hp) simp [models_iff, succInd, Semiformula.eval_substs, Semiformula.eval_rew_q Rew.toS, Function.comp, Matrix.constant_eq_singleton] at this exact this v -variable (L) - @[elab_as_elim] -lemma induction {P : M → Prop} - (hP : ∃ e : ℕ → M, ∃ p : Semiformula L ℕ 1, C p ∧ ∀ x, P x ↔ Semiformula.Evalm M ![x] e p) : +lemma induction {P : V → Prop} + (hP : ∃ e : ℕ → V, ∃ p : Semiformula ℒₒᵣ ℕ 1, C p ∧ ∀ x, P x ↔ Semiformula.Evalm V ![x] e p) : P 0 → (∀ x, P x → P (x + 1)) → ∀ x, P x := by - rcases hP with ⟨e, p, Cp, hp⟩; simpa [←hp] using induction_eval (M := M) Cp e + rcases hP with ⟨e, p, Cp, hp⟩; simpa [←hp] using induction_eval (V := V) Cp e end IndScheme section neg -variable (Γ : Polarity) (m : ℕ) [M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy Γ m)] +variable (Γ : Polarity) (m : ℕ) [V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy Γ m)] @[elab_as_elim] -lemma induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma induction_h {P : V → Prop} (hP : Γ-[m].BoldfacePred P) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := - induction (P := P) (C := Hierarchy Γ m) (L := L) (by + induction (P := P) (C := Hierarchy Γ m) (by rcases hP with ⟨p, hp⟩ - haveI : Inhabited M := Classical.inhabited_of_nonempty' + haveI : Inhabited V := Classical.inhabited_of_nonempty' exact ⟨p.val.fvEnumInv, (Rew.rewriteMap p.val.fvEnum).hom p.val, by simp [hp], by intro x; simp [Semiformula.eval_rewriteMap] - have : (Semiformula.Evalm M ![x] fun x => p.val.fvEnumInv (p.val.fvEnum x)) p.val ↔ (Semiformula.Evalm M ![x] id) p.val := + have : (Semiformula.Evalm V ![x] fun x => p.val.fvEnumInv (p.val.fvEnum x)) p.val ↔ (Semiformula.Evalm V ![x] id) p.val := Semiformula.eval_iff_of_funEqOn _ (by intro x hx; simp [Semiformula.fvEnumInv_fvEnum hx]) simp [this, hp.df.iff]⟩) zero succ @[elab_as_elim] -lemma order_induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma order_induction_h {P : V → Prop} (hP : Γ-[m].BoldfacePred P) (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := by suffices ∀ x, ∀ y < x, P y by intro x; exact this (x + 1) x (by simp only [lt_add_iff_pos_right, lt_one_iff_eq_zero]) intro x; induction x using induction_h · exact Γ · exact m - · suffices DefinablePred L (Γ, m) fun x => ∀ y < x, P y by exact this - exact Definable.ball_lt₀ (L := L) (by simp) (hP.retraction ![0]) + · suffices Γ-[m].BoldfacePred fun x => ∀ y < x, P y by exact this + exact HierarchySymbol.Boldface.ball_blt (by simp) (hP.retraction ![0]) case zero => simp case succ x IH => intro y hxy @@ -94,9 +92,8 @@ lemma order_induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) · exact IH y lt · exact ind y IH case inst => exact inferInstance - case inst => exact inferInstance -private lemma neg_induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +private lemma neg_induction_h {P : V → Prop} (hP : Γ-[m].BoldfacePred P) (nzero : ¬P 0) (nsucc : ∀ x, ¬P x → ¬P (x + 1)) : ∀ x, ¬P x := by by_contra A have : ∃ x, P x := by simpa using A @@ -105,8 +102,10 @@ private lemma neg_induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) intro x; induction x using induction_h · exact Γ · exact m - · suffices DefinablePred L (Γ, m) fun x => x ≤ a → P (a - x) by exact this - definability + · suffices Γ-[m].BoldfacePred fun x => x ≤ a → P (a - x) by exact this + apply HierarchySymbol.Boldface.imp + · apply HierarchySymbol.Boldface.bcomp₂ (by definability) (by definability) + · apply HierarchySymbol.Boldface.bcomp₁ (by definability) case zero => intro _; simpa using ha case succ x IH => @@ -117,28 +116,27 @@ private lemma neg_induction_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) · exact this · exact le_tsub_of_add_le_left hx) case inst => exact inferInstance - case inst => exact inferInstance have : P 0 := by simpa using this a (by rfl) contradiction -lemma models_indScheme_alt : M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy Γ.alt m) := by +lemma models_indScheme_alt : V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy Γ.alt m) := by simp [Theory.indH, Theory.indScheme] rintro _ p hp rfl simp [models_iff, succInd, Semiformula.eval_rew_q, Semiformula.eval_substs, Function.comp, Matrix.constant_eq_singleton] intro v H0 Hsucc x - have : Semiformula.Evalm M ![0] v p → - (∀ x, Semiformula.Evalm M ![x] v p → Semiformula.Evalm M ![x + 1] v p) → - ∀ x, Semiformula.Evalm M ![x] v p := by + have : Semiformula.Evalm V ![0] v p → + (∀ x, Semiformula.Evalm V ![x] v p → Semiformula.Evalm V ![x + 1] v p) → + ∀ x, Semiformula.Evalm V ![x] v p := by simpa using - neg_induction_h (L := L) Γ m (P := λ x ↦ ¬Semiformula.Evalm M ![x] v p) + neg_induction_h Γ m (P := λ x ↦ ¬Semiformula.Evalm V ![x] v p) (.mkPolarity (~(Rew.rewriteMap v).hom p) (by simpa using hp) (by intro x; simp [←Matrix.constant_eq_singleton', Semiformula.eval_rewriteMap])) exact this H0 Hsucc x -instance : M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy Γ.alt m) := models_indScheme_alt Γ m +instance : V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy Γ.alt m) := models_indScheme_alt Γ m -lemma least_number_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma least_number_h {P : V → Prop} (hP : Γ-[m].BoldfacePred P) {x} (h : P x) : ∃ y, P y ∧ ∀ z < y, ¬P z := by by_contra A have A : ∀ z, P z → ∃ w < z, P w := by simpa using A @@ -147,8 +145,10 @@ lemma least_number_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) induction z using induction_h · exact Γ.alt · exact m - · suffices DefinablePred L (Γ.alt, m) fun z ↦ ∀ w < z, ¬P w by exact this - definability + · suffices Γ.alt-[m].BoldfacePred fun z ↦ ∀ w < z, ¬P w by exact this + apply HierarchySymbol.Boldface.ball_blt (by definability) + apply HierarchySymbol.Boldface.not + apply HierarchySymbol.Boldface.bcomp₁ (hP := by simpa using hP) (by definability) case zero => simp case succ x IH => intro w hx hw @@ -158,121 +158,118 @@ lemma least_number_h {P : M → Prop} (hP : DefinablePred L (Γ, m) P) rcases this with ⟨v, hvw, hv⟩ exact IH v hvw hv case inst => exact inferInstance - case inst => exact inferInstance exact this (x + 1) x (by simp) h end neg section -variable (L) - -variable (Γ : SigmaPiDelta) (m : ℕ) [M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy 𝚺 m)] +variable (Γ : SigmaPiDelta) (m : ℕ) [V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy 𝚺 m)] -lemma induction_hh {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma induction_hh {P : V → Prop} (hP : Γ-[m].BoldfacePred P) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := match Γ with | 𝚺 => induction_h 𝚺 m hP zero succ | 𝚷 => - haveI : M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m + haveI : V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m induction_h 𝚷 m hP zero succ | 𝚫 => induction_h 𝚺 m hP.of_delta zero succ -lemma order_induction_hh {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma order_induction_hh {P : V → Prop} (hP : Γ-[m].BoldfacePred P) (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := match Γ with | 𝚺 => order_induction_h 𝚺 m hP ind | 𝚷 => - haveI : M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m + haveI : V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m order_induction_h 𝚷 m hP ind | 𝚫 => order_induction_h 𝚺 m hP.of_delta ind -lemma least_number_hh {P : M → Prop} (hP : DefinablePred L (Γ, m) P) +lemma least_number_hh {P : V → Prop} (hP : Γ-[m].BoldfacePred P) {x} (h : P x) : ∃ y, P y ∧ ∀ z < y, ¬P z := match Γ with | 𝚺 => least_number_h 𝚺 m hP h | 𝚷 => - haveI : M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m + haveI : V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy 𝚷 m) := models_indScheme_alt 𝚺 m least_number_h 𝚷 m hP h | 𝚫 => least_number_h 𝚺 m hP.of_delta h end -instance [M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy 𝚺 m)] : - M ⊧ₘ* Theory.indScheme L (Arith.Hierarchy Γ m) := by +instance [V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy 𝚺 m)] : + V ⊧ₘ* Theory.indScheme ℒₒᵣ (Arith.Hierarchy Γ m) := by rcases Γ · exact inferInstance · exact models_indScheme_alt 𝚺 m end -def mod_iOpen_of_mod_indH (Γ n) [M ⊧ₘ* 𝐈𝐍𝐃Γ n] : M ⊧ₘ* 𝐈open := +def mod_IOpen_of_mod_indH (Γ n) [V ⊧ₘ* 𝐈𝐍𝐃Γ n] : V ⊧ₘ* 𝐈open := ModelsTheory.of_ss (U := 𝐈𝐍𝐃Γ n) inferInstance (Set.union_subset_union_right _ (Theory.indScheme_subset Hierarchy.of_open)) -def mod_iSigma_of_le {n₁ n₂} (h : n₁ ≤ n₂) [M ⊧ₘ* 𝐈𝚺 n₂] : M ⊧ₘ* 𝐈𝚺 n₁ := +def mod_ISigma_of_le {n₁ n₂} (h : n₁ ≤ n₂) [V ⊧ₘ* 𝐈𝚺 n₂] : V ⊧ₘ* 𝐈𝚺 n₁ := ModelsTheory.of_ss inferInstance (Theory.iSigma_subset_mono h) -instance [M ⊧ₘ* 𝐈open] : M ⊧ₘ* 𝐏𝐀⁻ := ModelsTheory.of_add_left M 𝐏𝐀⁻ (Theory.indScheme _ Semiformula.Open) +instance [V ⊧ₘ* 𝐈open] : V ⊧ₘ* 𝐏𝐀⁻ := ModelsTheory.of_add_left V 𝐏𝐀⁻ (Theory.indScheme _ Semiformula.Open) -instance [M ⊧ₘ* 𝐈𝚺₀] : M ⊧ₘ* 𝐈open := mod_iOpen_of_mod_indH 𝚺 0 +instance [V ⊧ₘ* 𝐈𝚺₀] : V ⊧ₘ* 𝐈open := mod_IOpen_of_mod_indH 𝚺 0 -instance [M ⊧ₘ* 𝐈𝚺₁] : M ⊧ₘ* 𝐈𝚺₀ := mod_iSigma_of_le (show 0 ≤ 1 from by simp) +instance [V ⊧ₘ* 𝐈𝚺₁] : V ⊧ₘ* 𝐈𝚺₀ := mod_ISigma_of_le (show 0 ≤ 1 from by simp) -instance [M ⊧ₘ* 𝐈𝚺 n] : M ⊧ₘ* 𝐈𝚷 n := - haveI : M ⊧ₘ* 𝐏𝐀⁻ := Arith.models_peanoMinus_of_models_indH 𝚺 n +instance [V ⊧ₘ* 𝐈𝚺 n] : V ⊧ₘ* 𝐈𝚷 n := + haveI : V ⊧ₘ* 𝐏𝐀⁻ := Arith.models_peanoMinus_of_models_indH 𝚺 n inferInstance -instance [M ⊧ₘ* 𝐈𝚷 n] : M ⊧ₘ* 𝐈𝚺 n := - haveI : M ⊧ₘ* 𝐏𝐀⁻ := Arith.models_peanoMinus_of_models_indH 𝚷 n - by simp [*]; simpa [Theory.iPi] using models_indScheme_alt (L := ℒₒᵣ) (M := M) 𝚷 n +instance [V ⊧ₘ* 𝐈𝚷 n] : V ⊧ₘ* 𝐈𝚺 n := + haveI : V ⊧ₘ* 𝐏𝐀⁻ := Arith.models_peanoMinus_of_models_indH 𝚷 n + by simp [*]; simpa [Theory.iPi] using models_indScheme_alt (V := V) 𝚷 n -lemma models_iSigma_iff_models_iPi {n} : M ⊧ₘ* 𝐈𝚺 n ↔ M ⊧ₘ* 𝐈𝚷 n := +lemma models_ISigma_iff_models_IPi {n} : V ⊧ₘ* 𝐈𝚺 n ↔ V ⊧ₘ* 𝐈𝚷 n := ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ -instance [M ⊧ₘ* 𝐈𝚺 n] : M ⊧ₘ* 𝐈𝐍𝐃Γ n := +instance [V ⊧ₘ* 𝐈𝚺 n] : V ⊧ₘ* 𝐈𝐍𝐃Γ n := match Γ with | 𝚺 => inferInstance | 𝚷 => inferInstance -@[elab_as_elim] lemma induction_iSigmaZero [M ⊧ₘ* 𝐈𝚺₀] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₀ P) +@[elab_as_elim] lemma induction_sigma0 [V ⊧ₘ* 𝐈𝚺₀] + {P : V → Prop} (hP : 𝚺₀.BoldfacePred P) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := induction_h 𝚺 0 hP zero succ -@[elab_as_elim] lemma induction_iSigmaOne [M ⊧ₘ* 𝐈𝚺₁] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₁ P) +@[elab_as_elim] lemma induction_sigma1 [V ⊧ₘ* 𝐈𝚺₁] + {P : V → Prop} (hP : 𝚺₁-Predicate P) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := induction_h 𝚺 1 hP zero succ -@[elab_as_elim] lemma induction_iPiOne [M ⊧ₘ* 𝐈𝚺₁] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚷₁ P) +@[elab_as_elim] lemma induction_pi1 [V ⊧ₘ* 𝐈𝚺₁] + {P : V → Prop} (hP : 𝚷₁-Predicate P) (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := induction_h 𝚷 1 hP zero succ -@[elab_as_elim] lemma order_induction_iSigmaZero [M ⊧ₘ* 𝐈𝚺₀] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₀ P) +@[elab_as_elim] lemma order_induction_sigma0 [V ⊧ₘ* 𝐈𝚺₀] + {P : V → Prop} (hP : 𝚺₀-Predicate P) (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := order_induction_h 𝚺 0 hP ind -@[elab_as_elim] lemma order_induction_iSigmaOne [M ⊧ₘ* 𝐈𝚺₁] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₁ P) +@[elab_as_elim] lemma order_induction_sigma1 [V ⊧ₘ* 𝐈𝚺₁] + {P : V → Prop} (hP : 𝚺₁-Predicate P) (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := order_induction_h 𝚺 1 hP ind -@[elab_as_elim] lemma order_induction_piOne [M ⊧ₘ* 𝐈𝚺₁] - {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚷₁ P) +@[elab_as_elim] lemma order_induction_pi1 [V ⊧ₘ* 𝐈𝚺₁] + {P : V → Prop} (hP : 𝚷₁-Predicate P) (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := order_induction_h 𝚷 1 hP ind -lemma least_number_iSigmaZero [M ⊧ₘ* 𝐈𝚺₀] {P : M → Prop} (hP : DefinablePred ℒₒᵣ 𝚺₀ P) +lemma least_number_sigma0 [V ⊧ₘ* 𝐈𝚺₀] {P : V → Prop} (hP : 𝚺₀-Predicate P) {x} (h : P x) : ∃ y, P y ∧ ∀ z < y, ¬P z := least_number_h 𝚺 0 hP h -@[elab_as_elim] lemma induction_h_iSigmaOne [M ⊧ₘ* 𝐈𝚺₁] (Γ) - {P : M → Prop} (hP : DefinablePred ℒₒᵣ (Γ, 1) P) - (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := induction_hh ℒₒᵣ Γ 1 hP zero succ +@[elab_as_elim] lemma induction_h_sigma1 [V ⊧ₘ* 𝐈𝚺₁] (Γ) + {P : V → Prop} (hP : Γ-[1]-Predicate P) + (zero : P 0) (succ : ∀ x, P x → P (x + 1)) : ∀ x, P x := induction_hh Γ 1 hP zero succ -@[elab_as_elim] lemma order_induction_h_iSigmaOne [M ⊧ₘ* 𝐈𝚺₁] (Γ) - {P : M → Prop} (hP : DefinablePred ℒₒᵣ (Γ, 1) P) - (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := order_induction_hh ℒₒᵣ Γ 1 hP ind +@[elab_as_elim] lemma order_induction_h_sigma1 [V ⊧ₘ* 𝐈𝚺₁] (Γ) + {P : V → Prop} (hP : Γ-[1]-Predicate P) + (ind : ∀ x, (∀ y < x, P y) → P x) : ∀ x, P x := order_induction_hh Γ 1 hP ind end LO.Arith diff --git a/Arithmetization/Basic/PeanoMinus.lean b/Arithmetization/Basic/PeanoMinus.lean index f354df7..93853a2 100644 --- a/Arithmetization/Basic/PeanoMinus.lean +++ b/Arithmetization/Basic/PeanoMinus.lean @@ -1,4 +1,4 @@ -import Arithmetization.Vorspiel.Lemmata +import Arithmetization.Definability.Absoluteness namespace LO.Arith @@ -6,22 +6,22 @@ open FirstOrder FirstOrder.Arith noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐏𝐀⁻] -variable {a b c : M} +variable {a b c : V} section sub -lemma sub_existsUnique (a b : M) : ∃! c, (a ≥ b → a = b + c) ∧ (a < b → c = 0) := by +lemma sub_existsUnique (a b : V) : ∃! c, (a ≥ b → a = b + c) ∧ (a < b → c = 0) := by have : b ≤ a ∨ a < b := le_or_lt b a rcases this with (hxy | hxy) <;> simp [hxy] have : ∃ c, a = b + c := exists_add_of_le hxy rcases this with ⟨c, rfl⟩ exact ExistsUnique.intro c rfl (fun a h => (add_left_cancel h).symm) -def sub (a b : M) : M := Classical.choose! (sub_existsUnique a b) +def sub (a b : V) : V := Classical.choose! (sub_existsUnique a b) -instance : Sub M := ⟨sub⟩ +instance : Sub V := ⟨sub⟩ lemma sub_spec_of_ge (h : a ≥ b) : a = b + (a - b) := (Classical.choose!_spec (sub_existsUnique a b)).1 h @@ -29,28 +29,28 @@ lemma sub_spec_of_lt (h : a < b) : a - b = 0 := (Classical.choose!_spec (sub_exi lemma sub_eq_iff : c = a - b ↔ ((a ≥ b → a = b + c) ∧ (a < b → c = 0)) := Classical.choose!_eq_iff (sub_existsUnique a b) -@[simp] lemma sub_le_self (a b : M) : a - b ≤ a := by +@[simp] lemma sub_le_self (a b : V) : a - b ≤ a := by have : b ≤ a ∨ a < b := le_or_lt b a rcases this with (hxy | hxy) <;> simp[hxy] · simpa [← sub_spec_of_ge hxy] using show a - b ≤ b + (a - b) from le_add_self · simp[sub_spec_of_lt hxy] -open Definability +open FirstOrder.Arith.HierarchySymbol.Boldface -def _root_.LO.FirstOrder.Arith.subDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.subDef : 𝚺₀.Semisentence 3 := .mkSigma “z x y | (x ≥ y → x = y + z) ∧ (x < y → z = 0)” (by simp[Hierarchy.pi_zero_iff_sigma_zero]) -lemma sub_defined : 𝚺₀-Function₂ ((· - ·) : M → M → M) via subDef := by +lemma sub_defined : 𝚺₀-Function₂ ((· - ·) : V → V → V) via subDef := by intro v; simp [FirstOrder.Arith.subDef, sub_eq_iff] @[simp] lemma sub_defined_iff (v) : - Semiformula.Evalbm M v subDef.val ↔ v 0 = v 1 - v 2 := sub_defined.df.iff v + Semiformula.Evalbm V v subDef.val ↔ v 0 = v 1 - v 2 := sub_defined.df.iff v -instance sub_definable (Γ) : DefinableFunction₂ ℒₒᵣ Γ ((· - ·) : M → M → M) := Defined.to_definable₀ _ sub_defined +instance sub_definable (ℌ : HierarchySymbol) : ℌ.BoldfaceFunction₂ ((· - ·) : V → V → V) := sub_defined.to_definable₀ -instance sub_polybounded : Bounded₂ ℒₒᵣ ((· - ·) : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance sub_polybounded : Bounded₂ ((· - ·) : V → V → V) := ⟨#0, λ _ ↦ by simp⟩ -@[simp] lemma sub_self (a : M) : a - a = 0 := +@[simp] lemma sub_self (a : V) : a - a = 0 := add_right_eq_self.mp (sub_spec_of_ge (a := a) (b := a) (by rfl)).symm lemma sub_spec_of_le (h : a ≤ b) : a - b = 0 := by @@ -65,9 +65,9 @@ lemma add_tsub_self_of_le (h : b ≤ a) : b + (a - b) = a := by symm; exact sub_ @[simp] lemma add_sub_self' : (b + a) - b = a := by simp [add_comm] -@[simp] lemma zero_sub (a : M) : 0 - a = 0 := sub_spec_of_le (by simp) +@[simp] lemma zero_sub (a : V) : 0 - a = 0 := sub_spec_of_le (by simp) -@[simp] lemma sub_zero (a : M) : a - 0 = a := by +@[simp] lemma sub_zero (a : V) : a - 0 = a := by simpa using sub_add_self_of_le (show 0 ≤ a from zero_le a) lemma sub_remove_left (e : a = b + c) : a - c = b := by simp[e] @@ -97,7 +97,7 @@ lemma sub_sub : a - b - c = a - (b + c) := by @[simp] lemma sub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := not_iff_not.mp (by simp [←pos_iff_ne_zero]) -instance : OrderedSub M where +instance : OrderedSub V where tsub_le_iff_right := by intro a b c by_cases h : b ≤ a @@ -107,7 +107,7 @@ instance : OrderedSub M where · simp [sub_spec_of_lt (show a < b from by simpa using h)] exact le_trans (le_of_lt $ show a < b from by simpa using h) (by simp) -lemma zero_or_succ (a : M) : a = 0 ∨ ∃ a', a = a' + 1 := by +lemma zero_or_succ (a : V) : a = 0 ∨ ∃ a', a = a' + 1 := by rcases zero_le a with (rfl | pos) · simp · right; exact ⟨a - 1, by rw [sub_add_self_of_le]; exact pos_iff_one_le.mp pos⟩ @@ -128,14 +128,14 @@ lemma sub_mul (h : b ≤ a) : (a - b) * c = a * c - b * c := by lemma mul_sub (h : b ≤ a) : c * (a - b) = c * a - c * b := by simp [mul_comm c, sub_mul, h] -lemma add_sub_of_le (h : c ≤ b) (a : M) : a + b - c = a + (b - c) := add_tsub_assoc_of_le h a +lemma add_sub_of_le (h : c ≤ b) (a : V) : a + b - c = a + (b - c) := add_tsub_assoc_of_le h a -lemma sub_succ_add_succ {x y : M} (h : y < x) (z) : x - (y + 1) + (z + 1) = x - y + z := calc +lemma sub_succ_add_succ {x y : V} (h : y < x) (z) : x - (y + 1) + (z + 1) = x - y + z := calc x - (y + 1) + (z + 1) = x - (y + 1) + 1 + z := by simp [add_assoc, add_comm] _ = x - y - 1 + 1 + z := by simp [sub_sub] _ = x - y + z := by simp; rw [sub_add_self_of_le (one_le_of_zero_lt _ (pos_sub_iff_lt.mpr h))] -lemma le_sub_one_of_lt {a b : M} (h : a < b) : a ≤ b - 1 := by +lemma le_sub_one_of_lt {a b : V} (h : a < b) : a ≤ b - 1 := by have : 1 ≤ b := one_le_of_zero_lt _ (pos_of_gt h) simp [le_iff_lt_succ, sub_add_self_of_le this, h] @@ -150,23 +150,23 @@ lemma le_mul_self_of_pos_left (hy : 0 < b) : a ≤ b * a := by lemma le_mul_self_of_pos_right (hy : 0 < b) : a ≤ a * b := by simpa [mul_comm a b] using le_mul_self_of_pos_left hy -lemma dvd_iff_bounded {a b : M} : a ∣ b ↔ ∃ c ≤ b, b = a * c := by +lemma dvd_iff_bounded {a b : V} : a ∣ b ↔ ∃ c ≤ b, b = a * c := by by_cases hx : a = 0 · simp[hx]; rintro rfl; exact ⟨0, by simp⟩ · constructor · rintro ⟨c, rfl⟩; exact ⟨c, le_mul_self_of_pos_left (pos_iff_ne_zero.mpr hx), rfl⟩ · rintro ⟨c, hz, rfl⟩; exact dvd_mul_right a c -def _root_.LO.FirstOrder.Arith.dvd : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.dvd : 𝚺₀.Semisentence 2 := .mkSigma “x y | ∃ z <⁺ y, y = x * z” (by simp) -lemma dvd_defined : 𝚺₀-Relation (fun a b : M ↦ a ∣ b) via dvd := +lemma dvd_defined : 𝚺₀-Relation (fun a b : V ↦ a ∣ b) via dvd := fun v ↦ by simp [dvd_iff_bounded, Matrix.vecHead, Matrix.vecTail, dvd] @[simp] lemma dvd_defined_iff (v) : - Semiformula.Evalbm M v dvd.val ↔ v 0 ∣ v 1 := dvd_defined.df.iff v + Semiformula.Evalbm V v dvd.val ↔ v 0 ∣ v 1 := dvd_defined.df.iff v -instance dvd_definable (Γ) : DefinableRel ℒₒᵣ Γ ((· ∣ ·) : M → M → Prop) := Defined.to_definable₀ _ dvd_defined +instance dvd_definable (ℌ : HierarchySymbol) : ℌ.BoldfaceRel ((· ∣ ·) : V → V → Prop) := dvd_defined.to_definable₀ section @@ -197,22 +197,22 @@ lemma dvd_antisymm : a ∣ b → b ∣ a → a = b := by lemma dvd_one_iff : a ∣ 1 ↔ a = 1 := ⟨by { intro hx; exact dvd_antisymm hx (by simp) }, by rintro rfl; simp⟩ -theorem units_eq_one (u : Mˣ) : u = 1 := +theorem units_eq_one (u : Vˣ) : u = 1 := Units.ext <| dvd_one_iff.mp ⟨u.inv, u.val_inv.symm⟩ -@[simp] lemma unit_iff_eq_one {a : M} : IsUnit a ↔ a = 1 := +@[simp] lemma unit_iff_eq_one {a : V} : IsUnit a ↔ a = 1 := ⟨by rintro ⟨u, rfl⟩; simp [units_eq_one u], by rintro rfl; simp⟩ section Prime -lemma eq_one_or_eq_of_dvd_of_prime {p a : M} (pp : Prime p) (hxp : a ∣ p) : a = 1 ∨ a = p := by +lemma eq_one_or_eq_of_dvd_of_prime {p a : V} (pp : Prime p) (hxp : a ∣ p) : a = 1 ∨ a = p := by have : p ∣ a ∨ a ∣ 1 := pp.left_dvd_or_dvd_right_of_dvd_mul (show a ∣ p * 1 from by simpa using hxp) rcases this with (hx | hx) · right; exact dvd_antisymm hxp hx · left; exact dvd_one_iff.mp hx /- -lemma irreducible_iff_bounded {a : M} : Irreducible a ↔ 1 < a ∧ ∀ b ≤ a, (b ∣ a → b = 1 ∨ b = a) := by +lemma irreducible_iff_bounded {a : V} : Irreducible a ↔ 1 < a ∧ ∀ b ≤ a, (b ∣ a → b = 1 ∨ b = a) := by constructor · intro ha have : 1 < a := by @@ -220,7 +220,7 @@ lemma irreducible_iff_bounded {a : M} : Irreducible a ↔ 1 < a ∧ ∀ b ≤ a, simp [Irreducible.ne_one ha, Irreducible.ne_zero ha, le_one_iff_eq_zero_or_one] at A exact ⟨this, by { }⟩ -lemma prime_iff_bounded {a : M} : Prime a ↔ 1 < a ∧ ∀ b ≤ a, (b ∣ a → b = 1 ∨ b = a) := by +lemma prime_iff_bounded {a : V} : Prime a ↔ 1 < a ∧ ∀ b ≤ a, (b ∣ a → b = 1 ∨ b = a) := by constructor · intro prim have : 1 < a := by @@ -236,13 +236,13 @@ lemma prime_iff_bounded {a : M} : Prime a ↔ 1 < a ∧ ∀ b ≤ a, (b ∣ a · intro b c h -/ -def IsPrime (a : M) : Prop := 1 < a ∧ ∀ b ≤ a, b ∣ a → b = 1 ∨ b = a +def IsPrime (a : V) : Prop := 1 < a ∧ ∀ b ≤ a, b ∣ a → b = 1 ∨ b = a -- TODO: prove IsPrime a ↔ Prime a -def _root_.LO.FirstOrder.Arith.isPrime : 𝚺₀-Semisentence 1 := +def _root_.LO.FirstOrder.Arith.isPrime : 𝚺₀.Semisentence 1 := .mkSigma “x | 1 < x ∧ ∀ y <⁺ x, !dvd.val y x → y = 1 ∨ y = x” (by simp [Hierarchy.pi_zero_iff_sigma_zero]) -lemma isPrime_defined : 𝚺₀-Predicate (λ a : M ↦ IsPrime a) via isPrime := by +lemma isPrime_defined : 𝚺₀-Predicate (λ a : V ↦ IsPrime a) via isPrime := by intro v simp [Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.vecHead, Matrix.constant_eq_singleton, IsPrime, isPrime] @@ -251,10 +251,10 @@ end Prime section min -def _root_.LO.FirstOrder.Arith.min : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.min : 𝚺₀.Semisentence 3 := .mkSigma “z x y | (x ≤ y → z = x) ∧ (x ≥ y → z = y)” (by simp) -lemma min_defined : 𝚺₀-Function₂ (min : M → M → M) via min := by +lemma min_defined : 𝚺₀-Function₂ (min : V → V → V) via min := by intro v; simp [FirstOrder.Arith.min] rcases le_total (v 1) (v 2) with (h | h) <;> simp [h] · intro h₀₁ h₂₁ @@ -263,20 +263,20 @@ lemma min_defined : 𝚺₀-Function₂ (min : M → M → M) via min := by exact le_antisymm (by simpa [h₀₂] using h) (by simpa [h₀₂] using h₁₂) @[simp] lemma eval_minDef (v) : - Semiformula.Evalbm M v min.val ↔ v 0 = min (v 1) (v 2) := min_defined.df.iff v + Semiformula.Evalbm V v min.val ↔ v 0 = min (v 1) (v 2) := min_defined.df.iff v -instance min_definable (Γ) : Γ-Function₂ (min : M → M → M) := Defined.to_definable₀ _ min_defined +instance min_definable (ℌ) : ℌ-Function₂ (min : V → V → V) := HierarchySymbol.Defined.to_definable₀ min_defined -instance min_polybounded : Bounded₂ ℒₒᵣ (min : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance min_polybounded : Bounded₂ (min : V → V → V) := ⟨#0, λ _ ↦ by simp⟩ end min section max -def _root_.LO.FirstOrder.Arith.max : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.max : 𝚺₀.Semisentence 3 := .mkSigma “z x y | (x ≥ y → z = x) ∧ (x ≤ y → z = y)” (by simp) -lemma max_defined : 𝚺₀-Function₂ (max : M → M → M) via max := by +lemma max_defined : 𝚺₀-Function₂ (max : V → V → V) via max := by intro v; simp [Arith.max] rcases le_total (v 1) (v 2) with (h | h) <;> simp [h] · intro h₀₂ h₂₁ @@ -285,11 +285,11 @@ lemma max_defined : 𝚺₀-Function₂ (max : M → M → M) via max := by exact le_antisymm (by simpa [h₀₁] using h₁₂) (by simpa [h₀₁] using h) @[simp] lemma eval_maxDef (v) : - Semiformula.Evalbm M v max.val ↔ v 0 = max (v 1) (v 2) := max_defined.df.iff v + Semiformula.Evalbm V v max.val ↔ v 0 = max (v 1) (v 2) := max_defined.df.iff v -instance max_definable (Γ) : Γ-Function₂ (max : M → M → M) := Defined.to_definable₀ _ max_defined +instance max_definable (Γ) : Γ-Function₂ (max : V → V → V) := HierarchySymbol.Defined.to_definable₀ max_defined -instance max_polybounded : Bounded₂ ℒₒᵣ (max : M → M → M) := ⟨‘#0 + #1’, λ v ↦ by simp⟩ +instance max_polybounded : Bounded₂ (max : V → V → V) := ⟨‘#0 + #1’, λ v ↦ by simp⟩ end max diff --git a/Arithmetization/Definability/Absoluteness.lean b/Arithmetization/Definability/Absoluteness.lean index 5073a9e..59784bc 100644 --- a/Arithmetization/Definability/Absoluteness.lean +++ b/Arithmetization/Definability/Absoluteness.lean @@ -1,4 +1,4 @@ -import Arithmetization.Definability.Basic +import Arithmetization.Definability.BoundedBoldface noncomputable section @@ -10,50 +10,50 @@ lemma nat_modelsWithParam_iff_models_substs {v : Fin k → ℕ} {p : Semisentenc ℕ ⊧ₘ[v] p ↔ ℕ ⊧ₘ (Rew.substs (fun i ↦ Semiterm.Operator.numeral ℒₒᵣ (v i)) |>.hom p) := by simp [models_iff] -variable (M : Type*) [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐏𝐀⁻] +variable (M : Type*) [ORingStruc M] [M ⊧ₘ* 𝐏𝐀⁻] lemma modelsWithParam_iff_models_substs {v : Fin k → ℕ} {p : Semisentence ℒₒᵣ k} : M ⊧ₘ[fun i ↦ v i] p ↔ M ⊧ₘ (Rew.substs (fun i ↦ Semiterm.Operator.numeral ℒₒᵣ (v i)) |>.hom p) := by simp [models_iff, numeral_eq_natCast] -lemma shigmaZero_absolute {k} (p : 𝚺₀-Semisentence k) (v : Fin k → ℕ) : +lemma shigmaZero_absolute {k} (p : 𝚺₀.Semisentence k) (v : Fin k → ℕ) : ℕ ⊧ₘ[v] p.val ↔ M ⊧ₘ[fun i ↦ (v i)] p.val := ⟨by simp [nat_modelsWithParam_iff_models_substs, modelsWithParam_iff_models_substs]; exact nat_extention_sigmaOne M (by simp), by simp [nat_modelsWithParam_iff_models_substs, modelsWithParam_iff_models_substs]; exact nat_extention_piOne M (by simp)⟩ -lemma Defined.shigmaZero_absolute {k} {R : (Fin k → ℕ) → Prop} {R' : (Fin k → M) → Prop} {p : 𝚺₀-Semisentence k} - (hR : Defined R p) (hR' : Defined R' p) (v : Fin k → ℕ) : +lemma Defined.shigmaZero_absolute {k} {R : (Fin k → ℕ) → Prop} {R' : (Fin k → M) → Prop} {p : 𝚺₀.Semisentence k} + (hR : 𝚺₀.Defined R p) (hR' : 𝚺₀.Defined R' p) (v : Fin k → ℕ) : R v ↔ R' (fun i ↦ (v i : M)) := by simpa [hR.iff, hR'.iff] using Arith.shigmaZero_absolute M p v -lemma DefinedFunction.shigmaZero_absolute_func {k} {f : (Fin k → ℕ) → ℕ} {f' : (Fin k → M) → M} {p : 𝚺₀-Semisentence (k + 1)} - (hf : DefinedFunction f p) (hf' : DefinedFunction f' p) (v : Fin k → ℕ) : +lemma DefinedFunction.shigmaZero_absolute_func {k} {f : (Fin k → ℕ) → ℕ} {f' : (Fin k → M) → M} {p : 𝚺₀.Semisentence (k + 1)} + (hf : 𝚺₀.DefinedFunction f p) (hf' : 𝚺₀.DefinedFunction f' p) (v : Fin k → ℕ) : (f v : M) = f' (fun i ↦ (v i)) := by simpa using Defined.shigmaZero_absolute M hf hf' (f v :> v) -lemma sigmaOne_upward_absolute {k} (p : 𝚺₁-Semisentence k) (v : Fin k → ℕ) : +lemma sigmaOne_upward_absolute {k} (p : 𝚺₁.Semisentence k) (v : Fin k → ℕ) : ℕ ⊧ₘ[v] p.val → M ⊧ₘ[fun i ↦ (v i)] p.val := by simp [nat_modelsWithParam_iff_models_substs, modelsWithParam_iff_models_substs] exact nat_extention_sigmaOne M (by simp) -lemma piOne_downward_absolute {k} (p : 𝚷₁-Semisentence k) (v : Fin k → ℕ) : +lemma piOne_downward_absolute {k} (p : 𝚷₁.Semisentence k) (v : Fin k → ℕ) : M ⊧ₘ[fun i ↦ (v i)] p.val → ℕ ⊧ₘ[v] p.val := by simp [nat_modelsWithParam_iff_models_substs, modelsWithParam_iff_models_substs] exact nat_extention_piOne M (by simp) -lemma deltaOne_absolute {k} (p : 𝚫₁-Semisentence k) +lemma deltaOne_absolute {k} (p : 𝚫₁.Semisentence k) (properNat : p.ProperOn ℕ) (proper : p.ProperOn M) (v : Fin k → ℕ) : ℕ ⊧ₘ[v] p.val ↔ M ⊧ₘ[fun i ↦ (v i)] p.val := - ⟨by simpa [HSemiformula.val_sigma] using sigmaOne_upward_absolute M p.sigma v, + ⟨by simpa [HierarchySymbol.Semiformula.val_sigma] using sigmaOne_upward_absolute M p.sigma v, by simpa [proper.iff', properNat.iff'] using piOne_downward_absolute M p.pi v⟩ -lemma Defined.shigmaOne_absolute {k} {R : (Fin k → ℕ) → Prop} {R' : (Fin k → M) → Prop} {p : 𝚫₁-Semisentence k} - (hR : Defined R p) (hR' : Defined R' p) (v : Fin k → ℕ) : +lemma Defined.shigmaOne_absolute {k} {R : (Fin k → ℕ) → Prop} {R' : (Fin k → M) → Prop} {p : 𝚫₁.Semisentence k} + (hR : 𝚫₁.Defined R p) (hR' : 𝚫₁.Defined R' p) (v : Fin k → ℕ) : R v ↔ R' (fun i ↦ (v i : M)) := by simpa [hR.df.iff, hR'.df.iff] using deltaOne_absolute M p hR.proper hR'.proper v -lemma DefinedFunction.shigmaOne_absolute_func {k} {f : (Fin k → ℕ) → ℕ} {f' : (Fin k → M) → M} {p : 𝚺₁-Semisentence (k + 1)} - (hf : DefinedFunction f p) (hf' : DefinedFunction f' p) (v : Fin k → ℕ) : +lemma DefinedFunction.shigmaOne_absolute_func {k} {f : (Fin k → ℕ) → ℕ} {f' : (Fin k → M) → M} {p : 𝚺₁.Semisentence (k + 1)} + (hf : 𝚺₁.DefinedFunction f p) (hf' : 𝚺₁.DefinedFunction f' p) (v : Fin k → ℕ) : (f v : M) = f' (fun i ↦ (v i)) := by simpa using Defined.shigmaOne_absolute M hf.graph_delta hf'.graph_delta (f v :> v) diff --git a/Arithmetization/Definability/Boldface.lean b/Arithmetization/Definability/Boldface.lean index 1336f97..9fbf8a1 100644 --- a/Arithmetization/Definability/Boldface.lean +++ b/Arithmetization/Definability/Boldface.lean @@ -41,7 +41,7 @@ section variable (ℌ) -class Definable {k} (P : (Fin k → V) → Prop) : Prop where +class Lightface {k} (P : (Fin k → V) → Prop) : Prop where definable : ∃ p : ℌ.Semisentence k, Defined P p class Boldface {k} (P : (Fin k → V) → Prop) : Prop where @@ -107,6 +107,43 @@ abbrev BoldfaceFunction₅ (f : V → V → V → V → V → V) : Prop := ℌ.B variable {ℌ} +notation Γ "-Predicate " P " via " p => DefinedPred Γ P p + +notation Γ "-Relation " P " via " p => DefinedRel Γ P p + +notation Γ "-Relation₃ " P " via " p => DefinedRel₃ Γ P p + +notation Γ "-Relation₄ " P " via " p => DefinedRel₄ Γ P p + +notation Γ "-Function₁ " f " via " p => DefinedFunction₁ Γ f p + +notation Γ "-Function₂ " f " via " p => DefinedFunction₂ Γ f p + +notation Γ "-Function₃ " f " via " p => DefinedFunction₃ Γ f p + +notation Γ "-Function₄ " f " via " p => DefinedFunction₄ Γ f p + +notation Γ "-Function₅ " f " via " p => DefinedFunction₅ Γ f p + +notation Γ "-Predicate " P => BoldfacePred Γ P + +notation Γ "-Relation " P => BoldfaceRel Γ P + +notation Γ "-Relation₃ " P => BoldfaceRel₃ Γ P + +notation Γ "-Relation₄ " P => BoldfaceRel₄ Γ P + +notation Γ "-Relation₅ " P => BoldfaceRel₅ Γ P + +notation Γ "-Function₁ " f => BoldfaceFunction₁ Γ f + +notation Γ "-Function₂ " f => BoldfaceFunction₂ Γ f + +notation Γ "-Function₃ " f => BoldfaceFunction₃ Γ f + +notation Γ "-Function₄ " f => BoldfaceFunction₄ Γ f + + end section @@ -146,14 +183,14 @@ lemma to_definable (p : ℌ.Semisentence k) (hP : Defined P p) : ℌ.Boldface P fun v ↦ by rcases p; simpa [HierarchySymbol.Semiformula.rew] using hP.proper.rew Rew.emb v, by intro; simp [hP.df.iff]⟩⟩ -lemma to_definable₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : +lemma to_definable₀ {p : 𝚺₀.Semisentence k} (hP : Defined P p) : ℌ.Boldface P := Defined.to_definable (p.ofZero ℌ) hP.of_zero lemma to_definable_oRing (p : ℌ.Semisentence k) (hP : Defined P p) : ℌ.Boldface P := Defined.to_definable p.emb hP.emb lemma to_definable_oRing₀ (p : 𝚺₀.Semisentence k) (hP : Defined P p) : - ℌ.Boldface P := Defined.to_definable₀ p.emb hP.emb + ℌ.Boldface P := Defined.to_definable₀ hP.emb end Defined diff --git a/Arithmetization/Definability/BoundedBoldface.lean b/Arithmetization/Definability/BoundedBoldface.lean index 7e05dc3..96a367d 100644 --- a/Arithmetization/Definability/BoundedBoldface.lean +++ b/Arithmetization/Definability/BoundedBoldface.lean @@ -134,7 +134,7 @@ namespace HierarchySymbol.Boldface variable {P Q : (Fin k → V) → Prop} -lemma ball_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma ball_blt {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : ℌ.Boldface fun v ↦ ∀ x < f v, P v x := by rcases hf.bounded with ⟨bf, hbf⟩ @@ -147,7 +147,7 @@ lemma ball_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (F exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) -lemma bex_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma bex_blt {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : ℌ.Boldface fun v ↦ ∃ x < f v, P v x := by rcases hf.bounded with ⟨bf, hbf⟩ @@ -159,7 +159,7 @@ lemma bex_lt_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fi simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex #0)).bex ‘!!bf + 1’ exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) -lemma ball_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma ball_ble {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : ℌ.Boldface fun v ↦ ∀ x ≤ f v, P v x := by rcases hf.bounded with ⟨bf, hbf⟩ @@ -171,7 +171,7 @@ lemma ball_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (F simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).ball ‘x | x + 1’)).bex ‘!!bf + 1’ exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) -lemma bex_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma bex_ble {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : ℌ.Boldface fun w ↦ P (w ·.succ) (w 0)) : ℌ.Boldface fun v ↦ ∃ x ≤ f v, P v x := by rcases hf.bounded with ⟨bf, hbf⟩ @@ -183,21 +183,21 @@ lemma bex_le_boldfaceBoundedFunction {P : (Fin k → V) → V → Prop} {f : (Fi simpa [←le_iff_lt_succ] using (hf_graph.and ((hp.retraction (0 :> (·.succ.succ))).bex ‘x | x + 1’)).bex ‘!!bf + 1’ exact .of_iff ⟨_, this⟩ (fun v ↦ ⟨fun h ↦ ⟨f v, hbf v, rfl, h⟩, by rintro ⟨y, hy, rfl, h⟩; exact h⟩) -lemma ball_lt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma ball_blt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : - Γ-[0].Boldface fun v ↦ ∀ x < f v, P v x := ball_lt_boldfaceBoundedFunction hf h + Γ-[0].Boldface fun v ↦ ∀ x < f v, P v x := ball_blt hf h -lemma bex_lt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma bex_blt_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : - Γ-[0].Boldface fun v ↦ ∃ x < f v, P v x := bex_lt_boldfaceBoundedFunction hf h + Γ-[0].Boldface fun v ↦ ∃ x < f v, P v x := bex_blt hf h -lemma ball_le_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma ball_ble_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : - Γ-[0].Boldface fun v ↦ ∀ x ≤ f v, P v x := ball_le_boldfaceBoundedFunction hf h + Γ-[0].Boldface fun v ↦ ∀ x ≤ f v, P v x := ball_ble hf h -lemma bex_le_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} +lemma bex_ble_zero {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} (hf : BoldfaceBoundedFunction f) (h : Γ-[0].Boldface fun w ↦ P (w ·.succ) (w 0)) : - Γ-[0].Boldface fun v ↦ ∃ x ≤ f v, P v x := bex_le_boldfaceBoundedFunction hf h + Γ-[0].Boldface fun v ↦ ∃ x ≤ f v, P v x := bex_ble hf h lemma bex_vec_le_boldfaceBoundedFunction {k} {p : Fin l → (Fin k → V) → V} {P : (Fin k → V) → (Fin l → V) → Prop} (pp : ∀ i, BoldfaceBoundedFunction (p i)) (hP : ℌ.Boldface fun w : Fin (k + l) → V ↦ P (fun i ↦ w (i.castAdd l)) (fun j ↦ w (j.natAdd k))) : @@ -206,7 +206,7 @@ lemma bex_vec_le_boldfaceBoundedFunction {k} {p : Fin l → (Fin k → V) → V} case zero => simpa [Matrix.empty_eq (α := V)] using hP case succ l ih => simp only [exists_le_vec_iff_exists_le_exists_vec] - apply bex_le_boldfaceBoundedFunction (pp 0) + apply bex_ble (pp 0) apply ih · intro i; apply BoldfaceBoundedFunction.retraction (pp i.succ) · let g : Fin (k + (l + 1)) → Fin (k + 1 + l) := Matrix.vecAppend rfl (fun x ↦ x.succ.castAdd l) (Fin.castAdd l 0 :> fun j ↦ j.natAdd (k + 1)) @@ -259,22 +259,43 @@ namespace HierarchySymbol.Boldface open BoldfaceBoundedFunction -lemma bcomp₁ {k} {P : V → Prop} {f : (Fin k → V) → V} [hP : Γ-[0].BoldfacePred P] (hf : BoldfaceBoundedFunction f) : +lemma bcomp₁ {k} {P : V → Prop} {f : (Fin k → V) → V} [hP : ℌ.BoldfacePred P] (hf : BoldfaceBoundedFunction f) : + ℌ.Boldface fun v ↦ P (f v) := + substitution_boldfaceBoundedFunction (f := ![f]) hP (by simp [*]) + +lemma bcomp₂ {k} {R : V → V → Prop} {f₁ f₂ : (Fin k → V) → V} [hR : ℌ.BoldfaceRel R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : + ℌ.Boldface fun v ↦ R (f₁ v) (f₂ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma bcomp₃ {k} {R : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} [hR : ℌ.BoldfaceRel₃ R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) : + ℌ.Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂, f₃]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma bcomp₄ {k} {R : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} [hR : ℌ.BoldfaceRel₄ R] + (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) + (hf₃ : BoldfaceBoundedFunction f₃) (hf₄ : BoldfaceBoundedFunction f₄) : + ℌ.Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) (f₄ v) := + substitution_boldfaceBoundedFunction (f := ![f₁, f₂, f₃, f₄]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) + +lemma bcomp₁_zero {k} {P : V → Prop} {f : (Fin k → V) → V} [hP : Γ-[0].BoldfacePred P] (hf : BoldfaceBoundedFunction f) : Γ-[0].Boldface fun v ↦ P (f v) := substitution_boldfaceBoundedFunction (f := ![f]) hP (by simp [*]) -lemma bcomp₂ {k} {R : V → V → Prop} {f₁ f₂ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel R] +lemma bcomp₂_zero {k} {R : V → V → Prop} {f₁ f₂ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel R] (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) : Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) := substitution_boldfaceBoundedFunction (f := ![f₁, f₂]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) -lemma bcomp₃ {k} {R : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₃ R] +lemma bcomp₃_zero {k} {R : V → V → V → Prop} {f₁ f₂ f₃ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₃ R] (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) (hf₃ : BoldfaceBoundedFunction f₃) : Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) := substitution_boldfaceBoundedFunction (f := ![f₁, f₂, f₃]) hR (by simp [forall_fin_iff_zero_and_forall_succ, *]) -lemma bcomp₄ {k} {R : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₄ R] +lemma bcomp₄_zero {k} {R : V → V → V → V → Prop} {f₁ f₂ f₃ f₄ : (Fin k → V) → V} [hR : Γ-[0].BoldfaceRel₄ R] (hf₁ : BoldfaceBoundedFunction f₁) (hf₂ : BoldfaceBoundedFunction f₂) (hf₃ : BoldfaceBoundedFunction f₃) (hf₄ : BoldfaceBoundedFunction f₄) : Γ-[0].Boldface fun v ↦ R (f₁ v) (f₂ v) (f₃ v) (f₄ v) := @@ -366,26 +387,20 @@ attribute [aesop 6 (rule_sets := [Definability]) safe] Boldface.comp₃ Boldface.comp₄ Boldface.const - -attribute [aesop 7 (rule_sets := [Definability]) safe] - Boldface.bcomp₁ - Boldface.bcomp₂ - Boldface.bcomp₃ - Boldface.bcomp₄ + Boldface.bcomp₁_zero + Boldface.bcomp₂_zero + Boldface.bcomp₃_zero + Boldface.bcomp₄_zero attribute [aesop 8 (rule_sets := [Definability]) safe] Boldface.ball_lt Boldface.ball_le - Boldface.ball_lt' - Boldface.ball_le' Boldface.bex_lt Boldface.bex_le - -attribute [aesop 9 (rule_sets := [Definability]) safe] - Boldface.ball_lt_zero - Boldface.ball_le_zero - Boldface.bex_lt_zero - Boldface.bex_le_zero + Boldface.ball_blt_zero + Boldface.ball_ble_zero + Boldface.bex_blt_zero + Boldface.bex_ble_zero attribute [aesop 10 (rule_sets := [Definability]) safe] Boldface.not diff --git a/Arithmetization/ISigmaOne/Bit.lean b/Arithmetization/ISigmaOne/Bit.lean index c1e6b0f..a496471 100644 --- a/Arithmetization/ISigmaOne/Bit.lean +++ b/Arithmetization/ISigmaOne/Bit.lean @@ -15,7 +15,7 @@ def Bit (i a : M) : Prop := LenBit (exp i) a instance : Membership M M := ⟨Bit⟩ -def _root_.LO.FirstOrder.Arith.bitDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.bitDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | ∃ z <⁺ y, !expDef z x ∧ !lenbitDef z y” (by simp) lemma bit_defined : 𝚺₀-Relation ((· ∈ ·) : M → M → Prop) via bitDef := by @@ -97,10 +97,10 @@ def bexIn (t : Semiterm ℒₒᵣ ξ n) (p : Semiformula ℒₒᵣ ξ (n + 1)) : intros simp [Semiformula.Operator.operator, operator_mem_def] -def memRel : 𝚺₀-Semisentence 3 := .mkSigma +def memRel : 𝚺₀.Semisentence 3 := .mkSigma “R x y | ∃ p <⁺ (x + y + 1)², !pairDef p x y ∧ p ∈ R” (by simp) -def memRel₃ : 𝚺₀-Semisentence 4 := .mkSigma +def memRel₃ : 𝚺₀.Semisentence 4 := .mkSigma “R x y z | ∃ yz <⁺ (y + z + 1)², !pairDef yz y z ∧ ∃ xyz <⁺ (x + yz + 1)², !pairDef xyz x yz ∧ xyz ∈ R” (by simp) def memRelOpr : Semiformula.Operator ℒₒᵣ 3 := ⟨memRel.val⟩ @@ -269,7 +269,7 @@ lemma insert_graph (b i a : M) : not_false_eq_true, true_and, false_or, forall_exists_index, and_imp] rintro x _ rfl rfl; rfl ⟩ -def _root_.LO.FirstOrder.Arith.insertDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.insertDef : 𝚺₀.Semisentence 3 := .mkSigma “b i a | (i ∈ a ∧ b = a) ∨ (i ∉ a ∧ ∃ e <⁺ b, !expDef e i ∧ b = a + e)” (by simp) lemma insert_defined : 𝚺₀-Function₂ (insert : M → M → M) via insertDef := by @@ -328,7 +328,7 @@ lemma lt_exp_iff {a i : M} : a < exp i ↔ ∀ j ∈ a, j < i := instance : HasSubset M := ⟨fun a b ↦ ∀ ⦃i⦄, i ∈ a → i ∈ b⟩ -def _root_.LO.FirstOrder.Arith.bitSubsetDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.bitSubsetDef : 𝚺₀.Semisentence 2 := .mkSigma “a b | ∀ i < a, i ∈ a → i ∈ b” (by simp) lemma bitSubset_defined : 𝚺₀-Relation ((· ⊆ ·) : M → M → Prop) via bitSubsetDef := by @@ -386,7 +386,7 @@ def under (a : M) : M := exp a - 1 private lemma under_graph (x y : M) : y = under x ↔ y + 1 = exp x := ⟨by rintro rfl; simp [under, sub_add_self_of_le], by intro h; have := congr_arg (· - 1) h; simp [under] at this ⊢; exact this⟩ -def _root_.LO.FirstOrder.Arith.underDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.underDef : 𝚺₀.Semisentence 2 := .mkSigma “y x | !expDef.val (y + 1) x” (by simp) lemma under_defined : 𝚺₀-Function₁ (under : M → M) via underDef := by @@ -481,9 +481,9 @@ section variable {m : ℕ} [Fact (1 ≤ m)] [M ⊧ₘ* 𝐈𝐍𝐃𝚺 m] private lemma finset_comprehension_aux (Γ : Polarity) {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i := by - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) have : ∃ s < exp a, ∀ i < a, P i → i ∈ s := ⟨under a, pred_lt_self_of_pos (by simp), fun i hi _ ↦ by simpa [mem_under_iff] using hi⟩ rcases this with ⟨s, hsn, hs⟩ @@ -506,7 +506,7 @@ private lemma finset_comprehension_aux (Γ : Polarity) {P : M → Prop} (hP : ( exact ⟨t, lt_of_le_of_lt t_le_s hsn, fun i hi ↦ ⟨this i hi, ht i hi⟩⟩ theorem finset_comprehension {Γ} {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i := match Γ with | 𝚺 => finset_comprehension_aux 𝚺 hP a @@ -514,9 +514,9 @@ theorem finset_comprehension {Γ} {P : M → Prop} (hP : (Γ, m)-Predicate P) (a | 𝚫 => finset_comprehension_aux 𝚺 hP.of_delta a theorem finset_comprehension_exists_unique {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃! s, s < exp a ∧ ∀ i < a, i ∈ s ↔ P i := by - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) rcases finset_comprehension hP a with ⟨s, hs, Hs⟩ exact ExistsUnique.intro s ⟨hs, Hs⟩ (by intro t ⟨ht, Ht⟩ diff --git a/Arithmetization/ISigmaOne/HFS/Basic.lean b/Arithmetization/ISigmaOne/HFS/Basic.lean index ce228f4..28f413b 100644 --- a/Arithmetization/ISigmaOne/HFS/Basic.lean +++ b/Arithmetization/ISigmaOne/HFS/Basic.lean @@ -77,7 +77,7 @@ lemma sUnion_graph {u s : V} : u = ⋃ʰᶠ s ↔ ∀ x < u + s, (x ∈ u ↔ · rintro ⟨c, hc, hx⟩ exact h x (lt_of_lt_of_le (lt_trans (lt_of_mem hx) (lt_of_mem hc)) (by simp)) |>.mpr ⟨c, hc, hx⟩⟩ -def _root_.LO.FirstOrder.Arith.sUnionDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.sUnionDef : 𝚺₀.Semisentence 2 := .mkSigma “u s | ∀ x < u + s, (x ∈ u ↔ ∃ t ∈' s, x ∈ t)” (by simp) lemma sUnion_defined : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) via sUnionDef := by @@ -86,7 +86,7 @@ lemma sUnion_defined : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) via sUnionD @[simp] lemma sUnion_defined_iff (v) : Semiformula.Evalbm V v sUnionDef.val ↔ v 0 = ⋃ʰᶠ v 1 := sUnion_defined.df.iff v -instance sUnion_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ ((⋃ʰᶠ ·) : V → V) := Defined.to_definable _ sUnion_defined +instance sUnion_definable : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) := Defined.to_definable _ sUnion_defined instance sUnion_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ ((⋃ʰᶠ ·) : V → V) := .of_zero sUnion_definable _ @@ -110,7 +110,7 @@ private lemma union_graph {u s t : V} : u = s ∪ t ↔ ∀ x < u + s + t, (x · exact h x (lt_of_lt_of_le (lt_of_mem hx) (by simp )) |>.mpr (Or.inl hx) · exact h x (lt_of_lt_of_le (lt_of_mem hx) (by simp )) |>.mpr (Or.inr hx)⟩ -def _root_.LO.FirstOrder.Arith.unionDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.unionDef : 𝚺₀.Semisentence 3 := .mkSigma “∀[#0 < #1 + #2 + #3](#0 ∈ #1 ↔ #0 ∈ #2 ∨ #0 ∈ #3)” (by simp) lemma union_defined : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) via unionDef := by @@ -119,7 +119,7 @@ lemma union_defined : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) via unio @[simp] lemma union_defined_iff (v) : Semiformula.Evalbm V v unionDef.val ↔ v 0 = v 1 ∪ v 2 := union_defined.df.iff v -instance union_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ ((· ∪ ·) : V → V → V) := Defined.to_definable _ union_defined +instance union_definable : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) := Defined.to_definable _ union_defined instance union_definable' (Γ) : DefinableFunction₂ ℒₒᵣ Γ ((· ∪ ·) : V → V → V) := .of_zero union_definable _ @@ -137,7 +137,7 @@ lemma insert_eq_union_singleton (a s : V) : insert a s = {a} ∪ s := mem_ext (f _ < log (2 * (a + b)) := by simp [log_two_mul_of_pos (show 0 < a + b from by simp [pos_of_nonempty hi])] _ ≤ log (2 * (a + b) + 1) := log_monotone (by simp)) -instance : Bounded₂ ℒₒᵣ ((· ∪ ·) : V → V → V) := ⟨‘x y | 2 * (x + y)’, fun _ ↦ by simp⟩ +instance : Bounded₂ ((· ∪ ·) : V → V → V) := ⟨‘x y | 2 * (x + y)’, fun _ ↦ by simp⟩ lemma union_comm (a b : V) : a ∪ b = b ∪ a := mem_ext (by simp [or_comm]) @@ -232,7 +232,7 @@ private lemma product_graph {u a b : V} : u = a ×ʰᶠ b ↔ ∀ x < u + (a + b exact h ⟪y, z⟫ (lt_of_lt_of_le (pair_lt_pair (lt_of_mem hy) (lt_of_mem hz)) (le_trans (pair_polybound a b) <| by simp)) |>.mpr ⟨y, hy, z, hz, rfl⟩⟩ -def _root_.LO.FirstOrder.Arith.productDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.productDef : 𝚺₀.Semisentence 3 := .mkSigma “u a b | ∀ x < u + (a + b + 1)², (x ∈ u ↔ ∃ y ∈' a, ∃ z ∈' b, !pairDef x y z)” (by simp) lemma product_defined : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) via productDef := by @@ -241,7 +241,7 @@ lemma product_defined : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) vi @[simp] lemma product_defined_iff (v) : Semiformula.Evalbm V v productDef.val ↔ v 0 = v 1 ×ʰᶠ v 2 := product_defined.df.iff v -instance product_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ ((· ×ʰᶠ ·) : V → V → V) := Defined.to_definable _ product_defined +instance product_definable : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) := Defined.to_definable _ product_defined instance product_definable' (Γ) : DefinableFunction₂ ℒₒᵣ Γ ((· ×ʰᶠ ·) : V → V → V) := .of_zero product_definable _ @@ -276,7 +276,7 @@ private lemma domain_graph {u s : V} : u = domain s ↔ ∀ x < u + s, (x ∈ u exact h x (lt_of_lt_of_le (lt_of_le_of_lt (le_pair_left x y) (lt_of_mem hy)) (by simp)) |>.mpr ⟨y, lt_of_le_of_lt (le_pair_right x y) (lt_of_mem hy), _, hy, rfl⟩⟩ -def _root_.LO.FirstOrder.Arith.domainDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.domainDef : 𝚺₀.Semisentence 2 := .mkSigma “u s | ∀ x < u + s, (x ∈ u ↔ ∃ y < s, ∃ z ∈' s, !pairDef z x y)” (by simp) lemma domain_defined : 𝚺₀-Function₁ (domain : V → V) via domainDef := by @@ -285,7 +285,7 @@ lemma domain_defined : 𝚺₀-Function₁ (domain : V → V) via domainDef := b @[simp] lemma domain_defined_iff (v) : Semiformula.Evalbm V v domainDef.val ↔ v 0 = domain (v 1) := domain_defined.df.iff v -instance domain_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (domain : V → V) := Defined.to_definable _ domain_defined +instance domain_definable : 𝚺₀-Function₁ (domain : V → V) := Defined.to_definable _ domain_defined instance domain_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ (domain : V → V) := .of_zero domain_definable _ @@ -311,7 +311,7 @@ instance domain_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ (domain : V exact lt_of_le_of_lt (le_trans (le_pair_left i x) (le_log_of_mem hix)) (by simp [log_two_mul_add_one_of_pos (pos_of_nonempty hix)])) -instance : Bounded₁ ℒₒᵣ (domain : V → V) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ +instance : Bounded₁ (domain : V → V) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ lemma mem_domain_of_pair_mem {x y s : V} (h : ⟪x, y⟫ ∈ s) : x ∈ domain s := mem_domain_iff.mpr ⟨y, h⟩ @@ -366,7 +366,7 @@ private lemma range_graph {s' s : V} : s' = range s ↔ ∀ y < s' + s, (y ∈ s exact h y (lt_of_lt_of_le (lt_of_mem_rng hx) (by simp)) |>.mpr ⟨x, lt_of_mem_dom hx, _, hx, rfl⟩⟩ -def _root_.LO.FirstOrder.Arith.rangeDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.rangeDef : 𝚺₀.Semisentence 2 := .mkSigma “s' s | ∀ y < s' + s, (y ∈ s' ↔ ∃ x < s, ∃ z ∈' s, !pairDef z x y)” (by simp) lemma range_defined : 𝚺₀-Function₁ (range : V → V) via rangeDef := by @@ -375,7 +375,7 @@ lemma range_defined : 𝚺₀-Function₁ (range : V → V) via rangeDef := by @[simp] lemma range_defined_iff (v) : Semiformula.Evalbm V v rangeDef.val ↔ v 0 = range (v 1) := range_defined.df.iff v -instance range_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (range : V → V) := Defined.to_definable _ range_defined +instance range_definable : 𝚺₀-Function₁ (range : V → V) := Defined.to_definable _ range_defined @[definability, simp] instance range_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ (range : V → V) := .of_zero range_definable _ @@ -418,7 +418,7 @@ private lemma isMapping_iff {m : V} : rcases h x hx with ⟨y, _, hxy, h⟩ exact ExistsUnique.intro y hxy (fun y' hxy' ↦ h y' (lt_of_mem_rng hxy') hxy')⟩ -def _root_.LO.FirstOrder.Arith.isMappingDef : 𝚺₀-Semisentence 1 := .mkSigma +def _root_.LO.FirstOrder.Arith.isMappingDef : 𝚺₀.Semisentence 1 := .mkSigma “m | ∃ d <⁺ 2 * m, !domainDef d m ∧ ∀ x ∈' d, ∃ y < m, x ~[m] y ∧ ∀ y' < m, x ~[m] y' → y' = y” (by simp) lemma isMapping_defined : 𝚺₀-Predicate (IsMapping : V → Prop) via isMappingDef := by @@ -593,7 +593,7 @@ def fstIdx (p : V) : V := π₁ (p - 1) @[simp] lemma fstIdx_le_self (p : V) : fstIdx p ≤ p := le_trans (by simp [fstIdx]) (show p - 1 ≤ p by simp) -def _root_.LO.FirstOrder.Arith.fstIdxDef : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.fstIdxDef : 𝚺₀.Semisentence 2 := .mkSigma “n p | ∃ p' <⁺ p, !subDef p' p 1 ∧ !pi₁Def n p'” (by simp) lemma fstIdx_defined : 𝚺₀-Function₁ (fstIdx : V → V) via fstIdxDef := by @@ -617,7 +617,7 @@ def sndIdx (p : V) : V := π₂ (p - 1) @[simp] lemma sndIdx_le_self (p : V) : sndIdx p ≤ p := le_trans (by simp [sndIdx]) (show p - 1 ≤ p by simp) -def _root_.LO.FirstOrder.Arith.sndIdxDef : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.sndIdxDef : 𝚺₀.Semisentence 2 := .mkSigma “n p | ∃ p' <⁺ p, !subDef p' p 1 ∧ !pi₂Def n p'” (by simp) lemma sndIdx_defined : 𝚺₀-Function₁ (sndIdx : V → V) via sndIdxDef := by diff --git a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean index f969aad..4a5ff9d 100644 --- a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean +++ b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean @@ -136,7 +136,7 @@ lemma mem_limSeq_succ_iff {x s : M} : x ∈ c.limSeq v (s + 1) ↔ x ≤ s ∧ c.Φ v {z | z ∈ c.limSeq v s} x := by simp [limSeq_succ, mem_succ_iff] lemma limSeq_cumulative {s s' : M} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' := by - induction s' using induction_iSigmaOne generalizing s + induction s' using induction_sigma1 generalizing s · apply Definable.ball_le' (by definability) apply Definable.comp₂_infer · exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ @@ -153,7 +153,7 @@ lemma limSeq_cumulative {s s' : M} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' lemma mem_limSeq_self [c.StrongFinite] {u s : M} : u ∈ c.limSeq v s → u ∈ c.limSeq v (u + 1) := by - induction u using order_induction_piOne generalizing s + induction u using order_induction_pi1 generalizing s · apply Definable.all apply Definable.imp · apply Definable.comp₂_infer (by definability) diff --git a/Arithmetization/ISigmaOne/HFS/PRF.lean b/Arithmetization/ISigmaOne/HFS/PRF.lean index f85c843..247d0a9 100644 --- a/Arithmetization/ISigmaOne/HFS/PRF.lean +++ b/Arithmetization/ISigmaOne/HFS/PRF.lean @@ -90,7 +90,7 @@ lemma unique {s₁ s₂ : M} (H₁ : c.CSeq v s₁) (H₂ : c.CSeq v s₂) (h₁ suffices ∀ z₁ < s₁, ∀ z₂ < s₂, ⟪i, z₁⟫ ∈ s₁ → ⟪i, z₂⟫ ∈ s₂ → z₁ = z₂ by intro z₁ z₂ hz₁ hz₂; exact this z₁ (lt_of_mem_rng hz₁) z₂ (lt_of_mem_rng hz₂) hz₁ hz₂ intro z₁ hz₁ z₂ hz₂ h₁ h₂ - induction i using induction_iSigmaOne generalizing z₁ z₂ + induction i using induction_sigma1 generalizing z₁ z₂ · definability case zero => have : z₁ = c.zero v := H₁.seq.isMapping.uniq h₁ H₁.zero @@ -134,7 +134,7 @@ lemma CSeq.successor {s l z : M} (Hs : c.CSeq v s) (hl : l + 1 = lh s) (hz : ⟪ variable (c v) lemma CSeq.exists (l : M) : ∃ s, c.CSeq v s ∧ l + 1 = lh s := by - induction l using induction_iSigmaOne + induction l using induction_sigma1 · apply Definable.ex apply Definable.and · exact ⟨p.cseqDef.rew (Rew.embSubsts <| #0 :> fun i ↦ &(v i)), by diff --git a/Arithmetization/ISigmaOne/HFS/Seq.lean b/Arithmetization/ISigmaOne/HFS/Seq.lean index 3f8adea..6c5373f 100644 --- a/Arithmetization/ISigmaOne/HFS/Seq.lean +++ b/Arithmetization/ISigmaOne/HFS/Seq.lean @@ -26,7 +26,7 @@ private lemma seq_iff (s : M) : Seq s ↔ IsMapping s ∧ ∃ l ≤ 2 * s, ∃ d _ ≤ 2 * s := by simp), ⟨domain s , by simp, rfl, h⟩⟩, by rintro ⟨hs, l, _, _, _, rfl, h⟩; exact ⟨hs, l, h⟩⟩ -def _root_.LO.FirstOrder.Arith.seqDef : 𝚺₀-Semisentence 1 := .mkSigma +def _root_.LO.FirstOrder.Arith.seqDef : 𝚺₀.Semisentence 1 := .mkSigma “s | !isMappingDef s ∧ ∃ l <⁺ 2 * s, ∃ d <⁺ 2 * s, !domainDef d s ∧ !underDef d l” (by simp) lemma seq_defined : 𝚺₀-Predicate (Seq : M → Prop) via seqDef := by @@ -83,7 +83,7 @@ private lemma lh_graph (l s : M) : l = lh s ↔ (Seq s → ∃ d ≤ 2 * s, d = · rcases h Hs with ⟨_, _, rfl, h⟩; simpa [h] using Hs.domain_eq · simp [lh_prop_of_not_seq Hs, hn Hs]⟩ -def _root_.LO.FirstOrder.Arith.lhDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.lhDef : 𝚺₀.Semisentence 2 := .mkSigma “l s | (!seqDef s → ∃ d <⁺ 2 * s, !domainDef d s ∧ !underDef d l) ∧ (¬!seqDef s → l = 0)” (by simp) lemma lh_defined : 𝚺₀-Function₁ (lh : M → M) via lhDef := by @@ -96,7 +96,7 @@ instance lh_definable : 𝚺₀-Function₁ (lh : M → M) := Defined.to_definab instance lh_definable' (Γ) : Γ-Function₁ (lh : M → M) := .of_zero lh_definable _ -instance : Bounded₁ ℒₒᵣ (lh : M → M) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ +instance : Bounded₁ (lh : M → M) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ lemma Seq.exists {s : M} (h : Seq s) {x : M} (hx : x < lh s) : ∃ y, ⟪x, y⟫ ∈ s := h.isMapping x (by simpa [h.domain_eq] using hx) |>.exists @@ -137,7 +137,7 @@ lemma Seq.znth_eq_of_mem {s i : M} (h : Seq s) (hi : ⟪i, x⟫ ∈ s) : znth s lemma znth_prop_not {s i : M} (h : ¬Seq s ∨ lh s ≤ i) : znth s i = 0 := Classical.choose!_spec (znth_existsUnique s i) |>.2 (by simpa [-not_and, not_and_or] using h) -def _root_.LO.FirstOrder.Arith.znthDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.znthDef : 𝚺₀.Semisentence 3 := .mkSigma “x s i | ∃ l <⁺ 2 * s, !lhDef l s ∧ (:Seq s ∧ i < l → i ~[s] x) ∧ (¬(:Seq s ∧ i < l) → x = 0)” (by simp) private lemma znth_graph {x s i : M} : x = znth s i ↔ ∃ l ≤ 2 * s, l = lh s ∧ (Seq s ∧ i < l → ⟪i, x⟫ ∈ s) ∧ (¬(Seq s ∧ i < l) → x = 0) := by @@ -211,7 +211,7 @@ lemma seqCons_graph (t x s : M) : le_trans (pair_le_pair_left (by simp) x) (pair_polybound (2 * s) x), rfl, by rfl⟩, by rintro ⟨l, _, rfl, p, _, rfl, rfl⟩; rfl⟩ -def _root_.LO.FirstOrder.Arith.seqConsDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.seqConsDef : 𝚺₀.Semisentence 3 := .mkSigma “t s x | ∃ l <⁺ 2 * s, !lhDef l s ∧ ∃ p <⁺ (2 * s + x + 1)², !pairDef p l x ∧ !insertDef t p s” (by simp) lemma seqCons_defined : 𝚺₀-Function₂ (seqCons : M → M → M) via seqConsDef := by @@ -320,7 +320,7 @@ theorem seq_induction (Γ) {P : M → Prop} (hP : DefinablePred ℒₒᵣ (Γ, 1 (hnil : P ∅) (hcons : ∀ s x, Seq s → P s → P (s ⁀' x)) : ∀ {s : M}, Seq s → P s := by intro s sseq - induction s using order_induction_h_iSigmaOne + induction s using order_induction_h_sigma1 · exact Γ · definability case ind s ih => @@ -352,7 +352,7 @@ def vecConsUnexpander : Lean.PrettyPrinter.Unexpander section -def _root_.LO.FirstOrder.Arith.mkSeq₁Def : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.mkSeq₁Def : 𝚺₀.Semisentence 2 := .mkSigma “s x | !seqConsDef s 0 x” (by simp) lemma mkSeq₁_defined : 𝚺₀-Function₁ (fun x : M ↦ !⟦x⟧) via mkSeq₁Def := by @@ -436,7 +436,7 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function ⟪0, y⟫ ∈ W ∧ ∀ l < k, ∀ m < W, ∀ m' < W, ⟪l, m⟫ ∈ W → ⟪l + 1, m'⟫ ∈ W → ∀ x' ≤ x - l, ∀ y' ≤ m, f x' y' ≤ m' := by intro k hk - induction k using induction_iSigmaOne + induction k using induction_sigma1 · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) apply Definable.ex apply Definable.and (Definable.comp₁_infer (DefinableFunction.var _)) @@ -481,7 +481,7 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function rcases this x (by rfl) with ⟨W, SW, hxW, hW₀, hWₛ⟩ have : ∀ i ≤ x, ∀ m < W, ⟪x - i, m⟫ ∈ W → ∀ x' ≤ i, ∀ y' ≤ m, P x' y' := by intro i - induction i using induction_iSigmaOne + induction i using induction_sigma1 · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) apply Definable.ball_lt (DefinableFunction.const _) apply Definable.imp diff --git a/Arithmetization/ISigmaOne/HFS/Supplemental.lean b/Arithmetization/ISigmaOne/HFS/Supplemental.lean index fb26675..1755d7f 100644 --- a/Arithmetization/ISigmaOne/HFS/Supplemental.lean +++ b/Arithmetization/ISigmaOne/HFS/Supplemental.lean @@ -113,7 +113,7 @@ instance seqExp_definable : 𝚺₁-Function₂ (seqExp : M → M → M) := Defi .of_sigmaOne seqExp_definable _ _ lemma mem_seqExp_iff {s a k : M} : s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ (∀ i z, ⟪i, z⟫ ∈ s → z ∈ a) := by - induction k using induction_iPiOne generalizing s + induction k using induction_pi1 generalizing s · suffices 𝚷₁-Predicate fun {k} => ∀ {s : M}, s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ ∀ i < s, ∀ z < s, ⟪i, z⟫ ∈ s → z ∈ a by exact this.of_iff (fun k ↦ forall_congr' <| fun s ↦ iff_congr (by rfl) <| and_congr (by rfl) <| and_congr (by rfl) diff --git a/Arithmetization/ISigmaOne/HFS/Vec.lean b/Arithmetization/ISigmaOne/HFS/Vec.lean index ac4996e..9897c92 100644 --- a/Arithmetization/ISigmaOne/HFS/Vec.lean +++ b/Arithmetization/ISigmaOne/HFS/Vec.lean @@ -64,7 +64,7 @@ lemma cons_le_cons {x₁ x₂ v₁ v₂ : V} (hx : x₁ ≤ x₂) (hv : v₁ ≤ section -def _root_.LO.FirstOrder.Arith.consDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.consDef : 𝚺₀.Semisentence 3 := .mkSigma “w x v | ∃ xv < w, !pairDef xv x v ∧ w = xv + 1” (by simp) lemma cons_defined : 𝚺₀-Function₂ (cons : V → V → V) via consDef := by @@ -80,7 +80,7 @@ instance cons_definable : 𝚺₀-Function₂ (cons : V → V → V) := Defined. instance cons_definable' (Γ) : Γ-Function₂ (cons : V → V → V) := .of_zero cons_definable _ -def _root_.LO.FirstOrder.Arith.mkVec₁Def : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.mkVec₁Def : 𝚺₀.Semisentence 2 := .mkSigma “s x | !consDef s x 0” (by simp) lemma mkVec₁_defined : 𝚺₀-Function₁ (fun x : V ↦ ?[x]) via mkVec₁Def := by @@ -200,7 +200,7 @@ lemma graph_succ {v i x : V} : lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by suffices ∀ i' ≤ i, ∀ v' ≤ v, ∃ x, Graph ⟪v', i', x⟫ from this i (by simp) v (by simp) intro i' hi' - induction i' using induction_iSigmaOne + induction i' using induction_sigma1 · definability case zero => intro v' _ @@ -211,7 +211,7 @@ lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by exact ⟨x, graph_case.mpr <| Or.inr ⟨v', i', x, rfl, hx⟩⟩ lemma graph_unique {v i x₁ x₂ : V} : Graph ⟪v, i, x₁⟫ → Graph ⟪v, i, x₂⟫ → x₁ = x₂ := by - induction i using induction_iPiOne generalizing v x₁ x₂ + induction i using induction_pi1 generalizing v x₁ x₂ · definability case zero => simp [graph_zero] @@ -307,13 +307,13 @@ lemma pi₁_zero : π₁ (0 : V) = 0 := nonpos_iff_eq_zero.mp (pi₁_le_self 0) lemma pi₂_zero : π₂ (0 : V) = 0 := nonpos_iff_eq_zero.mp (pi₂_le_self 0) @[simp] lemma nth_zero_idx (i : V) : (0).[i] = 0 := by - induction i using induction_iSigmaOne + induction i using induction_sigma1 · definability case zero => simp [nth_zero, fstIdx, pi₁_zero] case succ i ih => simp [nth_succ, sndIdx, pi₂_zero, ih] lemma nth_lt_of_pos {v} (hv : 0 < v) (i : V) : v.[i] < v := by - induction i using induction_iPiOne generalizing v + induction i using induction_pi1 generalizing v · definability case zero => rcases zero_or_succ v with (rfl | ⟨v, rfl⟩) @@ -626,7 +626,7 @@ theorem sigmaOne_skolem_vec {R : V → V → Prop} (hP : 𝚺₁-Relation R) {l} (H : ∀ x < l, ∃ y, R x y) : ∃ v, len v = l ∧ ∀ i < l, R i v.[i] := by have : ∀ k ≤ l, ∃ v, len v = k ∧ ∀ i < k, R (l - k + i) v.[i] := by intro k hk - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => exact ⟨0, by simp⟩ case succ k ih => @@ -942,7 +942,7 @@ instance repeatVec_definable : 𝚺₁-Function₂ (repeatVec : V → V → V) : end @[simp] lemma len_repeatVec (x k : V) : len (repeatVec x k) = k := by - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simp case succ k ih => simp [ih] @@ -951,7 +951,7 @@ end simpa using len_le (repeatVec x k) lemma nth_repeatVec (x k : V) {i} (h : i < k) : (repeatVec x k).[i] = x := by - induction k using induction_iSigmaOne generalizing i + induction k using induction_sigma1 generalizing i · definability case zero => simp at h case succ k ih => diff --git a/Arithmetization/ISigmaOne/Ind.lean b/Arithmetization/ISigmaOne/Ind.lean index 7458dd9..14496dc 100644 --- a/Arithmetization/ISigmaOne/Ind.lean +++ b/Arithmetization/ISigmaOne/Ind.lean @@ -19,7 +19,7 @@ variable (m : ℕ) [Fact (1 ≤ m)] [V ⊧ₘ* 𝐈𝐍𝐃𝚺 m] lemma induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ : (𝚷, m)-Predicate Q) (zero : P 0 ∨ Q 0) (succ : ∀ x, P x ∨ Q x → P (x + 1) ∨ Q (x + 1)) : ∀ x, P x ∨ Q x := by - haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) intro a have : ∃ p < exp (a + 1), ∀ x ≤ a, x ∈ p ↔ P x := by simpa [lt_succ_iff_le] using finset_comprehension hP (a + 1) @@ -29,7 +29,7 @@ lemma induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ rcases this with ⟨q, _, hq⟩ have : ∀ x ≤ a, x ∈ p ∨ x ∈ q := by intro x hx - induction x using induction_iSigmaOne + induction x using induction_sigma1 · clear hp hq zero succ definability case zero => simpa [hp, hq] using zero @@ -42,7 +42,7 @@ lemma induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ : (𝚷, m)-Predicate Q) (ind : ∀ x, (∀ y < x, P y ∨ Q y) → P x ∨ Q x) : ∀ x, P x ∨ Q x := by - haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_iSigma_of_le (show 1 ≤ m from Fact.out) + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) intro a have : ∃ p < exp (a + 1), ∀ x ≤ a, x ∈ p ↔ P x := by simpa [lt_succ_iff_le] using finset_comprehension hP (a + 1) @@ -52,7 +52,7 @@ lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P rcases this with ⟨q, _, hq⟩ have : ∀ x ≤ a, x ∈ p ∨ x ∈ q := by intro x hx - induction x using order_induction_iSigmaOne + induction x using order_induction_sigma1 · clear hp hq ind apply LO.FirstOrder.Arith.Definable.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index 57e7a15..58edbaf 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -56,28 +56,28 @@ scoped notation "^∃ " p:64 => qqEx 0 p section -def _root_.LO.FirstOrder.Arith.qqRelDef : 𝚺₀-Semisentence 5 := +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.qqNRelDef : 𝚺₀-Semisentence 5 := +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.qqVerumDef : 𝚺₀-Semisentence 2 := +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.qqFalsumDef : 𝚺₀-Semisentence 2 := +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.qqAndDef : 𝚺₀-Semisentence 4 := +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.qqOrDef : 𝚺₀-Semisentence 4 := +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.qqAllDef : 𝚺₀-Semisentence 3 := +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.qqExDef : 𝚺₀-Semisentence 3 := +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) lemma ss (x : V) : x < x + 1 := by exact lt_add_one x @@ -268,7 +268,7 @@ private lemma phi_iff (C p : V) : · left; exact ⟨n, q, hq, rfl⟩ · right; exact ⟨n, q, hq, rfl⟩ -def formulaAux : 𝚺₀-Semisentence 2 := .mkSigma +def formulaAux : 𝚺₀.Semisentence 2 := .mkSigma “p C | (∃ n < p, !qqVerumDef p n) ∨ (∃ n < p, !qqFalsumDef p n) ∨ @@ -573,7 +573,7 @@ lemma Language.Semiformula.induction_sigma₁ {P : V → V → Prop} (hP : 𝚺 ∀ n p, L.Semiformula n p → P n p := Language.Semiformula.induction 𝚺 hP hrel hnrel hverum hfalsum hand hor hall hex -lemma Language.Semiformula.induction_pi₁ {P : V → V → Prop} (hP : 𝚷₁-Relation P) +lemma Language.Semiformula.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]) diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index 031b711..726c28e 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -546,7 +546,7 @@ lemma substs_not_uformula {m w x} (h : ¬L.UFormula x) : 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 revert m w - apply Language.Semiformula.induction_pi₁ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp · definability · intros; simp [*] · intros; simp [*] @@ -564,7 +564,7 @@ lemma substs_neg {p} (hp : L.Semiformula n p) : 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 revert m w - apply Language.Semiformula.induction_pi₁ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp · definability · intro n k R v hR hv m w hw simp only [substs_rel, Language.SemitermVec.termSubstVec, shift_rel, @@ -604,7 +604,7 @@ lemma shift_substs {p} (hp : L.Semiformula n p) : 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 revert m w n v - apply Language.Semiformula.induction_pi₁ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp · apply Definable.all apply Definable.all apply Definable.all @@ -648,7 +648,7 @@ lemma substs_substs {p} (hp : L.Semiformula l p) : 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 revert w - apply Language.Semiformula.induction_pi₁ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp + apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp · definability · intro n k R v hR hv w _ H simp only [substs_rel, qqRel_inj, true_and, hR, hv] diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean index efaba03..523857d 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean @@ -198,14 +198,14 @@ instance substItr_definable₁' (n w p : V) (Γ m) : (Γ, m + 1)-Function₁ (su end @[simp] lemma len_substItr (n w p k : V) : len (substItr n w p k) = k := by - induction k using induction_iSigmaOne + 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 - induction k using induction_iSigmaOne generalizing i + induction k using induction_sigma1 generalizing i · definability case zero => simp at hi case succ k ih => @@ -216,7 +216,7 @@ end 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simp case succ k ih => @@ -231,7 +231,7 @@ lemma neg_conj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simp case succ k ih => @@ -246,7 +246,7 @@ lemma neg_disj_substItr {n w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula (n + 1) 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simp case succ k ih => @@ -263,7 +263,7 @@ lemma substs_conj_substItr {n m l w p k : V} (hp : ⌜ℒₒᵣ⌝.Semiformula ( 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simp case succ k ih => diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean index 810f323..1e8a59a 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean @@ -151,7 +151,7 @@ def root (s p : V) : V := ⟪s, 9, p⟫ + 1 section -def _root_.LO.FirstOrder.Arith.axLDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.axLDef : 𝚺₀.Semisentence 3 := .mkSigma “y s p | ∃ y' < y, !pair₃Def y' s 0 p ∧ y = y' + 1” (by simp) lemma axL_defined : 𝚺₀-Function₂ (axL : V → V → V) via axLDef := by @@ -163,7 +163,7 @@ lemma axL_defined : 𝚺₀-Function₂ (axL : V → V → V) via axLDef := by @[simp] lemma eval_axLDef (v) : Semiformula.Evalbm V v axLDef.val ↔ v 0 = axL (v 1) (v 2) := axL_defined.df.iff v -def _root_.LO.FirstOrder.Arith.verumIntroDef : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.verumIntroDef : 𝚺₀.Semisentence 2 := .mkSigma “y s | ∃ y' < y, !pair₃Def y' s 1 0 ∧ y = y' + 1” (by simp) lemma verumIntro_defined : 𝚺₀-Function₁ (verumIntro : V → V) via verumIntroDef := by @@ -175,7 +175,7 @@ lemma verumIntro_defined : 𝚺₀-Function₁ (verumIntro : V → V) via verumI @[simp] lemma eval_verumIntroDef (v) : Semiformula.Evalbm V v verumIntroDef.val ↔ v 0 = verumIntro (v 1) := verumIntro_defined.df.iff v -def _root_.LO.FirstOrder.Arith.andIntroDef : 𝚺₀-Semisentence 6 := +def _root_.LO.FirstOrder.Arith.andIntroDef : 𝚺₀.Semisentence 6 := .mkSigma “y s p q dp dq | ∃ y' < y, !pair₆Def y' s 2 p q dp dq ∧ y = y' + 1” (by simp) lemma andIntro_defined : 𝚺₀-Function₅ (andIntro : V → V → V → V → V → V) via andIntroDef := by @@ -187,7 +187,7 @@ lemma andIntro_defined : 𝚺₀-Function₅ (andIntro : V → V → V → V → @[simp] lemma eval_andIntroDef (v) : Semiformula.Evalbm V v andIntroDef.val ↔ v 0 = andIntro (v 1) (v 2) (v 3) (v 4) (v 5) := andIntro_defined.df.iff v -def _root_.LO.FirstOrder.Arith.orIntroDef : 𝚺₀-Semisentence 5 := +def _root_.LO.FirstOrder.Arith.orIntroDef : 𝚺₀.Semisentence 5 := .mkSigma “y s p q d | ∃ y' < y, !pair₅Def y' s 3 p q d ∧ y = y' + 1” (by simp) lemma orIntro_defined : 𝚺₀-Function₄ (orIntro : V → V → V → V → V) via orIntroDef := by @@ -199,7 +199,7 @@ lemma orIntro_defined : 𝚺₀-Function₄ (orIntro : V → V → V → V → V @[simp] lemma eval_orIntroDef (v) : Semiformula.Evalbm V v orIntroDef.val ↔ v 0 = orIntro (v 1) (v 2) (v 3) (v 4) := orIntro_defined.df.iff v -def _root_.LO.FirstOrder.Arith.allIntroDef : 𝚺₀-Semisentence 4 := +def _root_.LO.FirstOrder.Arith.allIntroDef : 𝚺₀.Semisentence 4 := .mkSigma “y s p d | ∃ y' < y, !pair₄Def y' s 4 p d ∧ y = y' + 1” (by simp) lemma allIntro_defined : 𝚺₀-Function₃ (allIntro : V → V → V → V) via allIntroDef := by @@ -211,7 +211,7 @@ lemma allIntro_defined : 𝚺₀-Function₃ (allIntro : V → V → V → V) vi @[simp] lemma eval_allIntroDef (v) : Semiformula.Evalbm V v allIntroDef.val ↔ v 0 = allIntro (v 1) (v 2) (v 3) := allIntro_defined.df.iff v -def _root_.LO.FirstOrder.Arith.exIntroDef : 𝚺₀-Semisentence 5 := +def _root_.LO.FirstOrder.Arith.exIntroDef : 𝚺₀.Semisentence 5 := .mkSigma “y s p t d | ∃ y' < y, !pair₅Def y' s 5 p t d ∧ y = y' + 1” (by simp) lemma exIntro_defined : 𝚺₀-Function₄ (exIntro : V → V → V → V → V) via exIntroDef := by @@ -223,7 +223,7 @@ lemma exIntro_defined : 𝚺₀-Function₄ (exIntro : V → V → V → V → V @[simp] lemma eval_exIntroDef (v) : Semiformula.Evalbm V v exIntroDef.val ↔ v 0 = exIntro (v 1) (v 2) (v 3) (v 4) := exIntro_defined.df.iff v -def _root_.LO.FirstOrder.Arith.wkRuleDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.wkRuleDef : 𝚺₀.Semisentence 3 := .mkSigma “y s d | ∃ y' < y, !pair₃Def y' s 6 d ∧ y = y' + 1” (by simp) lemma wkRule_defined : 𝚺₀-Function₂ (wkRule : V → V → V) via wkRuleDef := by @@ -235,7 +235,7 @@ lemma wkRule_defined : 𝚺₀-Function₂ (wkRule : V → V → V) via wkRuleDe @[simp] lemma eval_wkRuleDef (v) : Semiformula.Evalbm V v wkRuleDef.val ↔ v 0 = wkRule (v 1) (v 2) := wkRule_defined.df.iff v -def _root_.LO.FirstOrder.Arith.shiftRuleDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.shiftRuleDef : 𝚺₀.Semisentence 3 := .mkSigma “y s d | ∃ y' < y, !pair₃Def y' s 7 d ∧ y = y' + 1” (by simp) lemma shiftRule_defined : 𝚺₀-Function₂ (shiftRule : V → V → V) via shiftRuleDef := by @@ -247,7 +247,7 @@ lemma shiftRule_defined : 𝚺₀-Function₂ (shiftRule : V → V → V) via sh @[simp] lemma eval_shiftRuleDef (v) : Semiformula.Evalbm V v shiftRuleDef.val ↔ v 0 = shiftRule (v 1) (v 2) := shiftRule_defined.df.iff v -def _root_.LO.FirstOrder.Arith.cutRuleDef : 𝚺₀-Semisentence 5 := +def _root_.LO.FirstOrder.Arith.cutRuleDef : 𝚺₀.Semisentence 5 := .mkSigma “y s p d₁ d₂ | ∃ y' < y, !pair₅Def y' s 8 p d₁ d₂ ∧ y = y' + 1” (by simp) lemma cutRule_defined : 𝚺₀-Function₄ (cutRule : V → V → V → V → V) via cutRuleDef := by @@ -259,7 +259,7 @@ lemma cutRule_defined : 𝚺₀-Function₄ (cutRule : V → V → V → V → V @[simp] lemma eval_cutRuleDef (v) : Semiformula.Evalbm V v cutRuleDef.val ↔ v 0 = cutRule (v 1) (v 2) (v 3) (v 4) := cutRule_defined.df.iff v -def _root_.LO.FirstOrder.Arith.rootDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.rootDef : 𝚺₀.Semisentence 3 := .mkSigma “y s p | ∃ y' < y, !pair₃Def y' s 9 p ∧ y = y' + 1” (by simp) lemma root_defined : 𝚺₀-Function₂ (root : V → V → V) via rootDef := by @@ -787,7 +787,7 @@ lemma conj (ps : V) {s} (hs : L.FormulaSet 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simpa using verum (by simp [hs]) (by simp) case succ k ih => @@ -801,7 +801,7 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( have : ∀ k ≤ len ps, ∀ s' ≤ vecToSet ps, s' ⊆ vecToSet ps → (∀ i < len ps - k, ps.[i] ∈ s') → T.Derivable (insert (^⋁ takeLast ps k) (s' ∪ s)) := by intro k hk - induction k using induction_iSigmaOne + induction k using induction_sigma1 · apply Definable.imp (by definability) apply Definable.ball_le (by definability) apply Definable.imp (by definability) diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean index 09374cf..4a9e30c 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean @@ -99,7 +99,7 @@ lemma conj (ps : V) {s} (hs : L.FormulaSet 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 - induction k using induction_iSigmaOne + induction k using induction_sigma1 · definability case zero => simpa using verum (by simp [hs]) (by simp) case succ k ih => @@ -113,7 +113,7 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( have : ∀ k ≤ len ps, ∀ s' ≤ vecToSet ps, s' ⊆ vecToSet ps → (∀ i < len ps - k, ps.[i] ∈ s') → T.Derivable (insert (^⋁ takeLast ps k) (s' ∪ s)) := by intro k hk - induction k using induction_iSigmaOne + induction k using induction_sigma1 · apply Definable.imp (by definability) apply Definable.ball_le (by definability) apply Definable.imp (by definability) diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean index 334a8cf..bbcb15e 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean @@ -47,7 +47,7 @@ lemma lt_qqFunc_of_mem {i b k f v : V} (hi : ⟪i, b⟫ ∈ v) : b < ^func k f v @[simp] lemma qqFunc_inj {k f v k' f' w : V} : ^func k f v = ^func k' f' w ↔ k = k' ∧ f = f' ∧ v = w := by simp [qqFunc] -def _root_.LO.FirstOrder.Arith.qqBvarDef : 𝚺₀-Semisentence 2 := .mkSigma “t z | ∃ t' < t, !pairDef t' 0 z ∧ t = t' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqBvarDef : 𝚺₀.Semisentence 2 := .mkSigma “t z | ∃ t' < t, !pairDef t' 0 z ∧ t = t' + 1” (by simp) lemma qqBvar_defined : 𝚺₀-Function₁ (qqBvar : V → V) via qqBvarDef := by intro v; simp [qqBvarDef] @@ -58,7 +58,7 @@ lemma qqBvar_defined : 𝚺₀-Function₁ (qqBvar : V → V) via qqBvarDef := b @[simp] lemma eval_qqBvarDef (v) : Semiformula.Evalbm V v qqBvarDef.val ↔ v 0 = ^#(v 1) := qqBvar_defined.df.iff v -def _root_.LO.FirstOrder.Arith.qqFvarDef : 𝚺₀-Semisentence 2 := .mkSigma “t x | ∃ t' < t, !pairDef t' 1 x ∧ t = t' + 1” (by simp) +def _root_.LO.FirstOrder.Arith.qqFvarDef : 𝚺₀.Semisentence 2 := .mkSigma “t x | ∃ t' < t, !pairDef t' 1 x ∧ t = t' + 1” (by simp) lemma qqFvar_defined : 𝚺₀-Function₁ (qqFvar : V → V) via qqFvarDef := by intro v; simp [qqFvarDef] @@ -77,7 +77,7 @@ private lemma qqFunc_graph {x k f v : V} : ⟪2, k, f, v⟫, by simp [qqFunc], rfl, rfl⟩, by rintro ⟨_, _, rfl, _, _, rfl, _, _, rfl, rfl⟩; rfl⟩ -def _root_.LO.FirstOrder.Arith.qqFuncDef : 𝚺₀-Semisentence 4 := .mkSigma +def _root_.LO.FirstOrder.Arith.qqFuncDef : 𝚺₀.Semisentence 4 := .mkSigma “x k f v | ∃ fv < x, !pairDef fv f v ∧ ∃ kfv < x, !pairDef kfv k fv ∧ ∃ x' < x, !pairDef x' 2 kfv ∧ x = x' + 1” (by simp) lemma qqFunc_defined : 𝚺₀-Function₃ (qqFunc : V → V → V → V) via qqFuncDef := by diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean index 90fdb35..00afeff 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean @@ -589,7 +589,7 @@ lemma numeralAux_defined : 𝚺₁-Function₁ (numeralAux : V → V) via numera end @[simp] lemma numeralAux_semiterm (n x : V) : ⌜ℒₒᵣ⌝.Semiterm n (numeralAux x) := by - induction x using induction_iSigmaOne + induction x using induction_sigma1 · definability case zero => simp case succ x ih => simp [qqAdd, ih] @@ -639,7 +639,7 @@ end @[simp] lemma numeral_substs {w : V} (hw : ⌜ℒₒᵣ⌝.SemitermVec n m w) (x : V) : ⌜ℒₒᵣ⌝.termSubst n m w (numeral x) = numeral x := by - induction x using induction_iSigmaOne + induction x using induction_sigma1 · definability case zero => simp [hw, Formalized.zero, qqFunc_absolute] case succ x ih => @@ -649,7 +649,7 @@ end @[simp] lemma numeral_shift (x : V) : ⌜ℒₒᵣ⌝.termShift n (numeral x) = numeral x := by - induction x using induction_iSigmaOne + induction x using induction_sigma1 · definability case zero => simp [Formalized.zero, qqFunc_absolute] case succ x ih => @@ -659,7 +659,7 @@ end @[simp] lemma numeral_bShift (x : V) : ⌜ℒₒᵣ⌝.termBShift n (numeral x) = numeral x := by - induction x using induction_iSigmaOne + induction x using induction_sigma1 · definability case zero => simp [Formalized.zero, qqFunc_absolute] case succ x ih => diff --git a/Arithmetization/ISigmaZero/Exponential/Exp.lean b/Arithmetization/ISigmaZero/Exponential/Exp.lean index 4043586..eb35f39 100644 --- a/Arithmetization/ISigmaZero/Exponential/Exp.lean +++ b/Arithmetization/ISigmaZero/Exponential/Exp.lean @@ -19,19 +19,19 @@ lemma ext_graph (a b c : M) : a = ext b c ↔ ∃ x ≤ c, x = c / b ∧ a = x % · rintro rfl; exact ⟨c / b, by simp, rfl, by rfl⟩ · rintro ⟨_, _, rfl, rfl⟩; simp -def _root_.LO.FirstOrder.Arith.extDef : 𝚺₀-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.extDef : 𝚺₀.Semisentence 3 := .mkSigma “a b c | ∃ x <⁺ c, !divDef x c b ∧ !remDef a x b” (by simp) lemma ext_defined : 𝚺₀-Function₂ (λ a b : M ↦ ext a b) via extDef := by intro v; simp [Matrix.vecHead, Matrix.vecTail, extDef, ext_graph, Semiformula.eval_substs, div_defined.df.iff, rem_defined.df.iff, le_iff_lt_succ] -instance ext_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ (ext : M → M → M) := Defined.to_definable _ ext_defined +instance ext_definable : 𝚺₀-Function₂ (ext : M → M → M) := Defined.to_definable _ ext_defined @[simp] lemma ext_le_add (u z : M) : ext u z ≤ z := le_trans (mod_le (z / u) u) (by simp [add_comm]) -instance : Bounded₂ ℒₒᵣ (ext : M → M → M) := ⟨#1, by intro v; simp⟩ +instance : Bounded₂ (ext : M → M → M) := ⟨#1, by intro v; simp⟩ @[simp] lemma ext_lt {u} (z : M) (pos : 0 < u) : ext u z < u := by simp [ext, pos] @@ -77,7 +77,7 @@ lemma Exponential.Seqₛ.iff (y X Y : M) : · exact Or.inl ⟨by simp [hx, hy], by simp [hx, hy]⟩ · exact Or.inr ⟨by simp [hx, hy], by simp [hx, hy]⟩⟩ -def Exponential.Seqₛ.def : 𝚺₀-Semisentence 3 := .mkSigma +def Exponential.Seqₛ.def : 𝚺₀.Semisentence 3 := .mkSigma “ y X Y | ∀ u <⁺ y, u ≠ 2 → !ppow2Def u → ( (∃ ext_u_X <⁺ X, !extDef ext_u_X u X ∧ !extDef (2 * ext_u_X) u² X) ∧ @@ -102,7 +102,7 @@ lemma Exponential.graph_iff (x y : M) : · exact Or.inl H · exact Or.inr ⟨X, bX, Y, bY, ⟨H₀.1.symm, H₀.2.symm⟩, Hₛ, ⟨u, hu, ne2, ppu, hX.symm, hY.symm⟩⟩⟩ -def _root_.LO.FirstOrder.Arith.exponentialDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.exponentialDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | (x = 0 ∧ y = 1) ∨ ∃ X <⁺ y⁴, ∃ Y <⁺ y⁴, (!extDef 1 4 X ∧ !extDef 2 4 Y) ∧ @@ -203,7 +203,7 @@ lemma Seqₛ.append {z x y X Y i : M} (h : Seqₛ z X Y) (ppi : PPow2 i) (hz : z lemma pow2_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : Pow2 (ext i Y) := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ei : i = 4 @@ -224,7 +224,7 @@ lemma range_pow2 {x y : M} (h : Exponential x y) : Pow2 y := by lemma le_sq_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : i ≤ (ext i Y)^2 := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ei : i = 4 @@ -245,7 +245,7 @@ example {a b c : ℕ} : a * (b * c) = b * (a * c) := by exact Nat.mul_left_comm lemma two_mul_ext_le_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 2 * ext i Y ≤ i := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ei : i = 4 @@ -427,7 +427,7 @@ lemma exponential_odd_two_mul_sq {x y : M} : Exponential (2 * x + 1) (2 * y ^ 2) lemma two_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 2 ≤ ext i Y := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ei : i = 4 @@ -448,7 +448,7 @@ lemma two_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) lemma ext_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : ext i X < ext i Y := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ne4 : i = 4 @@ -503,7 +503,7 @@ lemma exponential_succ {x y : M} : Exponential (x + 1) y ↔ ∃ z, y = 2 * z rintro z rfl exact not_exponential_of_le (le_trans le_two_mul_left $ by simpa using hxy) · revert x - induction y using order_induction_iSigmaZero + induction y using order_induction_sigma0 · definability case ind y IH => intro x hxy @@ -547,7 +547,7 @@ alias ⟨of_succ_two_mul, succ⟩ := exponential_succ_mul_two lemma one_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 1 ≤ ext i X := by - induction i using order_induction_iSigmaZero + induction i using order_induction_sigma0 · definability case ind i IH => by_cases ne4 : i = 4 @@ -582,7 +582,7 @@ protected lemma uniq {x y₁ y₂ : M} : Exponential x y₁ → Exponential x y revert x h y₁ suffices ∀ x < y₂, ∀ y₁ ≤ y₂, Exponential x y₁ → Exponential x y₂ → y₁ = y₂ by intro x y₁ h₁ h₂ hy; exact this x h₂.lt y₁ hy h₁ h₂ - induction y₂ using order_induction_iSigmaZero + induction y₂ using order_induction_sigma0 · definability case ind y₂ IH => intro x _ y₁ h h₁ h₂ @@ -599,7 +599,7 @@ protected lemma inj {x₁ x₂ y : M} : Exponential x₁ y → Exponential x₂ revert x₁ x₂ h₁ h₂ suffices ∀ x₁ < y, ∀ x₂ < y, Exponential x₁ y → Exponential x₂ y → x₁ = x₂ by intro x₁ x₂ h₁ h₂; exact this x₁ h₁.lt x₂ h₂.lt h₁ h₂ - induction y using order_induction_iSigmaZero + induction y using order_induction_sigma0 · definability case ind y IH => intro x₁ _ x₂ _ h₁ h₂ @@ -633,7 +633,7 @@ lemma monotone {x₁ x₂ y₁ y₂ : M} : Exponential x₁ y₁ → Exponential intro h₁ h₂; contrapose; simp intro hy exact this x₁ h₁.lt y₂ hy x₂ h₂.lt h₁ h₂ - induction y₁ using order_induction_iSigmaZero + induction y₁ using order_induction_sigma0 · definability case ind y₁ IH => intro x₁ _ y₂ hy x₂ _ h₁ h₂ @@ -671,7 +671,7 @@ lemma add_mul {x₁ x₂ y₁ y₂ : M} (h₁ : Exponential x₁ y₁) (h₂ : E revert y₂ suffices ∀ y₂ ≤ y₁, Exponential x₂ y₂ → Exponential (x₁ + x₂) (y₁ * y₂) by intro y₂ h₂ hy; exact this y₂ hy h₂ - induction x₂ using induction_iSigmaZero + induction x₂ using induction_sigma0 · definability case zero => intro y₂ _ h₂ @@ -693,7 +693,7 @@ variable [M ⊧ₘ* 𝐈𝚺₁] namespace Exponential lemma range_exists (x : M) : ∃ y, Exponential x y := by - induction x using induction_iSigmaOne + induction x using induction_sigma1 · apply Definable.ex definability case zero => exact ⟨1, by simp⟩ @@ -715,7 +715,7 @@ lemma exponential_exp (a : M) : Exponential a (exp a) := Classical.choose!_spec lemma exponential_graph {a b : M} : a = exp b ↔ Exponential b a := Classical.choose!_eq_iff _ -def _root_.LO.FirstOrder.Arith.expDef : 𝚺₀-Semisentence 2 := .mkSigma “x y | !exponentialDef.val y x” (by simp) +def _root_.LO.FirstOrder.Arith.expDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | !exponentialDef.val y x” (by simp) lemma exp_defined_deltaZero : 𝚺₀-Function₁ (Exp.exp : M → M) via expDef := by intro v; simp [expDef, exponential_graph] @@ -723,7 +723,7 @@ lemma exp_defined_deltaZero : 𝚺₀-Function₁ (Exp.exp : M → M) via expDef @[simp] lemma exp_defined_iff (v) : Semiformula.Evalbm M v expDef.val ↔ v 0 = Exp.exp (v 1) := exp_defined_deltaZero.df.iff v -instance exp_definable_deltaZero : DefinableFunction₁ ℒₒᵣ 𝚺₀ (Exp.exp : M → M) := Defined.to_definable _ exp_defined_deltaZero +instance exp_definable_deltaZero : 𝚺₀-Function₁ (Exp.exp : M → M) := Defined.to_definable _ exp_defined_deltaZero lemma exp_of_exponential {a b : M} (h : Exponential a b) : exp a = b := Eq.symm <| exponential_graph.mpr h diff --git a/Arithmetization/ISigmaZero/Exponential/Log.lean b/Arithmetization/ISigmaZero/Exponential/Log.lean index 609c420..ab4fc4d 100644 --- a/Arithmetization/ISigmaZero/Exponential/Log.lean +++ b/Arithmetization/ISigmaZero/Exponential/Log.lean @@ -65,7 +65,7 @@ lemma log_lt_self_of_pos {y : M} (pos : 0 < y) : log y < y := lemma log_graph {x y : M} : x = log y ↔ (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y') := Classical.choose!_eq_iff _ -def _root_.LO.FirstOrder.Arith.logDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.logDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' <⁺ y, !exponentialDef x y' ∧ y < 2 * y')” (by simp) lemma log_defined : 𝚺₀-Function₁ (log : M → M) via logDef := by @@ -74,9 +74,9 @@ lemma log_defined : 𝚺₀-Function₁ (log : M → M) via logDef := by @[simp] lemma log_defined_iff (v) : Semiformula.Evalbm M v logDef.val ↔ v 0 = log (v 1) := log_defined.df.iff v -instance log_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (log : M → M) := Defined.to_definable _ log_defined +instance log_definable : 𝚺₀-Function₁ (log : M → M) := Defined.to_definable _ log_defined -instance : Bounded₁ ℒₒᵣ (log : M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₁ (log : M → M) := ⟨#0, λ _ ↦ by simp⟩ lemma log_eq_of_pos {x y : M} (pos : 0 < y) {y'} (H : Exponential x y') (hy' : y' ≤ y) (hy : y < 2 * y') : log y = x := (log_exists_unique_pos pos).unique ⟨log_lt_self_of_pos pos, log_pos pos⟩ ⟨lt_of_lt_of_le H.lt hy', y', hy', H, hy⟩ @@ -160,7 +160,7 @@ lemma length_graph {i a : M} : i = ‖a‖ ↔ (0 < a → ∃ k ≤ a, k = log a · rintro rfl; exact ⟨log a, by simp⟩ · rintro ⟨_, _, rfl, rfl⟩; rfl -def _root_.LO.FirstOrder.Arith.lengthDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.lengthDef : 𝚺₀.Semisentence 2 := .mkSigma “i a | (0 < a → ∃ k <⁺ a, !logDef k a ∧ i = k + 1) ∧ (a = 0 → i = 0)” (by simp) lemma length_defined : 𝚺₀-Function₁ (‖·‖ : M → M) via lengthDef := by @@ -169,9 +169,9 @@ lemma length_defined : 𝚺₀-Function₁ (‖·‖ : M → M) via lengthDef := @[simp] lemma length_defined_iff (v) : Semiformula.Evalbm M v lengthDef.val ↔ v 0 = ‖v 1‖ := length_defined.df.iff v -instance length_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (‖·‖ : M → M) := Defined.to_definable _ length_defined +instance length_definable : 𝚺₀-Function₁ (‖·‖ : M → M) := Defined.to_definable _ length_defined -instance : Bounded₁ ℒₒᵣ (‖·‖ : M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₁ (‖·‖ : M → M) := ⟨#0, λ _ ↦ by simp⟩ @[simp] lemma length_one : ‖(1 : M)‖ = 1 := by simp [length_eq_binaryLength] @@ -285,7 +285,7 @@ lemma brange_exists_unique (a : M) : ∀ x < ‖a‖, ∃! y, Exponential x y := intro x hx; rcases this x hx with ⟨_, _, H⟩ exact ExistsUnique.intro _ H (fun y' H' ↦ H'.uniq H) intro x - induction x using induction_iSigmaZero + induction x using induction_sigma0 · definability case zero => intro ha @@ -335,7 +335,7 @@ lemma bexp_graph {y a x : M} : y = bexp a x ↔ ∃ l ≤ a, l = ‖a‖ ∧ (x · exact (hlt lt).uniq (exp_bexp_of_lt lt) · rcases hle le; simp [bexp_eq_zero_of_le le]⟩ -def _root_.LO.FirstOrder.Arith.bexpDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.bexpDef : 𝚺₀.Semisentence 3 := .mkSigma “y a x | ∃ l <⁺ a, !lengthDef l a ∧ (x < l → !exponentialDef x y) ∧ (l ≤ x → y = 0)” (by simp) lemma bexp_defined : 𝚺₀-Function₂ (bexp : M → M → M) via bexpDef := by @@ -344,9 +344,9 @@ lemma bexp_defined : 𝚺₀-Function₂ (bexp : M → M → M) via bexpDef := b @[simp] lemma bexp_defined_iff (v) : Semiformula.Evalbm M v bexpDef.val ↔ v 0 = bexp (v 1) (v 2) := bexp_defined.df.iff v -instance bexp_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ (bexp : M → M → M) := Defined.to_definable _ bexp_defined +instance bexp_definable : 𝚺₀-Function₂ (bexp : M → M → M) := Defined.to_definable _ bexp_defined -instance : Bounded₂ ℒₒᵣ (bexp : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₂ (bexp : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ lemma bexp_monotone_iff {a i j : M} (hi : i < ‖a‖) (hj : j < ‖a‖) : bexp a i < bexp a j ↔ i < j := Iff.symm <| Exponential.monotone_iff (by simp [hi]) (by simp [hj]) @@ -415,7 +415,7 @@ lemma fbit_eq_zero_iff {a i : M} : fbit a i = 0 ↔ ¬LenBit (bexp a i) a := by lemma fbit_eq_zero_of_le {a i : M} (hi : ‖a‖ ≤ i) : fbit a i = 0 := by simp [fbit, bexp_eq_zero_of_le hi] -def _root_.LO.FirstOrder.Arith.fbitDef : 𝚺₀-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.fbitDef : 𝚺₀.Semisentence 3 := .mkSigma “b a i | ∃ x <⁺ a, !bexpDef x a i ∧ ∃ y <⁺ a, !divDef y a x ∧ !remDef b y 2” (by simp) lemma fbit_defined : 𝚺₀-Function₂ (fbit : M → M → M) via fbitDef := by @@ -427,9 +427,9 @@ lemma fbit_defined : 𝚺₀-Function₂ (fbit : M → M → M) via fbitDef := b @[simp] lemma fbit_defined_iff (v) : Semiformula.Evalbm M v fbitDef.val ↔ v 0 = fbit (v 1) (v 2) := fbit_defined.df.iff v -instance fbit_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ (fbit : M → M → M) := Defined.to_definable _ fbit_defined +instance fbit_definable : 𝚺₀-Function₂ (fbit : M → M → M) := Defined.to_definable _ fbit_defined -instance : Bounded₂ ℒₒᵣ (fbit : M → M → M) := ⟨‘1’, λ _ ↦ by simp⟩ +instance : Bounded₂ (fbit : M → M → M) := ⟨‘1’, λ _ ↦ by simp⟩ @[simp] lemma fbit_zero (i : M) : fbit 0 i = 0 := by simp [fbit] diff --git a/Arithmetization/ISigmaZero/Exponential/PPow2.lean b/Arithmetization/ISigmaZero/Exponential/PPow2.lean index 48e038f..7da1d6e 100644 --- a/Arithmetization/ISigmaZero/Exponential/PPow2.lean +++ b/Arithmetization/ISigmaZero/Exponential/PPow2.lean @@ -14,7 +14,7 @@ variable [M ⊧ₘ* 𝐈𝚺₀] def SPPow2 (m : M) : Prop := ¬LenBit 1 m ∧ LenBit 2 m ∧ ∀ i ≤ m, Pow2 i → 2 < i → (LenBit i m ↔ (√i)^2 = i ∧ LenBit (√i) m) -def _root_.LO.FirstOrder.Arith.sppow2Def : 𝚺₀-Semisentence 1 := +def _root_.LO.FirstOrder.Arith.sppow2Def : 𝚺₀.Semisentence 1 := .mkSigma “ m | ¬!lenbitDef 1 m ∧ !lenbitDef 2 m ∧ ∀ i <⁺ m, !pow2Def i → 2 < i → (!lenbitDef i m ↔ ∃ s <⁺ i, !sqrtDef s i ∧ s * s = i ∧ !lenbitDef s m) @@ -32,7 +32,7 @@ lemma sppow2_defined : 𝚺₀-Predicate (SPPow2 : M → Prop) via sppow2Def := def PPow2 (i : M) : Prop := Pow2 i ∧ ∃ m < 2 * i, SPPow2 m ∧ LenBit i m -def _root_.LO.FirstOrder.Arith.ppow2Def : 𝚺₀-Semisentence 1 := +def _root_.LO.FirstOrder.Arith.ppow2Def : 𝚺₀.Semisentence 1 := .mkSigma “i | !pow2Def i ∧ ∃ m < 2 * i, !sppow2Def m ∧ !lenbitDef i m” (by simp) lemma ppow2_defined : 𝚺₀-Predicate (PPow2 : M → Prop) via ppow2Def := by @@ -92,7 +92,7 @@ lemma sq_le_of_lt {i j : M} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : intro hij suffices ∀ i < j, Pow2 i → Pow2 j → LenBit i m → LenBit j m → i^2 ≤ j from this i hij pi pj hi hj clear i pi hi hij pj hj - induction j using order_induction_iSigmaZero + induction j using order_induction_sigma0 · definability case ind j IH => intro i hij pi pj hi hj @@ -263,7 +263,7 @@ lemma sq_le_of_lt {i j : M} (hi : PPow2 i) (hj : PPow2 j) : i < j → i^2 ≤ j intro hij suffices ∀ i < j, PPow2 i → PPow2 j → i^2 ≤ j from this i hij hi hj clear hi hij hj - induction j using order_induction_iSigmaZero + induction j using order_induction_sigma0 · definability case ind j IH => intro i hij hi hj diff --git a/Arithmetization/ISigmaZero/Exponential/Pow2.lean b/Arithmetization/ISigmaZero/Exponential/Pow2.lean index 692f7d6..25c4f38 100644 --- a/Arithmetization/ISigmaZero/Exponential/Pow2.lean +++ b/Arithmetization/ISigmaZero/Exponential/Pow2.lean @@ -14,7 +14,7 @@ variable [M ⊧ₘ* 𝐈open] def Pow2 (a : M) : Prop := 0 < a ∧ ∀ r ≤ a, 1 < r → r ∣ a → 2 ∣ r -def _root_.LO.FirstOrder.Arith.pow2Def : 𝚺₀-Semisentence 1 := +def _root_.LO.FirstOrder.Arith.pow2Def : 𝚺₀.Semisentence 1 := .mkSigma “a | 0 < a ∧ ∀ r <⁺ a, 1 < r → r ∣ a → 2 ∣ r” (by simp [Hierarchy.pi_zero_iff_sigma_zero]) lemma pow2_defined : 𝚺₀-Predicate (Pow2 : M → Prop) via pow2Def := by @@ -108,7 +108,7 @@ section LenBit /-- $\mathrm{LenBit} (2^i, a) \iff \text{$i$th-bit of $a$ is $1$}$. -/ def LenBit (i a : M) : Prop := ¬2 ∣ (a / i) -def _root_.LO.FirstOrder.Arith.lenbitDef : 𝚺₀-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.lenbitDef : 𝚺₀.Semisentence 2 := .mkSigma “i a | ∃ z <⁺ a, !divDef.val z a i ∧ ¬2 ∣ z” (by simp) lemma lenbit_defined : 𝚺₀-Relation (LenBit : M → M → Prop) via lenbitDef := by @@ -204,7 +204,7 @@ lemma mul {a b : M} (ha : Pow2 a) (hb : Pow2 b) : Pow2 (a * b) := by suffices ∀ b : M, ∀ a ≤ b, Pow2 a → Pow2 b → Pow2 (a * b) by exact this b a hab ha hb intro b - induction b using order_induction_iSigmaZero + induction b using order_induction_sigma0 · definability case ind IH a b IH => intro a hab ha hb @@ -234,7 +234,7 @@ lemma sq {a : M} : Pow2 a → Pow2 (a^2) := by lemma dvd_of_le {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b → a ∣ b := by suffices ∀ b : M, ∀ a ≤ b, Pow2 a → Pow2 b → a ∣ b by intro hab; exact this b a hab ha hb - intro b; induction b using order_induction_iSigmaZero + intro b; induction b using order_induction_sigma0 · definability case ind b IH => intro a hab ha hb @@ -282,7 +282,7 @@ lemma sq_or_dsq {a : M} (pa : Pow2 a) : ∃ b, a = b^2 ∨ a = 2 * b^2 := by suffices ∃ b ≤ a, a = b^2 ∨ a = 2 * b^2 by rcases this with ⟨b, _, h⟩ exact ⟨b, h⟩ - induction a using order_induction_iSigmaZero + induction a using order_induction_sigma0 · definability case ind a IH => rcases Pow2.elim'.mp pa with (rfl | ⟨ha, a, rfl, pa'⟩) diff --git a/Arithmetization/OmegaOne/Basic.lean b/Arithmetization/OmegaOne/Basic.lean index 0a99881..0f052dc 100644 --- a/Arithmetization/OmegaOne/Basic.lean +++ b/Arithmetization/OmegaOne/Basic.lean @@ -61,7 +61,7 @@ lemma exponential_hash (a b : M) : Exponential (‖a‖ * ‖b‖) (a # b) := Cl lemma exponential_hash_one (a : M) : Exponential ‖a‖ (a # 1) := by simpa using exponential_hash a 1 -def hashDef : 𝚺₀-Semisentence 3 := .mkSigma +def hashDef : 𝚺₀.Semisentence 3 := .mkSigma “z x y | ∃ lx <⁺ x, ∃ ly <⁺ y, !lengthDef lx x ∧ !lengthDef ly y ∧ !exponentialDef (lx * ly) z” (by simp) lemma hash_defined : 𝚺₀-Function₂ (Hash.hash : M → M → M) via hashDef := by @@ -70,7 +70,7 @@ lemma hash_defined : 𝚺₀-Function₂ (Hash.hash : M → M → M) via hashDef · intro h; exact ⟨‖v 1‖, by simp, ‖v 2‖, by simp, rfl, rfl, by rw [h]; exact exponential_hash _ _⟩ · rintro ⟨_, _, _, _, rfl, rfl, h⟩; exact h.uniq (exponential_hash (v 1) (v 2)) -instance hash_definable : DefinableFunction₂ ℒₒᵣ 𝚺₀ (Hash.hash : M → M → M) := Defined.to_definable _ hash_defined +instance hash_definable : 𝚺₀-Function₂ (Hash.hash : M → M → M) := Defined.to_definable _ hash_defined @[simp] lemma hash_pow2 (a b : M) : Pow2 (a # b) := (exponential_hash a b).range_pow2 diff --git a/Arithmetization/OmegaOne/Nuon.lean b/Arithmetization/OmegaOne/Nuon.lean index f508247..f877319 100644 --- a/Arithmetization/OmegaOne/Nuon.lean +++ b/Arithmetization/OmegaOne/Nuon.lean @@ -60,7 +60,7 @@ lemma ext_graph (z S L i : M) : z = S{L}[i] ↔ rcases h lt with ⟨b, hb, Hb, hL, _, HhL, _, _, rfl, rfl⟩ exact ⟨b, hb, Hb, by rw [HhL.uniq (exponential_hash_one L)]⟩ -def extDef : 𝚺₀-Semisentence 4 := .mkSigma +def extDef : 𝚺₀.Semisentence 4 := .mkSigma “z L S i | ∃ lS <⁺ S, !lengthDef lS S ∧ ∃ lL <⁺ L, !lengthDef lL L ∧ (lS ≤ i * lL → z = 0) ∧ @@ -213,7 +213,7 @@ lemma SeriesSegment.series {U I L A k n : M} (H : SeriesSegment U I L A k n) : lemma IsSegment.le_add {L A start intv S : M} (H : IsSegment L A start intv S) : ∀ i ≤ intv, S{L}[i] ≤ S{L}[0] + i := by intro i - induction i using induction_iSigmaZero + induction i using induction_sigma0 · definability case zero => simp case succ i IH => @@ -238,7 +238,7 @@ lemma Segment.uniq {U L A start intv nₛ nₑ₁ nₑ₂ : M} rcases H₁ with ⟨S₁, _, HS₁, Hₛ, rfl⟩ rcases H₂ with ⟨S₂, _, HS₂, rfl, rfl⟩ suffices ∀ i ≤ intv, S₁{L}[i] = S₂{L}[i] from this intv (by rfl) - intro i; induction i using induction_iSigmaZero + intro i; induction i using induction_sigma0 · definability case zero => intro _; exact Hₛ case succ i IH => @@ -249,7 +249,7 @@ lemma Segment.uniq {U L A start intv nₛ nₑ₁ nₑ₂ : M} lemma IsSeries.le_add {U I L A iter T : M} (H : IsSeries U I L A iter T) : ∀ l ≤ iter, T{L}[l] ≤ T{L}[0] + ‖I‖ * l := by intro l - induction l using induction_iSigmaZero + induction l using induction_sigma0 · definability case zero => simp case succ l IH => @@ -266,7 +266,7 @@ lemma Series.uniq {U I L A iter n₁ n₂ : M} (H₁ : Series U I L A iter n₁) rcases H₁ with ⟨T₁, _, HT₁, Hₛ₁, rfl⟩ rcases H₂ with ⟨T₂, _, HT₂, Hₛ₂, rfl⟩ suffices ∀ i ≤ iter, T₁{L}[i] = T₂{L}[i] from this iter (by rfl) - intro i; induction i using induction_iSigmaZero + intro i; induction i using induction_sigma0 · definability case zero => intro _; simp [Hₛ₁, Hₛ₂] case succ i IH => @@ -483,7 +483,7 @@ lemma sq_polyI_hash_polyL_polybounded {A : M} (pos : 0 < A) : ((polyI A) # (poly def NuonAux (A k n : M) : Prop := SeriesSegment (polyU A) (polyI A) (polyL A) A k n -def isSegmentDef : 𝚺₀-Semisentence 5 := .mkSigma +def isSegmentDef : 𝚺₀.Semisentence 5 := .mkSigma “L A start intv S | ∀ i < intv, ∃ S_L_i_succ <⁺ S, !extDef S_L_i_succ L S (i + 1) ∧ @@ -499,14 +499,14 @@ lemma isSegmentDef_defined : Arith.Defined (M := M) (λ v ↦ IsSegment (v 0) (v · intro h; exact ⟨_, by simp, rfl, _, by simp, rfl, _, by simp, rfl, h⟩ · rintro ⟨_, _, rfl, _, _, rfl, _, _, rfl, h⟩; exact h -def segmentDef : 𝚺₀-Semisentence 7 := .mkSigma +def segmentDef : 𝚺₀.Semisentence 7 := .mkSigma “U L A start intv nₛ nₑ | ∃ S < U, !isSegmentDef L A start intv S ∧ !extDef nₛ L S 0 ∧ !extDef nₑ L S intv” (by simp) lemma segmentDef_defined : Arith.Defined (M := M) (λ v ↦ Segment (v 0) (v 1) (v 2) (v 3) (v 4) (v 5) (v 6)) segmentDef := by intro v; simp [Segment, segmentDef, ext_defined.df.iff, isSegmentDef_defined.df.iff, @Eq.comm _ (v 5), @Eq.comm _ (v 6)] rfl -def isSeriesDef : 𝚺₀-Semisentence 6 := .mkSigma +def isSeriesDef : 𝚺₀.Semisentence 6 := .mkSigma “U I L A iter T | ∀ l < iter, ∃ lI <⁺ I, !lengthDef lI I ∧ @@ -529,7 +529,7 @@ lemma isSerieDef_defined : Arith.Defined (M := M) (λ v ↦ IsSeries (v 0) (v 1) simp; rfl -def seriesDef : 𝚺₀-Semisentence 6 := .mkSigma +def seriesDef : 𝚺₀.Semisentence 6 := .mkSigma “U I L A iter n | ∃ T < U, !isSeriesDef U I L A iter T ∧ !extDef 0 L T 0 ∧ !extDef n L T iter” (by simp) lemma seriesDef_defined : Arith.Defined (M := M) (λ v ↦ Series (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) seriesDef := by @@ -539,7 +539,7 @@ lemma seriesDef_defined : Arith.Defined (M := M) (λ v ↦ Series (v 0) (v 1) (v apply and_congr_right; intros simp [Eq.comm] -def seriesSegmentDef : 𝚺₀-Semisentence 6 := .mkSigma +def seriesSegmentDef : 𝚺₀.Semisentence 6 := .mkSigma “U I L A k n | ∃ nₖ <⁺ n, ∃ l <⁺ I, !lengthDef l I ∧ @@ -555,7 +555,7 @@ lemma seriesSegmentDef_defined : Arith.Defined (M := M) (λ v ↦ SeriesSegment apply and_congr_right; intros rw [bex_eq_le_iff, bex_eq_le_iff, bex_eq_le_iff]; simp; rfl -def nuonAuxDef : 𝚺₀-Semisentence 3 := .mkSigma +def nuonAuxDef : 𝚺₀.Semisentence 3 := .mkSigma “A k n | ∃ lA <⁺ A, !lengthDef lA A ∧ ∃ sA <⁺ lA, !sqrtDef sA lA ∧ @@ -591,7 +591,7 @@ lemma NuonAux.exists {k : M} (hk : k ≤ ‖A‖) : ∃ n, NuonAux A k n := by suffices ∃ n ≤ k, NuonAux A k n by rcases this with ⟨n, _, h⟩; exact ⟨n, h⟩ revert hk - induction k using induction_iSigmaZero + induction k using induction_sigma0 · definability case zero => intro _; exact ⟨0, by simp⟩ @@ -618,7 +618,7 @@ lemma NuonAux.two_mul {k n : M} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux suffices ∀ n ≤ k, k ≤ ‖A‖ → NuonAux A k n → NuonAux (2 * A) (k + 1) n by intro n hk H exact this n H.le hk H - induction k using induction_iSigmaZero + induction k using induction_sigma0 · definability case zero => simp; simpa using (NuonAux.initial (2 * A)).succ (by simp) @@ -633,7 +633,7 @@ lemma NuonAux.two_mul_add_one {k n : M} (hk : k ≤ ‖A‖) : NuonAux A k n → suffices ∀ n ≤ k, k ≤ ‖A‖ → NuonAux A k n → NuonAux (2 * A + 1) (k + 1) (n + 1) by intro n hk H exact this n H.le hk H - induction k using induction_iSigmaZero + induction k using induction_sigma0 · definability case zero => simpa using (NuonAux.initial (2 * A + 1)).succ (by simp) @@ -673,7 +673,7 @@ lemma nuon_bit1 (a : M) : nuon (2 * a + 1) = nuon a + 1 := by @[simp] lemma nuon_zero : nuon (0 : M) = 0 := Nuon.nuon_eq (by simp [Nuon]) -def _root_.LO.FirstOrder.Arith.nuonDef : 𝚺₀-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.nuonDef : 𝚺₀.Semisentence 2 := .mkSigma “n A | ∃ l <⁺ A, !lengthDef l A ∧ !Nuon.nuonAuxDef A l n” (by simp) lemma nuon_defined : 𝚺₀-Function₁ (nuon : M → M) via nuonDef := by @@ -684,7 +684,7 @@ lemma nuon_defined : 𝚺₀-Function₁ (nuon : M → M) via nuonDef := by @[simp] lemma eval_nuon_iff (v) : Semiformula.Evalbm M v nuonDef.val ↔ v 0 = nuon (v 1) :=nuon_defined.df.iff v -instance nuon_definable : DefinableFunction₁ ℒₒᵣ 𝚺₀ (nuon : M → M) := Defined.to_definable _ nuon_defined +instance nuon_definable : 𝚺₀-Function₁ (nuon : M → M) := Defined.to_definable _ nuon_defined end From c3aef520a1ba4c573e10889e3e225e4cc8728f93 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Fri, 16 Aug 2024 05:03:21 +0900 Subject: [PATCH 07/10] wip --- Arithmetization/ISigmaOne/Bit.lean | 256 +++++++++-------- Arithmetization/ISigmaOne/HFS/Basic.lean | 46 +-- Arithmetization/ISigmaOne/HFS/Fixpoint.lean | 148 +++++----- Arithmetization/ISigmaOne/HFS/PRF.lean | 70 ++--- Arithmetization/ISigmaOne/HFS/Seq.lean | 269 +++++++++--------- .../ISigmaOne/HFS/Supplemental.lean | 58 ++-- Arithmetization/ISigmaOne/HFS/Vec.lean | 52 ++-- Arithmetization/ISigmaOne/Ind.lean | 10 +- .../ISigmaOne/Metamath/Formula/Basic.lean | 70 ++--- .../ISigmaOne/Metamath/Formula/Functions.lean | 42 +-- .../ISigmaOne/Metamath/Formula/Iteration.lean | 8 +- .../ISigmaOne/Metamath/Language.lean | 4 +- .../ISigmaOne/Metamath/Proof/Derivation.lean | 20 +- .../ISigmaOne/Metamath/Proof/Theory.lean | 8 +- .../ISigmaOne/Metamath/Term/Basic.lean | 24 +- .../ISigmaOne/Metamath/Term/Functions.lean | 16 +- .../ISigmaOne/Metamath/Theory/R.lean | 2 +- .../Metamath/Theory/SigmaOneDefinable.lean | 4 +- .../ISigmaZero/Exponential/Exp.lean | 213 +++++++------- .../ISigmaZero/Exponential/Log.lean | 234 +++++++-------- .../ISigmaZero/Exponential/PPow2.lean | 88 +++--- .../ISigmaZero/Exponential/Pow2.lean | 130 ++++----- Arithmetization/OmegaOne/Basic.lean | 64 ++--- Arithmetization/OmegaOne/Nuon.lean | 200 ++++++------- 24 files changed, 1014 insertions(+), 1022 deletions(-) diff --git a/Arithmetization/ISigmaOne/Bit.lean b/Arithmetization/ISigmaOne/Bit.lean index a496471..5ee8c2c 100644 --- a/Arithmetization/ISigmaOne/Bit.lean +++ b/Arithmetization/ISigmaOne/Bit.lean @@ -7,59 +7,55 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] -variable [M ⊧ₘ* 𝐈𝚺₁] +variable [V ⊧ₘ* 𝐈𝚺₁] -def Bit (i a : M) : Prop := LenBit (exp i) a +def Bit (i a : V) : Prop := LenBit (exp i) a -instance : Membership M M := ⟨Bit⟩ +instance : Membership V V := ⟨Bit⟩ def _root_.LO.FirstOrder.Arith.bitDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | ∃ z <⁺ y, !expDef z x ∧ !lenbitDef z y” (by simp) -lemma bit_defined : 𝚺₀-Relation ((· ∈ ·) : M → M → Prop) via bitDef := by +lemma bit_defined : 𝚺₀-Relation ((· ∈ ·) : V → V → Prop) via bitDef := by intro v; simp [bitDef, ←le_iff_lt_succ] constructor · intro h; exact ⟨exp (v 0), by simp [h.le], rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h @[simp] lemma bit_defined_iff (v) : - Semiformula.Evalbm M v bitDef.val ↔ v 0 ∈ v 1 := bit_defined.df.iff v + Semiformula.Evalbm V v bitDef.val ↔ v 0 ∈ v 1 := bit_defined.df.iff v -@[instance, definability] def mem_definable : DefinableRel ℒₒᵣ 𝚺₀ ((· ∈ ·) : M → M → Prop) := Defined.to_definable _ bit_defined +@[instance] def mem_definable : 𝚺₀-Relation ((· ∈ ·) : V → V → Prop) := bit_defined.to_definable -@[simp, instance, definability] def mem_definable' : DefinableRel ℒₒᵣ Γ ((· ∈ ·) : M → M → Prop) := .of_zero mem_definable _ +@[instance] def mem_definable' (ℌ : HierarchySymbol) : ℌ-Relation ((· ∈ ·) : V → V → Prop) := mem_definable.of_zero -lemma mem_absolute (i a : ℕ) : i ∈ a ↔ (i : M) ∈ (a : M) := by - simpa using Defined.shigmaZero_absolute M bit_defined bit_defined ![i, a] +lemma mem_absolute (i a : ℕ) : i ∈ a ↔ (i : V) ∈ (a : V) := by + simpa using Defined.shigmaZero_absolute V bit_defined bit_defined ![i, a] -lemma mem_iff_bit {i a : M} : i ∈ a ↔ Bit i a := iff_of_eq rfl +lemma mem_iff_bit {i a : V} : i ∈ a ↔ Bit i a := iff_of_eq rfl -lemma exp_le_of_mem {i a : M} (h : i ∈ a) : exp i ≤ a := LenBit.le h +lemma exp_le_of_mem {i a : V} (h : i ∈ a) : exp i ≤ a := LenBit.le h -lemma lt_of_mem {i a : M} (h : i ∈ a) : i < a := lt_of_lt_of_le (lt_exp i) (exp_le_of_mem h) +lemma lt_of_mem {i a : V} (h : i ∈ a) : i < a := lt_of_lt_of_le (lt_exp i) (exp_le_of_mem h) -lemma not_mem_of_lt_exp {i a : M} (h : a < exp i) : i ∉ a := fun H ↦ by have := lt_of_le_of_lt (exp_le_of_mem H) h; simp at this +lemma not_mem_of_lt_exp {i a : V} (h : a < exp i) : i ∉ a := fun H ↦ by have := lt_of_le_of_lt (exp_le_of_mem H) h; simp at this section -@[definability] lemma Definable.ball_mem (Γ m) {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction ℒₒᵣ (𝚺, m + 1) f) (h : Definable ℒₒᵣ (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable ℒₒᵣ (Γ, m + 1) (fun v ↦ ∀ x ∈ f v, P v x) := by - have : Definable ℒₒᵣ (Γ, m + 1) (fun v ↦ ∀ x < f v, x ∈ f v → P v x) := - .ball_lt hf (.imp (by simpa using Definable.comp₂ (by simp) (hf.retraction _) (by simp)) h) +@[definability] lemma HierarchySymbol.Boldface.ball_mem (Γ m) {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∀ x ∈ f v, P v x) := by + have : Γ-[m + 1].Boldface (fun v ↦ ∀ x < f v, x ∈ f v → P v x) := + .ball_lt hf (.imp (by simpa using HierarchySymbol.Boldface.comp₂ (by simp) (hf.retraction Fin.succ)) h) exact this.of_iff <| by intro v; exact ⟨fun h x _ hxv ↦ h x hxv, fun h x hx ↦ h x (lt_of_mem hx) hx⟩ -@[definability] lemma Definable.ball_mem' (Γ m) {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction ℒₒᵣ (𝚺, m + 1) f) (h : Definable ℒₒᵣ (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable ℒₒᵣ (Γ, m + 1) (fun v ↦ ∀ {x}, x ∈ f v → P v x) := Definable.ball_mem Γ m hf h - -@[definability] lemma Definable.bex_mem (Γ m) {P : (Fin k → M) → M → Prop} {f : (Fin k → M) → M} - (hf : DefinableFunction ℒₒᵣ (𝚺, m + 1) f) (h : Definable ℒₒᵣ (Γ, m + 1) (fun w ↦ P (w ·.succ) (w 0))) : - Definable ℒₒᵣ (Γ, m + 1) (fun v ↦ ∃ x ∈ f v, P v x) := by - have : Definable ℒₒᵣ (Γ, m + 1) (fun v ↦ ∃ x < f v, x ∈ f v ∧ P v x) := - .bex_lt hf (.and (by simpa using Definable.comp₂ (by simp) (hf.retraction _) (by simp)) h) +@[definability] lemma HierarchySymbol.Boldface.bex_mem (Γ m) {P : (Fin k → V) → V → Prop} {f : (Fin k → V) → V} + (hf : 𝚺-[m + 1].BoldfaceFunction f) (h : Γ-[m + 1].Boldface (fun w ↦ P (w ·.succ) (w 0))) : + Γ-[m + 1].Boldface (fun v ↦ ∃ x ∈ f v, P v x) := by + have : Γ-[m + 1].Boldface (fun v ↦ ∃ x < f v, x ∈ f v ∧ P v x) := + .bex_lt hf (.and (by simpa using HierarchySymbol.Boldface.comp₂ (by simp) (hf.retraction _)) h) exact this.of_iff <| by intro v; exact ⟨by rintro ⟨x, hx, hxv⟩; exact ⟨x, lt_of_mem hx, hx, hxv⟩, by rintro ⟨x, _, hx, hvx⟩; exact ⟨x, hx, hvx⟩⟩ @@ -146,44 +142,44 @@ end @[simp] lemma Hierarchy.memRel₃ {t₁ t₂ t₃ u : Semiterm ℒₒᵣ μ n} : Hierarchy Γ s “:⟪!!t₁, !!t₂, !!t₃⟫:∈ !!u” := by simp[Semiformula.Operator.operator, Matrix.fun_eq_vec₂, operator_mem_def, memRel₃Opr] -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] open LO.Arith -scoped instance : Structure.Mem ℒₒᵣ M := ⟨by intro a b; simp [Semiformula.Operator.val, operator_mem_def, bit_defined.df.iff]⟩ +scoped instance : Structure.Mem ℒₒᵣ V := ⟨by intro a b; simp [Semiformula.Operator.val, operator_mem_def, bit_defined.df.iff]⟩ @[simp] lemma eval_ballIn {t : Semiterm ℒₒᵣ ξ n} {p : Semiformula ℒₒᵣ ξ (n + 1)} {e ε} : - Semiformula.Evalm M e ε (ballIn t p) ↔ ∀ x ∈ t.valm M e ε, Semiformula.Evalm M (x :> e) ε p := by + Semiformula.Evalm V e ε (ballIn t p) ↔ ∀ x ∈ t.valm V e ε, Semiformula.Evalm V (x :> e) ε p := by simp [ballIn] constructor · intro h x hx; exact h x (lt_of_mem hx) hx · intro h x _ hx; exact h x hx @[simp] lemma eval_bexIn {t : Semiterm ℒₒᵣ ξ n} {p : Semiformula ℒₒᵣ ξ (n + 1)} {e ε} : - Semiformula.Evalm M e ε (bexIn t p) ↔ ∃ x ∈ t.valm M e ε, Semiformula.Evalm M (x :> e) ε p := by + Semiformula.Evalm V e ε (bexIn t p) ↔ ∃ x ∈ t.valm V e ε, Semiformula.Evalm V (x :> e) ε p := by simp [bexIn] constructor · rintro ⟨x, _, hx, h⟩; exact ⟨x, hx, h⟩ · rintro ⟨x, hx, h⟩; exact ⟨x, lt_of_mem hx, hx, h⟩ -lemma memRel_defined : 𝚺₀-Relation₃ (fun r x y : M ↦ ⟪x, y⟫ ∈ r) via memRel := by +lemma memRel_defined : 𝚺₀-Relation₃ (fun r x y : V ↦ ⟪x, y⟫ ∈ r) via memRel := by intro v; simp [memRel, pair_defined.df.iff, lt_succ_iff_le] constructor · intro h; exact ⟨⟪v 1, v 2⟫, by simp, rfl, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma memRel₃_defined : 𝚺₀-Relation₄ (fun r x y z : M ↦ ⟪x, y, z⟫ ∈ r) via memRel₃ := by +lemma memRel₃_defined : 𝚺₀-Relation₄ (fun r x y z : V ↦ ⟪x, y, z⟫ ∈ r) via memRel₃ := by intro v; simp [memRel₃, pair_defined.df.iff, lt_succ_iff_le] constructor · intro h; exact ⟨⟪v 2, v 3⟫, by simp, rfl, ⟪v 1, v 2, v 3⟫, by simp, rfl, h⟩ · rintro ⟨_, _, rfl, _, _, rfl, h⟩; exact h -@[simp] lemma eval_memRel {x y r : M} : +@[simp] lemma eval_memRel {x y r : V} : memRelOpr.val ![r, x, y] ↔ ⟪x, y⟫ ∈ r := by unfold Semiformula.Operator.val simp [memRelOpr, pair_defined.df.iff, memRel_defined.df.iff] -@[simp] lemma eval_memRel₃ {x y z r : M} : +@[simp] lemma eval_memRel₃ {x y z r : V} : memRel₃Opr.val ![r, x, y, z] ↔ ⟪x, y, z⟫ ∈ r := by unfold Semiformula.Operator.val simp [memRel₃Opr, pair_defined.df.iff, memRel₃_defined.df.iff] @@ -196,72 +192,72 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] -variable [M ⊧ₘ* 𝐈𝚺₁] +variable [V ⊧ₘ* 𝐈𝚺₁] -lemma mem_iff_mul_exp_add_exp_add {i a : M} : i ∈ a ↔ ∃ k, ∃ r < exp i, a = k * exp (i + 1) + exp i + r := by +lemma mem_iff_mul_exp_add_exp_add {i a : V} : i ∈ a ↔ ∃ k, ∃ r < exp i, a = k * exp (i + 1) + exp i + r := by simp [mem_iff_bit, exp_succ] exact lenbit_iff_add_mul (exp_pow2 i) (a := a) -lemma not_mem_iff_mul_exp_add {i a : M} : i ∉ a ↔ ∃ k, ∃ r < exp i, a = k * exp (i + 1) + r := by +lemma not_mem_iff_mul_exp_add {i a : V} : i ∉ a ↔ ∃ k, ∃ r < exp i, a = k * exp (i + 1) + r := by simp [mem_iff_bit, exp_succ] exact not_lenbit_iff_add_mul (exp_pow2 i) (a := a) section empty -scoped instance : EmptyCollection M := ⟨0⟩ +scoped instance : EmptyCollection V := ⟨0⟩ -lemma emptyset_def : (∅ : M) = 0 := rfl +lemma emptyset_def : (∅ : V) = 0 := rfl -@[simp] lemma not_mem_empty (i : M) : i ∉ (∅ : M) := by simp [emptyset_def, mem_iff_bit, Bit] +@[simp] lemma not_mem_empty (i : V) : i ∉ (∅ : V) := by simp [emptyset_def, mem_iff_bit, Bit] -@[simp] lemma not_mem_zero (i : M) : i ∉ (0 : M) := by simp [mem_iff_bit, Bit] +@[simp] lemma not_mem_zero (i : V) : i ∉ (0 : V) := by simp [mem_iff_bit, Bit] end empty section singleton -scoped instance : Singleton M M := ⟨fun a ↦ exp a⟩ +scoped instance : Singleton V V := ⟨fun a ↦ exp a⟩ -lemma singleton_def (a : M) : {a} = exp a := rfl +lemma singleton_def (a : V) : {a} = exp a := rfl end singleton section insert open Classical in -noncomputable def bitInsert (i a : M) : M := if i ∈ a then a else a + exp i +noncomputable def bitInsert (i a : V) : V := if i ∈ a then a else a + exp i open Classical in -noncomputable def bitRemove (i a : M) : M := if i ∈ a then a - exp i else a +noncomputable def bitRemove (i a : V) : V := if i ∈ a then a - exp i else a -scoped instance : Insert M M := ⟨bitInsert⟩ +scoped instance : Insert V V := ⟨bitInsert⟩ -lemma insert_eq {i a : M} : insert i a = bitInsert i a := rfl +lemma insert_eq {i a : V} : insert i a = bitInsert i a := rfl -lemma singleton_eq_insert (i : M) : ({i} : M) = insert i ∅ := by simp [singleton_def, insert, bitInsert, emptyset_def] +lemma singleton_eq_insert (i : V) : ({i} : V) = insert i ∅ := by simp [singleton_def, insert, bitInsert, emptyset_def] -instance : LawfulSingleton M M where +instance : LawfulSingleton V V where insert_emptyc_eq := fun x ↦ Eq.symm <| singleton_eq_insert x -@[simp] lemma mem_bitInsert_iff {i j a : M} : +@[simp] lemma mem_bitInsert_iff {i j a : V} : i ∈ insert j a ↔ i = j ∨ i ∈ a := by by_cases h : j ∈ a <;> simp [h, insert_eq, bitInsert] · by_cases e : i = j <;> simp [h, e] · simpa [exp_inj.eq_iff] using lenbit_add_pow2_iff_of_not_lenbit (exp_pow2 i) (exp_pow2 j) h -@[simp] lemma mem_bitRemove_iff {i j a : M} : +@[simp] lemma mem_bitRemove_iff {i j a : V} : i ∈ bitRemove j a ↔ i ≠ j ∧ i ∈ a := by by_cases h : j ∈ a <;> simp [h, bitRemove] · simpa [exp_inj.eq_iff] using lenbit_sub_pow2_iff_of_lenbit (exp_pow2 i) (exp_pow2 j) h · rintro _ rfl; contradiction -@[simp] lemma not_mem_bitRemove_self (i a : M) : i ∉ bitRemove i a := by simp +@[simp] lemma not_mem_bitRemove_self (i a : V) : i ∉ bitRemove i a := by simp -lemma insert_graph (b i a : M) : +lemma insert_graph (b i a : V) : b = insert i a ↔ (i ∈ a ∧ b = a) ∨ (i ∉ a ∧ ∃ e ≤ b, e = exp i ∧ b = a + e) := ⟨by rintro rfl; by_cases hi : i ∈ a <;> simp [hi, insert, bitInsert], by by_cases hi : i ∈ a <;> simp only [hi, true_and, not_true_eq_false, false_and, @@ -272,17 +268,17 @@ lemma insert_graph (b i a : M) : def _root_.LO.FirstOrder.Arith.insertDef : 𝚺₀.Semisentence 3 := .mkSigma “b i a | (i ∈ a ∧ b = a) ∨ (i ∉ a ∧ ∃ e <⁺ b, !expDef e i ∧ b = a + e)” (by simp) -lemma insert_defined : 𝚺₀-Function₂ (insert : M → M → M) via insertDef := by +lemma insert_defined : 𝚺₀-Function₂ (insert : V → V → V) via insertDef := by intro v; simp [insertDef, insert_graph] @[simp] lemma insert_defined_iff (v) : - Semiformula.Evalbm M v insertDef.val ↔ v 0 = insert (v 1) (v 2) := insert_defined.df.iff v + Semiformula.Evalbm V v insertDef.val ↔ v 0 = insert (v 1) (v 2) := insert_defined.df.iff v -instance insert_definable : 𝚺₀-Function₂ (insert : M → M → M) := Defined.to_definable _ insert_defined +instance insert_definable : 𝚺₀-Function₂ (insert : V → V → V) := insert_defined.to_definable -instance insert_definable' (Γ) : Γ-Function₂ (insert : M → M → M) := .of_zero insert_definable _ +instance insert_definable' (Γ) : Γ-Function₂ (insert : V → V → V) := insert_definable.of_zero -lemma insert_le_of_le_of_le {i j a b : M} (hij : i ≤ j) (hab : a ≤ b) : insert i a ≤ b + exp j := by +lemma insert_le_of_le_of_le {i j a b : V} (hij : i ≤ j) (hab : a ≤ b) : insert i a ≤ b + exp j := by simp [insert, bitInsert] by_cases hi : i ∈ a <;> simp [hi] · exact le_trans hab (by simp) @@ -290,66 +286,66 @@ lemma insert_le_of_le_of_le {i j a b : M} (hij : i ≤ j) (hab : a ≤ b) : inse end insert -lemma one_eq_singleton : (1 : M) = {∅} := by simp [singleton_eq_insert, insert, bitInsert, emptyset_def] +lemma one_eq_singleton : (1 : V) = {∅} := by simp [singleton_eq_insert, insert, bitInsert, emptyset_def] -@[simp] lemma mem_singleton_iff {i j : M} : - i ∈ ({j} : M) ↔ i = j := by simp [singleton_eq_insert] +@[simp] lemma mem_singleton_iff {i j : V} : + i ∈ ({j} : V) ↔ i = j := by simp [singleton_eq_insert] -lemma bitRemove_lt_of_mem {i a : M} (h : i ∈ a) : bitRemove i a < a := by +lemma bitRemove_lt_of_mem {i a : V} (h : i ∈ a) : bitRemove i a < a := by simp [h, bitRemove, tsub_lt_iff_left (exp_le_of_mem h)] -lemma pos_of_nonempty {i a : M} (h : i ∈ a) : 0 < a := by +lemma pos_of_nonempty {i a : V} (h : i ∈ a) : 0 < a := by by_contra A; simp at A; rcases A; simp_all -@[simp] lemma mem_insert (i a : M) : i ∈ insert i a := by simp +@[simp] lemma mem_insert (i a : V) : i ∈ insert i a := by simp -lemma insert_eq_self_of_mem {i a : M} (h : i ∈ a) : insert i a = a := by +lemma insert_eq_self_of_mem {i a : V} (h : i ∈ a) : insert i a = a := by simp [insert_eq, bitInsert, h] -lemma log_mem_of_pos {a : M} (h : 0 < a) : log a ∈ a := +lemma log_mem_of_pos {a : V} (h : 0 < a) : log a ∈ a := mem_iff_mul_exp_add_exp_add.mpr ⟨0, a - exp log a, (tsub_lt_iff_left (exp_log_le_self h)).mpr (by rw [←two_mul]; exact lt_two_mul_exponential_log h), by simp; exact Eq.symm <| add_tsub_self_of_le (exp_log_le_self h)⟩ -lemma le_log_of_mem {i a : M} (h : i ∈ a) : i ≤ log a := (exp_le_iff_le_log (pos_of_nonempty h)).mp (exp_le_of_mem h) +lemma le_log_of_mem {i a : V} (h : i ∈ a) : i ≤ log a := (exp_le_iff_le_log (pos_of_nonempty h)).mp (exp_le_of_mem h) -lemma succ_mem_iff_mem_div_two {i a : M} : i + 1 ∈ a ↔ i ∈ a / 2 := by simp [mem_iff_bit, Bit, LenBit.iff_rem, exp_succ, div_mul] +lemma succ_mem_iff_mem_div_two {i a : V} : i + 1 ∈ a ↔ i ∈ a / 2 := by simp [mem_iff_bit, Bit, LenBit.iff_rem, exp_succ, div_mul] -lemma lt_length_of_mem {i a : M} (h : i ∈ a) : i < ‖a‖ := by +lemma lt_length_of_mem {i a : V} (h : i ∈ a) : i < ‖a‖ := by simpa [length_of_pos (pos_of_nonempty h), ←le_iff_lt_succ] using le_log_of_mem h -lemma lt_exp_iff {a i : M} : a < exp i ↔ ∀ j ∈ a, j < i := +lemma lt_exp_iff {a i : V} : a < exp i ↔ ∀ j ∈ a, j < i := ⟨fun h j hj ↦ exp_monotone.mp <| lt_of_le_of_lt (exp_le_of_mem hj) h, by contrapose; simp intro (h : exp i ≤ a) have pos : 0 < a := lt_of_lt_of_le (by simp) h exact ⟨log a, log_mem_of_pos pos, (exp_le_iff_le_log pos).mp h⟩⟩ -instance : HasSubset M := ⟨fun a b ↦ ∀ ⦃i⦄, i ∈ a → i ∈ b⟩ +instance : HasSubset V := ⟨fun a b ↦ ∀ ⦃i⦄, i ∈ a → i ∈ b⟩ def _root_.LO.FirstOrder.Arith.bitSubsetDef : 𝚺₀.Semisentence 2 := .mkSigma “a b | ∀ i < a, i ∈ a → i ∈ b” (by simp) -lemma bitSubset_defined : 𝚺₀-Relation ((· ⊆ ·) : M → M → Prop) via bitSubsetDef := by +lemma bitSubset_defined : 𝚺₀-Relation ((· ⊆ ·) : V → V → Prop) via bitSubsetDef := by intro v; simp [bitSubsetDef] exact ⟨by intro h x _ hx; exact h hx, by intro h x hx; exact h x (lt_of_mem hx) hx⟩ @[simp] lemma bitSubset_defined_iff (v) : - Semiformula.Evalbm M v bitSubsetDef.val ↔ v 0 ⊆ v 1 := bitSubset_defined.df.iff v + Semiformula.Evalbm V v bitSubsetDef.val ↔ v 0 ⊆ v 1 := bitSubset_defined.df.iff v -instance bitSubset_definable : DefinableRel ℒₒᵣ 𝚺₀ ((· ⊆ ·) : M → M → Prop) := Defined.to_definable₀ _ bitSubset_defined +instance bitSubset_definable : 𝚺₀-Relation ((· ⊆ ·) : V → V → Prop) := bitSubset_defined.to_definable₀ -@[simp, definability] instance bitSubset_definable' : DefinableRel ℒₒᵣ Γ ((· ⊆ ·) : M → M → Prop) := Defined.to_definable₀ _ bitSubset_defined +@[simp, definability] instance bitSubset_definable' (ℌ : HierarchySymbol) : ℌ-Relation ((· ⊆ ·) : V → V → Prop) := bitSubset_defined.to_definable₀ -lemma subset_iff {a b : M} : a ⊆ b ↔ (∀ x ∈ a, x ∈ b) := by simp [HasSubset.Subset] +lemma subset_iff {a b : V} : a ⊆ b ↔ (∀ x ∈ a, x ∈ b) := by simp [HasSubset.Subset] -@[refl, simp] lemma subset_refl (a : M) : a ⊆ a := by intro x; simp +@[refl, simp] lemma subset_refl (a : V) : a ⊆ a := by intro x; simp -@[trans] lemma subset_trans {a b c : M} (hab : a ⊆ b) (hbc : b ⊆ c) : a ⊆ c := by +@[trans] lemma subset_trans {a b c : V} (hab : a ⊆ b) (hbc : b ⊆ c) : a ⊆ c := by intro x hx; exact hbc (hab hx) -lemma mem_exp_add_succ_sub_one (i j : M) : i ∈ exp (i + j + 1) - 1 := by +lemma mem_exp_add_succ_sub_one (i j : V) : i ∈ exp (i + j + 1) - 1 := by have : exp (i + j + 1) - 1 = (exp j - 1) * exp (i + 1) + exp i + (exp i - 1) := calc exp (i + j + 1) - 1 = exp j * exp (i + 1) - 1 := by simp [exp_add, ←mul_assoc, mul_comm] _ = exp j * exp (i + 1) - exp (i + 1) + exp (i + 1) - 1 := by rw [sub_add_self_of_le]; exact le_mul_of_pos_left (exp_pos j) @@ -360,12 +356,12 @@ lemma mem_exp_add_succ_sub_one (i j : M) : i ∈ exp (i + j + 1) - 1 := by exact mem_iff_mul_exp_add_exp_add.mpr ⟨exp j - 1, exp i - 1, (tsub_lt_iff_left (by simp)).mpr $ by simp, this⟩ /-- under a = {0, 1, 2, ..., a - 1} -/ -def under (a : M) : M := exp a - 1 +def under (a : V) : V := exp a - 1 -@[simp] lemma le_under (a : M) : a ≤ under a := +@[simp] lemma le_under (a : V) : a ≤ under a := le_iff_lt_succ.mpr (by simp [under, show exp a - 1 + 1 = exp a from sub_add_self_of_le (by simp)]) -@[simp] lemma mem_under_iff {i j : M} : i ∈ under j ↔ i < j := by +@[simp] lemma mem_under_iff {i j : V} : i ∈ under j ↔ i < j := by constructor · intro h have : exp i < exp j := calc @@ -381,39 +377,39 @@ def under (a : M) : M := exp a - 1 · exact le_tsub_of_add_le_left this rw [this]; exact mem_exp_add_succ_sub_one i k -@[simp] lemma not_mem_under_self (i : M) : i ∉ under i := by simp +@[simp] lemma not_mem_under_self (i : V) : i ∉ under i := by simp -private lemma under_graph (x y : M) : y = under x ↔ y + 1 = exp x := +private lemma under_graph (x y : V) : y = under x ↔ y + 1 = exp x := ⟨by rintro rfl; simp [under, sub_add_self_of_le], by intro h; have := congr_arg (· - 1) h; simp [under] at this ⊢; exact this⟩ def _root_.LO.FirstOrder.Arith.underDef : 𝚺₀.Semisentence 2 := .mkSigma “y x | !expDef.val (y + 1) x” (by simp) -lemma under_defined : 𝚺₀-Function₁ (under : M → M) via underDef := by +lemma under_defined : 𝚺₀-Function₁ (under : V → V) via underDef := by intro v; simp [underDef, under_graph] @[simp] lemma under_defined_iff (v) : - Semiformula.Evalbm M v underDef.val ↔ v 0 = under (v 1) := under_defined.df.iff v + Semiformula.Evalbm V v underDef.val ↔ v 0 = under (v 1) := under_defined.df.iff v -instance under_definable : 𝚺₀-Function₁ (under : M → M) := Defined.to_definable _ under_defined +instance under_definable : 𝚺₀-Function₁ (under : V → V) := under_defined.to_definable -instance under_definable' (Γ) : Γ-Function₁ (under : M → M) := .of_zero under_definable _ +instance under_definable' (Γ) : Γ-Function₁ (under : V → V) := under_definable.of_zero -lemma eq_zero_of_subset_zero {a : M} : a ⊆ 0 → a = 0 := by +lemma eq_zero_of_subset_zero {a : V} : a ⊆ 0 → a = 0 := by intro h; by_contra A - have : log a ∈ (0 : M) := h (log_mem_of_pos (pos_iff_ne_zero.mpr A)) + have : log a ∈ (0 : V) := h (log_mem_of_pos (pos_iff_ne_zero.mpr A)) simp_all -lemma subset_div_two {a b : M} : a ⊆ b → a / 2 ⊆ b / 2 := by +lemma subset_div_two {a b : V} : a ⊆ b → a / 2 ⊆ b / 2 := by intro ss i hi have : i + 1 ∈ a := succ_mem_iff_mem_div_two.mpr hi exact succ_mem_iff_mem_div_two.mp <| ss this -lemma zero_mem_iff {a : M} : 0 ∉ a ↔ 2 ∣ a := by simp [mem_iff_bit, Bit, LenBit] +lemma zero_mem_iff {a : V} : 0 ∉ a ↔ 2 ∣ a := by simp [mem_iff_bit, Bit, LenBit] -@[simp] lemma zero_not_mem (a : M) : 0 ∉ 2 * a := by simp [mem_iff_bit, Bit, LenBit] +@[simp] lemma zero_not_mem (a : V) : 0 ∉ 2 * a := by simp [mem_iff_bit, Bit, LenBit] -lemma le_of_subset {a b : M} (h : a ⊆ b) : a ≤ b := by +lemma le_of_subset {a b : V} (h : a ⊆ b) : a ≤ b := by induction b using hierarchy_polynomial_induction_oRing_pi₁ generalizing a · definability case zero => @@ -427,37 +423,37 @@ lemma le_of_subset {a b : M} (h : a ⊆ b) : a ≤ b := by have IH : a / 2 ≤ b := IH (by simpa [div_mul_add' b 2 one_lt_two] using subset_div_two h) exact le_trans (le_two_mul_div_two_add_one a) (by simpa using IH) -lemma mem_ext {a b : M} (h : ∀ i, i ∈ a ↔ i ∈ b) : a = b := +lemma mem_ext {a b : V} (h : ∀ i, i ∈ a ↔ i ∈ b) : a = b := le_antisymm (le_of_subset fun i hi ↦ (h i).mp hi) (le_of_subset fun i hi ↦ (h i).mpr hi) -lemma pos_iff_nonempty {s : M} : 0 < s ↔ s ≠ ∅ := pos_iff_ne_zero +lemma pos_iff_nonempty {s : V} : 0 < s ↔ s ≠ ∅ := pos_iff_ne_zero -lemma nonempty_of_pos {a : M} (h : 0 < a) : ∃ i, i ∈ a := by +lemma nonempty_of_pos {a : V} (h : 0 < a) : ∃ i, i ∈ a := by by_contra A have : a = 0 := mem_ext (by simpa using A) simp [this] at h -lemma eq_empty_or_nonempty (a : M) : a = ∅ ∨ ∃ i, i ∈ a := by +lemma eq_empty_or_nonempty (a : V) : a = ∅ ∨ ∃ i, i ∈ a := by rcases zero_le a with (rfl | pos) · simp [emptyset_def] · right; exact nonempty_of_pos pos -lemma nonempty_iff {s : M} : s ≠ ∅ ↔ ∃ x, x ∈ s := by +lemma nonempty_iff {s : V} : s ≠ ∅ ↔ ∃ x, x ∈ s := by rcases eq_empty_or_nonempty s with ⟨rfl, hy⟩ <;> simp simp [show s ≠ ∅ from by rintro rfl; simp_all]; assumption -lemma isempty_iff {s : M} : s = ∅ ↔ ∀ x, x ∉ s := by +lemma isempty_iff {s : V} : s = ∅ ↔ ∀ x, x ∉ s := by simpa using not_iff_not.mpr (nonempty_iff (s := s)) -@[simp] lemma empty_subset (s : M) : ∅ ⊆ s := by intro x; simp +@[simp] lemma empty_subset (s : V) : ∅ ⊆ s := by intro x; simp -lemma lt_of_lt_log {a b : M} (pos : 0 < b) (h : ∀ i ∈ a, i < log b) : a < b := by +lemma lt_of_lt_log {a b : V} (pos : 0 < b) (h : ∀ i ∈ a, i < log b) : a < b := by rcases zero_le a with (rfl | apos) · exact pos by_contra A exact not_lt_of_le (log_monotone <| show b ≤ a by simpa using A) (h (log a) (log_mem_of_pos apos)) -@[simp] lemma under_inj {i j : M} : under i = under j ↔ i = j := ⟨fun h ↦ by +@[simp] lemma under_inj {i j : V} : under i = under j ↔ i = j := ⟨fun h ↦ by by_contra ne wlog lt : i < j · exact this (Eq.symm h) (Ne.symm ne) (lt_of_le_of_ne (by simpa using lt) (Ne.symm ne)) @@ -465,12 +461,12 @@ lemma lt_of_lt_log {a b : M} (pos : 0 < b) (h : ∀ i ∈ a, i < log b) : a < b have : i ∈ under i := by rw [h]; simp [mem_under_iff, lt] contradiction, by rintro rfl; simp⟩ -@[simp] lemma under_zero : under (0 : M) = ∅ := mem_ext (by simp [mem_under_iff]) +@[simp] lemma under_zero : under (0 : V) = ∅ := mem_ext (by simp [mem_under_iff]) -@[simp] lemma under_succ (i : M) : under (i + 1) = insert i (under i) := +@[simp] lemma under_succ (i : V) : under (i + 1) = insert i (under i) := mem_ext (by simp [mem_under_iff, lt_succ_iff_le, le_iff_eq_or_lt]) -lemma insert_remove {i a : M} (h : i ∈ a) : insert i (bitRemove i a) = a := mem_ext <| by +lemma insert_remove {i a : V} (h : i ∈ a) : insert i (bitRemove i a) = a := mem_ext <| by simp; intro j constructor · rintro (rfl | ⟨_, hj⟩) <;> assumption @@ -478,19 +474,21 @@ lemma insert_remove {i a : M} (h : i ∈ a) : insert i (bitRemove i a) = a := me section -variable {m : ℕ} [Fact (1 ≤ m)] [M ⊧ₘ* 𝐈𝐍𝐃𝚺 m] +variable {m : ℕ} [Fact (1 ≤ m)] [V ⊧ₘ* 𝐈𝐍𝐃𝚺 m] -private lemma finset_comprehension_aux (Γ : Polarity) {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) +private lemma finset_comprehension_aux (Γ : Polarity) {P : V → Prop} (hP : Γ-[m]-Predicate P) (a : V) : + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i := by - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) have : ∃ s < exp a, ∀ i < a, P i → i ∈ s := ⟨under a, pred_lt_self_of_pos (by simp), fun i hi _ ↦ by simpa [mem_under_iff] using hi⟩ rcases this with ⟨s, hsn, hs⟩ - have : (Γ.alt, m)-Predicate (fun s : M ↦ ∀ i < a, P i → i ∈ s) := by - apply Definable.ball_lt₀; simp; apply Definable.imp <;> definability - have : ∃ t, (∀ i < a, P i → i ∈ t) ∧ ∀ t' < t, ∃ x < a, P x ∧ x ∉ (t' : M) := by - simpa using least_number_h (L := ℒₒᵣ) Γ.alt m this hs + have : Γ.alt-[m]-Predicate (fun s : V ↦ ∀ i < a, P i → i ∈ s) := by + apply HierarchySymbol.Boldface.ball_blt; simp; apply HierarchySymbol.Boldface.imp + · simpa using HierarchySymbol.Boldface.bcomp₁ (by definability) + · simpa using HierarchySymbol.Boldface.bcomp₂ (by definability) (by definability) + have : ∃ t, (∀ i < a, P i → i ∈ t) ∧ ∀ t' < t, ∃ x < a, P x ∧ x ∉ (t' : V) := by + simpa using least_number_h Γ.alt m this hs rcases this with ⟨t, ht, t_minimal⟩ have t_le_s : t ≤ s := not_lt.mp (by intro lt @@ -505,18 +503,18 @@ private lemma finset_comprehension_aux (Γ : Polarity) {P : M → Prop} (hP : ( rcases hm (ht j hjn Hj); contradiction exact ⟨t, lt_of_le_of_lt t_le_s hsn, fun i hi ↦ ⟨this i hi, ht i hi⟩⟩ -theorem finset_comprehension {Γ} {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) +theorem finset_comprehension {Γ} {P : V → Prop} (hP : Γ-[m]-Predicate P) (a : V) : + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i := match Γ with | 𝚺 => finset_comprehension_aux 𝚺 hP a | 𝚷 => finset_comprehension_aux 𝚷 hP a | 𝚫 => finset_comprehension_aux 𝚺 hP.of_delta a -theorem finset_comprehension_exists_unique {P : M → Prop} (hP : (Γ, m)-Predicate P) (a : M) : - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) +theorem finset_comprehension_exists_unique {P : V → Prop} (hP : Γ-[m]-Predicate P) (a : V) : + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) ∃! s, s < exp a ∧ ∀ i < a, i ∈ s ↔ P i := by - haveI : M ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) + haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) rcases finset_comprehension hP a with ⟨s, hs, Hs⟩ exact ExistsUnique.intro s ⟨hs, Hs⟩ (by intro t ⟨ht, Ht⟩ @@ -536,11 +534,11 @@ section ISigma₁ instance : Fact (1 ≤ 1) := ⟨by rfl⟩ -theorem finset_comprehension₁ {P : M → Prop} (hP : (Γ, 1)-Predicate P) (a : M) : +theorem finset_comprehension₁ {P : V → Prop} (hP : Γ-[1]-Predicate P) (a : V) : ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i := finset_comprehension hP a -theorem finset_comprehension₁! {P : M → Prop} (hP : (Γ, 1)-Predicate P) (a : M) : +theorem finset_comprehension₁! {P : V → Prop} (hP : Γ-[1]-Predicate P) (a : V) : ∃! s, s < exp a ∧ (∀ i < a, i ∈ s ↔ P i) := by rcases finset_comprehension₁ hP a with ⟨s, hs, Ha⟩ exact ExistsUnique.intro s ⟨hs, Ha⟩ @@ -556,8 +554,8 @@ theorem finset_comprehension₁! {P : M → Prop} (hP : (Γ, 1)-Predicate P) (a have : x < a := exp_monotone.mp <| LE.le.trans_lt (exp_le_of_mem hx) hs exact (Hb x this).mpr <| (Ha x this).mp hx) -theorem finite_comprehension₁! {P : M → Prop} (hP : (Γ, 1)-Predicate P) (fin : ∃ m, ∀ i, P i → i < m) : - ∃! s : M, ∀ i, i ∈ s ↔ P i := by +theorem finite_comprehension₁! {P : V → Prop} (hP : Γ-[1]-Predicate P) (fin : ∃ m, ∀ i, P i → i < m) : + ∃! s : V, ∀ i, i ∈ s ↔ P i := by rcases fin with ⟨m, mh⟩ rcases finset_comprehension₁ hP m with ⟨s, hs, Hs⟩ have H : ∀ i, i ∈ s ↔ P i := @@ -567,12 +565,12 @@ theorem finite_comprehension₁! {P : M → Prop} (hP : (Γ, 1)-Predicate P) (fi exact ExistsUnique.intro s H (fun s' H' ↦ mem_ext <| fun i ↦ by simp [H, H']) /- -def setExt {Γ} (p : 𝚫₁-Semisentence (n + 1)) : Γ-Semisentence (n + 1) := +def setExt {Γ} (p : 𝚫₁.Semisentence (n + 1)) : Γ.Semisentence (n + 1) := match Γ with | (𝚺, m) => .mkSigma “u | ∀ x < u, x ∈ u ↔ !p x ⋯” (by { }) -lemma set_iff {n} {f : (Fin n → M) → M} {R : (Fin (n + 1) → M) → Prop} - (hf : ∀ v x, x ∈ f v ↔ R (x :> v)) {Γ} (p : (Γ, 1)-Semisentence (n + 1)) : DefinedFunction ℒₒᵣ (Γ, 1) f p := by { +lemma set_iff {n} {f : (Fin n → V) → V} {R : (Fin (n + 1) → V) → Prop} + (hf : ∀ v x, x ∈ f v ↔ R (x :> v)) {Γ} (p : (Γ, 1).Semisentence (n + 1)) : DefinedFunction ℒₒᵣ (Γ, 1) f p := by { } -/ diff --git a/Arithmetization/ISigmaOne/HFS/Basic.lean b/Arithmetization/ISigmaOne/HFS/Basic.lean index 28f413b..ef82eee 100644 --- a/Arithmetization/ISigmaOne/HFS/Basic.lean +++ b/Arithmetization/ISigmaOne/HFS/Basic.lean @@ -13,7 +13,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] @[simp] lemma susbset_insert (x a : V) : a ⊆ insert x a := by intro z hz; simp [hz] @@ -86,9 +86,9 @@ lemma sUnion_defined : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) via sUnionD @[simp] lemma sUnion_defined_iff (v) : Semiformula.Evalbm V v sUnionDef.val ↔ v 0 = ⋃ʰᶠ v 1 := sUnion_defined.df.iff v -instance sUnion_definable : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) := Defined.to_definable _ sUnion_defined +instance sUnion_definable : 𝚺₀-Function₁ ((⋃ʰᶠ ·) : V → V) := sUnion_defined.to_definable -instance sUnion_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ ((⋃ʰᶠ ·) : V → V) := .of_zero sUnion_definable _ +instance sUnion_definable' (ℌ : HierarchySymbol) : ℌ-Function₁ ((⋃ʰᶠ ·) : V → V) := sUnion_definable.of_zero end sUnion @@ -119,9 +119,9 @@ lemma union_defined : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) via unio @[simp] lemma union_defined_iff (v) : Semiformula.Evalbm V v unionDef.val ↔ v 0 = v 1 ∪ v 2 := union_defined.df.iff v -instance union_definable : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) := Defined.to_definable _ union_defined +instance union_definable : 𝚺₀-Function₂ ((· ∪ ·) : V → V → V) := union_defined.to_definable -instance union_definable' (Γ) : DefinableFunction₂ ℒₒᵣ Γ ((· ∪ ·) : V → V → V) := .of_zero union_definable _ +instance union_definable' (ℌ : HierarchySymbol) : ℌ-Function₂ ((· ∪ ·) : V → V → V) := union_definable.of_zero lemma insert_eq_union_singleton (a s : V) : insert a s = {a} ∪ s := mem_ext (fun x ↦ by simp) @@ -241,9 +241,9 @@ lemma product_defined : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) vi @[simp] lemma product_defined_iff (v) : Semiformula.Evalbm V v productDef.val ↔ v 0 = v 1 ×ʰᶠ v 2 := product_defined.df.iff v -instance product_definable : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) := Defined.to_definable _ product_defined +instance product_definable : 𝚺₀-Function₂ ((· ×ʰᶠ ·) : V → V → V) := product_defined.to_definable -instance product_definable' (Γ) : DefinableFunction₂ ℒₒᵣ Γ ((· ×ʰᶠ ·) : V → V → V) := .of_zero product_definable _ +instance product_definable' (ℌ : HierarchySymbol) : ℌ-Function₂ ((· ×ʰᶠ ·) : V → V → V) := product_definable.of_zero end product @@ -252,10 +252,10 @@ section domain lemma domain_exists_unique (s : V) : ∃! d : V, ∀ x, x ∈ d ↔ ∃ y, ⟪x, y⟫ ∈ s := by have : 𝚺₁-Predicate fun x ↦ ∃ y, ⟪x, y⟫ ∈ s := - DefinablePred.of_iff (fun x ↦ ∃ y < s, ⟪x, y⟫ ∈ s) + HierarchySymbol.BoldfacePred.of_iff (Q := fun x ↦ ∃ y < s, ⟪x, y⟫ ∈ s) + (by definability) (fun x ↦ ⟨by rintro ⟨y, hy⟩; exact ⟨y, lt_of_le_of_lt (le_pair_right x y) (lt_of_mem hy), hy⟩, by rintro ⟨y, _, hy⟩; exact ⟨y, hy⟩⟩) - (by definability) exact finite_comprehension₁! this (⟨s, fun x ↦ by rintro ⟨y, hy⟩; exact lt_of_le_of_lt (le_pair_left x y) (lt_of_mem hy)⟩) @@ -285,9 +285,9 @@ lemma domain_defined : 𝚺₀-Function₁ (domain : V → V) via domainDef := b @[simp] lemma domain_defined_iff (v) : Semiformula.Evalbm V v domainDef.val ↔ v 0 = domain (v 1) := domain_defined.df.iff v -instance domain_definable : 𝚺₀-Function₁ (domain : V → V) := Defined.to_definable _ domain_defined +instance domain_definable : 𝚺₀-Function₁ (domain : V → V) := domain_defined.to_definable -instance domain_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ (domain : V → V) := .of_zero domain_definable _ +instance domain_definable' (ℌ : HierarchySymbol) : ℌ-Function₁ (domain : V → V) := domain_definable.of_zero @[simp] lemma domain_empty : domain (∅ : V) = ∅ := mem_ext (by simp [mem_domain_iff]) @@ -341,10 +341,10 @@ section range lemma range_exists_unique (s : V) : ∃! r : V, ∀ y, y ∈ r ↔ ∃ x, ⟪x, y⟫ ∈ s := by have : 𝚺₁-Predicate fun y ↦ ∃ x, ⟪x, y⟫ ∈ s := - DefinablePred.of_iff (fun y ↦ ∃ x < s, ⟪x, y⟫ ∈ s) + HierarchySymbol.BoldfacePred.of_iff (Q := fun y ↦ ∃ x < s, ⟪x, y⟫ ∈ s) + (by definability) (fun y ↦ ⟨by rintro ⟨x, hy⟩; exact ⟨x, lt_of_le_of_lt (le_pair_left x y) (lt_of_mem hy), hy⟩, by rintro ⟨y, _, hy⟩; exact ⟨y, hy⟩⟩) - (by definability) exact finite_comprehension₁! this (⟨s, fun y ↦ by rintro ⟨x, hx⟩; exact lt_of_le_of_lt (le_pair_right x y) (lt_of_mem hx)⟩) @@ -375,9 +375,9 @@ lemma range_defined : 𝚺₀-Function₁ (range : V → V) via rangeDef := by @[simp] lemma range_defined_iff (v) : Semiformula.Evalbm V v rangeDef.val ↔ v 0 = range (v 1) := range_defined.df.iff v -instance range_definable : 𝚺₀-Function₁ (range : V → V) := Defined.to_definable _ range_defined +instance range_definable : 𝚺₀-Function₁ (range : V → V) := range_defined.to_definable -@[definability, simp] instance range_definable' (Γ) : DefinableFunction₁ ℒₒᵣ Γ (range : V → V) := .of_zero range_definable _ +instance range_definable' (ℌ : HierarchySymbol) : ℌ-Function₁ (range : V → V) := range_definable.of_zero @[simp] lemma range_empty : range (∅ : V) = ∅ := mem_ext (by simp [mem_range_iff]) @@ -427,9 +427,9 @@ lemma isMapping_defined : 𝚺₀-Predicate (IsMapping : V → Prop) via isMappi @[simp] lemma isMapping_defined_iff (v) : Semiformula.Evalbm V v isMappingDef.val ↔ IsMapping (v 0) := isMapping_defined.df.iff v -instance isMapping_definable : 𝚺₀-Predicate (IsMapping : V → Prop) := Defined.to_definable _ isMapping_defined +instance isMapping_definable : 𝚺₀-Predicate (IsMapping : V → Prop) := isMapping_defined.to_definable -instance isMapping_definable' (Γ) : Γ-Predicate (IsMapping : V → Prop) := .of_zero isMapping_definable _ +instance isMapping_definable' (ℌ) : ℌ-Predicate (IsMapping : V → Prop) := isMapping_definable.of_zero end @@ -513,9 +513,9 @@ lemma domain_restr_of_subset_domain {f s : V} (h : s ⊆ domain f) : domain (f end restriction -theorem insert_induction {P : V → Prop} (hP : (Γ, 1)-Predicate P) +theorem insert_induction {P : V → Prop} (hP : Γ-[1]-Predicate P) (hempty : P ∅) (hinsert : ∀ a s, a ∉ s → P s → P (insert a s)) : ∀ s, P s := - order_induction_hh ℒₒᵣ Γ 1 hP <| by + order_induction_hh Γ 1 hP <| by intro s IH rcases eq_empty_or_nonempty s with (rfl | ⟨x, hx⟩) · exact hempty @@ -605,9 +605,9 @@ lemma fstIdx_defined : 𝚺₀-Function₁ (fstIdx : V → V) via fstIdxDef := b @[simp] lemma eval_fstIdxDef (v) : Semiformula.Evalbm V v fstIdxDef.val ↔ v 0 = fstIdx (v 1) := fstIdx_defined.df.iff v -instance fstIdx_definable : 𝚺₀-Function₁ (fstIdx : V → V) := Defined.to_definable _ fstIdx_defined +instance fstIdx_definable : 𝚺₀-Function₁ (fstIdx : V → V) := fstIdx_defined.to_definable -instance fstIdx_definable' (Γ) : Γ-Function₁ (fstIdx : V → V) := .of_zero fstIdx_definable _ +instance fstIdx_definable' (Γ) : Γ-Function₁ (fstIdx : V → V) := fstIdx_definable.of_zero end fstIdx @@ -629,9 +629,9 @@ lemma sndIdx_defined : 𝚺₀-Function₁ (sndIdx : V → V) via sndIdxDef := b @[simp] lemma eval_sndIdxDef (v) : Semiformula.Evalbm V v sndIdxDef.val ↔ v 0 = sndIdx (v 1) := sndIdx_defined.df.iff v -instance sndIdx_definable : 𝚺₀-Function₁ (sndIdx : V → V) := Defined.to_definable _ sndIdx_defined +instance sndIdx_definable : 𝚺₀-Function₁ (sndIdx : V → V) := sndIdx_defined.to_definable -instance sndIdx_definable' (Γ) : Γ-Function₁ (sndIdx : V → V) := .of_zero sndIdx_definable _ +instance sndIdx_definable' (Γ) : Γ-Function₁ (sndIdx : V → V) := sndIdx_definable.of_zero end sndIdx diff --git a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean index 4a5ff9d..9b18529 100644 --- a/Arithmetization/ISigmaOne/HFS/Fixpoint.lean +++ b/Arithmetization/ISigmaOne/HFS/Fixpoint.lean @@ -12,72 +12,72 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] namespace Fixpoint structure Blueprint (k : ℕ) where - core : 𝚫₁-Semisentence (k + 2) + core : 𝚫₁.Semisentence (k + 2) namespace Blueprint variable {k} (φ : Blueprint k) -instance : Coe (Blueprint k) (𝚫₁-Semisentence (k + 2)) := ⟨Blueprint.core⟩ +instance : Coe (Blueprint k) (𝚫₁.Semisentence (k + 2)) := ⟨Blueprint.core⟩ -def succDef : 𝚺₁-Semisentence (k + 3) := .mkSigma +def succDef : 𝚺₁.Semisentence (k + 3) := .mkSigma “u ih s | ∀ x < u + (s + 1), (x ∈ u → x ≤ s ∧ !φ.core.sigma x ih ⋯) ∧ (x ≤ s ∧ !φ.core.pi x ih ⋯ → x ∈ u)” (by simp) def prBlueprint : PR.Blueprint k where zero := .mkSigma “x | x = 0” (by simp) succ := φ.succDef -def limSeqDef : 𝚺₁-Semisentence (k + 2) := (φ.prBlueprint).resultDef +def limSeqDef : 𝚺₁.Semisentence (k + 2) := (φ.prBlueprint).resultDef -def fixpointDef : 𝚺₁-Semisentence (k + 1) := +def fixpointDef : 𝚺₁.Semisentence (k + 1) := .mkSigma “x | ∃ s L, !φ.limSeqDef L s ⋯ ∧ x ∈ L” (by simp) -def fixpointDefΔ₁ : 𝚫₁-Semisentence (k + 1) := .mkDelta +def fixpointDefΔ₁ : 𝚫₁.Semisentence (k + 1) := .mkDelta (.mkSigma “x | ∃ L, !φ.limSeqDef L (x + 1) ⋯ ∧ x ∈ L” (by simp)) (.mkPi “x | ∀ L, !φ.limSeqDef L (x + 1) ⋯ → x ∈ L” (by simp)) end Blueprint -variable (M) +variable (V) structure Construction {k : ℕ} (φ : Blueprint k) where - Φ : (Fin k → M) → Set M → M → Prop - defined : Arith.Defined (fun v ↦ Φ (v ·.succ.succ) {x | x ∈ v 1} (v 0)) φ.core - monotone {C C' : Set M} (h : C ⊆ C') {v x} : Φ v C x → Φ v C' x + Φ : (Fin k → V) → Set V → V → Prop + defined : 𝚫₁.Defined (fun v ↦ Φ (v ·.succ.succ) {x | x ∈ v 1} (v 0)) φ.core + monotone {C C' : Set V} (h : C ⊆ C') {v x} : Φ v C x → Φ v C' x -class Construction.Finite {k : ℕ} {φ : Blueprint k} (c : Construction M φ) where - finite {C : Set M} {v x} : c.Φ v C x → ∃ m, c.Φ v {y ∈ C | y < m} x +class Construction.Finite {k : ℕ} {φ : Blueprint k} (c : Construction V φ) where + finite {C : Set V} {v x} : c.Φ v C x → ∃ m, c.Φ v {y ∈ C | y < m} x -class Construction.StrongFinite {k : ℕ} {φ : Blueprint k} (c : Construction M φ) where - strong_finite {C : Set M} {v x} : c.Φ v C x → c.Φ v {y ∈ C | y < x} x +class Construction.StrongFinite {k : ℕ} {φ : Blueprint k} (c : Construction V φ) where + strong_finite {C : Set V} {v x} : c.Φ v C x → c.Φ v {y ∈ C | y < x} x -instance {k : ℕ} {φ : Blueprint k} (c : Construction M φ) [c.StrongFinite] : c.Finite where +instance {k : ℕ} {φ : Blueprint k} (c : Construction V φ) [c.StrongFinite] : c.Finite where finite {_ _ x} := fun h ↦ ⟨x, Construction.StrongFinite.strong_finite h⟩ -variable {M} +variable {V} namespace Construction -variable {k : ℕ} {φ : Blueprint k} (c : Construction M φ) (v : Fin k → M) +variable {k : ℕ} {φ : Blueprint k} (c : Construction V φ) (v : Fin k → V) -lemma eval_formula (v : Fin k.succ.succ → M) : - Semiformula.Evalbm M v (HSemiformula.val φ.core) ↔ c.Φ (v ·.succ.succ) {x | x ∈ v 1} (v 0) := c.defined.df.iff v +lemma eval_formula (v : Fin k.succ.succ → V) : + Semiformula.Evalbm V v φ.core.val ↔ c.Φ (v ·.succ.succ) {x | x ∈ v 1} (v 0) := c.defined.df.iff v -lemma succ_existsUnique (s ih : M) : - ∃! u : M, ∀ x, (x ∈ u ↔ x ≤ s ∧ c.Φ v {z | z ∈ ih} x) := by +lemma succ_existsUnique (s ih : V) : + ∃! u : V, ∀ x, (x ∈ u ↔ x ≤ s ∧ c.Φ v {z | z ∈ ih} x) := by have : 𝚺₁-Predicate fun x ↦ x ≤ s ∧ c.Φ v {z | z ∈ ih} x := by - apply Definable.and (by definability) + apply HierarchySymbol.Boldface.and (by definability) ⟨φ.core.sigma.rew <| Rew.embSubsts (#0 :> &ih :> fun i ↦ &(v i)), - by intro x; simp [HSemiformula.val_sigma, c.eval_formula]⟩ + by intro x; simp [HierarchySymbol.Semiformula.val_sigma, c.eval_formula]⟩ exact finite_comprehension₁! this ⟨s + 1, fun i ↦ by rintro ⟨hi, _⟩; exact lt_succ_iff_le.mpr hi⟩ -def succ (s ih : M) : M := Classical.choose! (c.succ_existsUnique v s ih) +def succ (s ih : V) : V := Classical.choose! (c.succ_existsUnique v s ih) variable {v} @@ -94,16 +94,16 @@ private lemma succ_graph {u v s ih} : exact h x (lt_of_lt_of_le (lt_succ_iff_le.mpr (c.mem_succ_iff.mp hx).1) (by simp)) |>.mpr (c.mem_succ_iff.mp hx)⟩ -lemma succ_defined : DefinedFunction (fun v : Fin (k + 2) → M ↦ c.succ (v ·.succ.succ) (v 1) (v 0)) φ.succDef := by +lemma succ_defined : 𝚺₁.DefinedFunction (fun v : Fin (k + 2) → V ↦ c.succ (v ·.succ.succ) (v 1) (v 0)) φ.succDef := by intro v - simp [Blueprint.succDef, succ_graph, HSemiformula.val_sigma, c.eval_formula, + simp [Blueprint.succDef, succ_graph, HierarchySymbol.Semiformula.val_sigma, c.eval_formula, c.defined.proper.iff', -and_imp, ←iff_iff_implies_and_implies] rfl lemma eval_succDef (v) : - Semiformula.Evalbm M v φ.succDef.val ↔ v 0 = c.succ (v ·.succ.succ.succ) (v 2) (v 1) := c.succ_defined.df.iff v + Semiformula.Evalbm V v φ.succDef.val ↔ v 0 = c.succ (v ·.succ.succ.succ) (v 2) (v 1) := c.succ_defined.df.iff v -def prConstruction : PR.Construction M φ.prBlueprint where +def prConstruction : PR.Construction V φ.prBlueprint where zero := fun _ ↦ ∅ succ := c.succ zero_defined := by intro v; simp [Blueprint.prBlueprint, emptyset_def] @@ -111,34 +111,34 @@ def prConstruction : PR.Construction M φ.prBlueprint where variable (v) -def limSeq (s : M) : M := c.prConstruction.result v s +def limSeq (s : V) : V := c.prConstruction.result v s variable {v} @[simp] lemma limSeq_zero : c.limSeq v 0 = ∅ := by simp [limSeq, prConstruction] -lemma limSeq_succ (s : M) : c.limSeq v (s + 1) = c.succ v s (c.limSeq v s) := by simp [limSeq, prConstruction] +lemma limSeq_succ (s : V) : c.limSeq v (s + 1) = c.succ v s (c.limSeq v s) := by simp [limSeq, prConstruction] -lemma termSet_defined : DefinedFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) φ.limSeqDef := +lemma termSet_defined : 𝚺₁.DefinedFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) φ.limSeqDef := fun v ↦ by simp [c.prConstruction.result_defined_iff, Blueprint.limSeqDef]; rfl @[simp] lemma eval_limSeqDef (v) : - Semiformula.Evalbm M v φ.limSeqDef.val ↔ v 0 = c.limSeq (v ·.succ.succ) (v 1) := c.termSet_defined.df.iff v + Semiformula.Evalbm V v φ.limSeqDef.val ↔ v 0 = c.limSeq (v ·.succ.succ) (v 1) := c.termSet_defined.df.iff v instance limSeq_definable : - DefinableFunction ℒₒᵣ 𝚺₁ (fun v ↦ c.limSeq (v ·.succ) (v 0)) := Defined.to_definable _ c.termSet_defined + 𝚺₁.BoldfaceFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) := c.termSet_defined.to_definable @[simp, definability] instance limSeq_definable' (Γ) : - DefinableFunction ℒₒᵣ (Γ, m + 1) (fun v ↦ c.limSeq (v ·.succ) (v 0)) := + Γ-[m + 1].BoldfaceFunction (fun v ↦ c.limSeq (v ·.succ) (v 0)) := .of_sigmaOne c.limSeq_definable _ _ -lemma mem_limSeq_succ_iff {x s : M} : +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] -lemma limSeq_cumulative {s s' : M} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' := by +lemma limSeq_cumulative {s s' : V} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' := by induction s' using induction_sigma1 generalizing s - · apply Definable.ball_le' (by definability) - apply Definable.comp₂_infer + · apply HierarchySymbol.Boldface.ball_le (by definability) + apply HierarchySymbol.Boldface.comp₂ · exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ · exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #2 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ case zero => @@ -151,14 +151,14 @@ lemma limSeq_cumulative {s s' : M} : s ≤ s' → c.limSeq v s ⊆ c.limSeq v s' rcases c.mem_limSeq_succ_iff.mp hu with ⟨hu, Hu⟩ exact c.mem_limSeq_succ_iff.mpr ⟨_root_.le_trans hu hs, c.monotone (fun z hz ↦ ih hs hz) Hu⟩ -lemma mem_limSeq_self [c.StrongFinite] {u s : M} : +lemma mem_limSeq_self [c.StrongFinite] {u s : V} : u ∈ c.limSeq v s → u ∈ c.limSeq v (u + 1) := by induction u using order_induction_pi1 generalizing s - · apply Definable.all - apply Definable.imp - · apply Definable.comp₂_infer (by definability) + · apply HierarchySymbol.Boldface.all + apply HierarchySymbol.Boldface.imp + · apply HierarchySymbol.Boldface.comp₂ (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ - · apply Definable.comp₂_infer (by definability) + · apply HierarchySymbol.Boldface.comp₂ (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> ‘#2 + 1’ :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩ case ind u ih => rcases zero_or_succ s with (rfl | ⟨s, rfl⟩) @@ -176,32 +176,32 @@ lemma mem_limSeq_self [c.StrongFinite] {u s : M} : variable (v) -def Fixpoint (x : M) : Prop := ∃ s, x ∈ c.limSeq v s +def Fixpoint (x : V) : Prop := ∃ s, x ∈ c.limSeq v s variable {v} -lemma fixpoint_iff [c.StrongFinite] {x : M} : c.Fixpoint v x ↔ x ∈ c.limSeq v (x + 1) := +lemma fixpoint_iff [c.StrongFinite] {x : V} : c.Fixpoint v x ↔ x ∈ c.limSeq v (x + 1) := ⟨by rintro ⟨s, hs⟩; exact c.mem_limSeq_self hs, fun h ↦ ⟨x + 1, h⟩⟩ -lemma fixpoint_iff_succ {x : M} : c.Fixpoint v x ↔ ∃ u, x ∈ c.limSeq v (u + 1) := +lemma fixpoint_iff_succ {x : V} : c.Fixpoint v x ↔ ∃ u, x ∈ c.limSeq v (u + 1) := ⟨by rintro ⟨u, h⟩ rcases zero_or_succ u with (rfl | ⟨u, rfl⟩) · simp at h · exact ⟨u, h⟩, by rintro ⟨u, h⟩; exact ⟨u + 1, h⟩⟩ -lemma finite_upperbound (m : M) : ∃ s, ∀ z < m, c.Fixpoint v z → z ∈ c.limSeq v s := by - have : ∃ F : M, ∀ x, x ∈ F ↔ x < m ∧ c.Fixpoint v x := by +lemma finite_upperbound (m : V) : ∃ s, ∀ z < m, c.Fixpoint v z → z ∈ c.limSeq v s := by + have : ∃ F : V, ∀ x, x ∈ F ↔ x < m ∧ c.Fixpoint v x := by have : 𝚺₁-Predicate fun x ↦ x < m ∧ c.Fixpoint v x := - Definable.and (by definability) - (Definable.ex (Definable.comp₂_infer (by definability) + HierarchySymbol.Boldface.and (by definability) + (HierarchySymbol.Boldface.ex (HierarchySymbol.Boldface.comp₂ (by definability) ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #1 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩)) exact finite_comprehension₁! this ⟨m, fun i hi ↦ hi.1⟩ |>.exists rcases this with ⟨F, hF⟩ have : ∀ x ∈ F, ∃ u, x ∈ c.limSeq v u := by intro x hx; exact hF x |>.mp hx |>.2 - have : ∃ f, IsMapping f ∧ domain f = F ∧ ∀ (x y : M), ⟪x, y⟫ ∈ f → x ∈ c.limSeq v y := sigmaOne_skolem - (by apply Definable.comp₂_infer (by definability) + have : ∃ f, IsMapping f ∧ domain f = F ∧ ∀ (x y : V), ⟪x, y⟫ ∈ f → x ∈ c.limSeq v y := sigmaOne_skolem + (by apply HierarchySymbol.Boldface.comp₂ (by definability) exact ⟨φ.limSeqDef.rew <| Rew.embSubsts (#0 :> #2 :> fun i ↦ &(v i)), by intro v; simp [c.eval_limSeqDef]⟩) this rcases this with ⟨f, mf, rfl, hf⟩ exact ⟨f, by @@ -231,51 +231,37 @@ theorem case [c.Finite] : c.Fixpoint v x ↔ c.Φ v {z | c.Fixpoint v z} x := section -lemma fixpoint_defined : Arith.Defined (fun v ↦ c.Fixpoint (v ·.succ) (v 0)) φ.fixpointDef := by +lemma fixpoint_defined : 𝚺₁.Defined (fun v ↦ c.Fixpoint (v ·.succ) (v 0)) φ.fixpointDef := by intro v; simp [Blueprint.fixpointDef, c.eval_limSeqDef]; rfl @[simp] lemma eval_fixpointDef (v) : - Semiformula.Evalbm M v φ.fixpointDef.val ↔ c.Fixpoint (v ·.succ) (v 0) := c.fixpoint_defined.df.iff v + Semiformula.Evalbm V v φ.fixpointDef.val ↔ c.Fixpoint (v ·.succ) (v 0) := c.fixpoint_defined.df.iff v -lemma fixpoint_definedΔ₁ [c.StrongFinite] : Arith.Defined (fun v ↦ c.Fixpoint (v ·.succ) (v 0)) φ.fixpointDefΔ₁ := +lemma fixpoint_definedΔ₁ [c.StrongFinite] : 𝚫₁.Defined (fun v ↦ c.Fixpoint (v ·.succ) (v 0)) φ.fixpointDefΔ₁ := ⟨by intro v; simp [Blueprint.fixpointDefΔ₁, c.eval_limSeqDef], by intro v; simp [Blueprint.fixpointDefΔ₁, c.eval_limSeqDef, fixpoint_iff]⟩ @[simp] lemma eval_fixpointDefΔ₁ [c.StrongFinite] (v) : - Semiformula.Evalbm M v φ.fixpointDefΔ₁.val ↔ c.Fixpoint (v ·.succ) (v 0) := c.fixpoint_definedΔ₁.df.iff v + Semiformula.Evalbm V v φ.fixpointDefΔ₁.val ↔ c.Fixpoint (v ·.succ) (v 0) := c.fixpoint_definedΔ₁.df.iff v end -theorem induction [c.StrongFinite] {P : M → Prop} (hP : (Γ, 1)-Predicate P) - (H : ∀ C : Set M, (∀ x ∈ C, c.Fixpoint v x ∧ P x) → ∀ x, c.Φ v C x → P x) : +theorem induction [c.StrongFinite] {P : V → Prop} (hP : Γ-[1]-Predicate P) + (H : ∀ C : Set V, (∀ x ∈ C, c.Fixpoint v x ∧ P x) → ∀ x, c.Φ v C x → P x) : ∀ x, c.Fixpoint v x → P x := by - apply @order_induction_hh M _ _ _ _ _ _ ℒₒᵣ _ _ _ _ Γ 1 _ - · apply Definable.imp - (Definable.comp₁ (by definability) (by - apply Definable.of_deltaOne - exact ⟨φ.fixpointDefΔ₁.rew <| Rew.embSubsts <| #0 :> fun x ↦ &(v x), c.fixpoint_definedΔ₁.proper.rew' _, - by intro v; simp [c.eval_fixpointDefΔ₁]⟩)) + apply order_induction_hh (Γ := Γ) (m := 1) (P := fun x ↦ c.Fixpoint v x → P x) + · apply HierarchySymbol.Boldface.imp + (HierarchySymbol.BoldfacePred.comp + (by + apply HierarchySymbol.Boldface.of_deltaOne + exact ⟨φ.fixpointDefΔ₁.rew <| Rew.embSubsts <| #0 :> fun x ↦ &(v x), c.fixpoint_definedΔ₁.proper.rew' _, + by intro v; simp [c.eval_fixpointDefΔ₁]⟩) + (by definability)) (by definability) intro x ih hx have : c.Φ v {y | c.Fixpoint v y ∧ y < x} x := StrongFinite.strong_finite (c.case.mp hx) exact H {y | c.Fixpoint v y ∧ y < x} (by intro y ⟨hy, hyx⟩; exact ⟨hy, ih y hyx hy⟩) x this -/- -theorem pi₁_induction [c.Finite] {P : M → Prop} (hP : 𝚷₁-Predicate P) - (H : ∀ C : Set M, (∀ x ∈ C, c.Fixpoint v x ∧ P x) → ∀ x, c.Φ v C x → P x) : - ∀ x, c.Fixpoint v x → P x := by - apply @order_induction_hh M _ _ _ _ _ _ ℒₒᵣ _ _ _ _ 𝚷 1 _ - · apply Definable.imp - (Definable.comp₁ (by definability) (by - exact ⟨φ.fixpointDef.rew <| Rew.embSubsts <| #0 :> fun x ↦ &(v x), - by intro v; simp [c.eval_fixpointDef]⟩)) - (by definability) - intro x ih hx - have := (c.case.mp hx) - have : c.Φ v {y | c.Fixpoint v y ∧ y < x} x := StrongFinite.strong_finite (c.case.mp hx) - exact H {y | c.Fixpoint v y ∧ y < x} (by intro y ⟨hy, hyx⟩; exact ⟨hy, ih y hyx hy⟩) x this --/ - end Construction end Fixpoint diff --git a/Arithmetization/ISigmaOne/HFS/PRF.lean b/Arithmetization/ISigmaOne/HFS/PRF.lean index 247d0a9..dbd0955 100644 --- a/Arithmetization/ISigmaOne/HFS/PRF.lean +++ b/Arithmetization/ISigmaOne/HFS/PRF.lean @@ -12,15 +12,15 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] namespace PR structure Blueprint (k : ℕ) where - zero : HSemisentence ℒₒᵣ (k + 1) 𝚺₁ - succ : HSemisentence ℒₒᵣ (k + 3) 𝚺₁ + zero : 𝚺₁.Semisentence (k + 1) + succ : 𝚺₁.Semisentence (k + 3) -def Blueprint.cseqDef (p : Blueprint k) : HSemisentence ℒₒᵣ (k + 1) 𝚺₁ := .mkSigma +def Blueprint.cseqDef (p : Blueprint k) : 𝚺₁.Semisentence (k + 1) := .mkSigma “s | :Seq s ∧ (∃ z < s, !p.zero z ⋯ ∧ 0 ~[s] z) @@ -28,28 +28,28 @@ def Blueprint.cseqDef (p : Blueprint k) : HSemisentence ℒₒᵣ (k + 1) 𝚺 (∃ l <⁺ 2 * s, !lhDef l s ∧ i + 1 < l) → ∀ z < s, i ~[s] z → ∃ u < s, !p.succ u z i ⋯ ∧ i + 1 ~[s] u)” (by simp) -def Blueprint.resultDef (p : Blueprint k) : HSemisentence ℒₒᵣ (k + 2) 𝚺₁ := .mkSigma +def Blueprint.resultDef (p : Blueprint k) : 𝚺₁.Semisentence (k + 2) := .mkSigma “z u | ∃ s, !p.cseqDef s ⋯ ∧ u ~[s] z” (by simp) -def Blueprint.resultDeltaDef (p : Blueprint k) : HSemisentence ℒₒᵣ (k + 2) 𝚫₁ := p.resultDef.graphDelta +def Blueprint.resultDeltaDef (p : Blueprint k) : 𝚫₁.Semisentence (k + 2) := p.resultDef.graphDelta -variable (M) +variable (V) structure Construction {k : ℕ} (p : Blueprint k) where - zero : (Fin k → M) → M - succ : (Fin k → M) → M → M → M - zero_defined : DefinedFunction zero p.zero - succ_defined : DefinedFunction (fun v ↦ succ (v ·.succ.succ) (v 1) (v 0)) p.succ + zero : (Fin k → V) → V + succ : (Fin k → V) → V → V → V + zero_defined : 𝚺₁.DefinedFunction zero p.zero + succ_defined : 𝚺₁.DefinedFunction (fun v ↦ succ (v ·.succ.succ) (v 1) (v 0)) p.succ -variable {M} +variable {V} namespace Construction -variable {k : ℕ} {p : Blueprint k} (c : Construction M p) (v : Fin k → M) +variable {k : ℕ} {p : Blueprint k} (c : Construction V p) (v : Fin k → V) -def CSeq (s : M) : Prop := Seq s ∧ ⟪0, c.zero v⟫ ∈ s ∧ ∀ i < lh s - 1, ∀ z, ⟪i, z⟫ ∈ s → ⟪i + 1, c.succ v i z⟫ ∈ s +def CSeq (s : V) : Prop := Seq s ∧ ⟪0, c.zero v⟫ ∈ s ∧ ∀ i < lh s - 1, ∀ z, ⟪i, z⟫ ∈ s → ⟪i + 1, c.succ v i z⟫ ∈ s -private lemma cseq_iff (s : M) : c.CSeq v s ↔ +private lemma cseq_iff (s : V) : c.CSeq v s ↔ Seq s ∧ (∃ z < s, z = c.zero v ∧ ⟪0, z⟫ ∈ s) ∧ (∀ i < 2 * s, @@ -66,17 +66,17 @@ private lemma cseq_iff (s : M) : c.CSeq v s ↔ ⟨lh s, by simp [lt_succ_iff_le], rfl, by simpa [lt_tsub_iff_right] using hi⟩ z (lt_of_mem_rng hiz) hiz with ⟨_, _, rfl, h⟩ exact h⟩⟩ -lemma cseq_defined : Arith.Defined (fun v ↦ c.CSeq (v ·.succ) (v 0) : (Fin (k + 1) → M) → Prop) p.cseqDef := by +lemma cseq_defined : 𝚺₁.Defined (fun v ↦ c.CSeq (v ·.succ) (v 0) : (Fin (k + 1) → V) → Prop) p.cseqDef := by intro v; simp [Blueprint.cseqDef, cseq_iff, c.zero_defined.df.iff, c.succ_defined.df.iff] @[simp] lemma cseq_defined_iff (v) : - Semiformula.Evalbm M v p.cseqDef.val ↔ c.CSeq (v ·.succ) (v 0) := c.cseq_defined.df.iff v + Semiformula.Evalbm V v p.cseqDef.val ↔ c.CSeq (v ·.succ) (v 0) := c.cseq_defined.df.iff v variable {c v} namespace CSeq -variable {s : M} (h : c.CSeq v s) +variable {s : V} (h : c.CSeq v s) lemma seq : Seq s := h.1 @@ -84,7 +84,7 @@ lemma zero : ⟪0, c.zero v⟫ ∈ s := h.2.1 lemma succ : ∀ i < lh s - 1, ∀ z, ⟪i, z⟫ ∈ s → ⟪i + 1, c.succ v i z⟫ ∈ s := h.2.2 -lemma unique {s₁ s₂ : M} (H₁ : c.CSeq v s₁) (H₂ : c.CSeq v s₂) (h₁₂ : lh s₁ ≤ lh s₂) {i} (hi : i < lh s₁) {z₁ z₂} : +lemma unique {s₁ s₂ : V} (H₁ : c.CSeq v s₁) (H₂ : c.CSeq v s₂) (h₁₂ : lh s₁ ≤ lh s₂) {i} (hi : i < lh s₁) {z₁ z₂} : ⟪i, z₁⟫ ∈ s₁ → ⟪i, z₂⟫ ∈ s₂ → z₁ = z₂ := by revert z₁ z₂ suffices ∀ z₁ < s₁, ∀ z₂ < s₂, ⟪i, z₁⟫ ∈ s₁ → ⟪i, z₂⟫ ∈ s₂ → z₁ = z₂ @@ -115,7 +115,7 @@ end CSeq lemma CSeq.initial : c.CSeq v !⟦c.zero v⟧ := ⟨by simp, by simp [seqCons], by simp⟩ -lemma CSeq.successor {s l z : M} (Hs : c.CSeq v s) (hl : l + 1 = lh s) (hz : ⟪l, z⟫ ∈ s) : +lemma CSeq.successor {s l z : V} (Hs : c.CSeq v s) (hl : l + 1 = lh s) (hz : ⟪l, z⟫ ∈ s) : c.CSeq v (s ⁀' c.succ v l z) := ⟨ Hs.seq.seqCons _, by simp [seqCons, Hs.zero], by simp [Hs.seq.lh_seqCons] @@ -133,10 +133,10 @@ lemma CSeq.successor {s l z : M} (Hs : c.CSeq v s) (hl : l + 1 = lh s) (hz : ⟪ variable (c v) -lemma CSeq.exists (l : M) : ∃ s, c.CSeq v s ∧ l + 1 = lh s := by +lemma CSeq.exists (l : V) : ∃ s, c.CSeq v s ∧ l + 1 = lh s := by induction l using induction_sigma1 - · apply Definable.ex - apply Definable.and + · apply HierarchySymbol.Boldface.ex + apply HierarchySymbol.Boldface.and · exact ⟨p.cseqDef.rew (Rew.embSubsts <| #0 :> fun i ↦ &(v i)), by intro w; simpa using c.cseq_defined_iff (w 0 :> v) |>.symm⟩ · definability @@ -149,7 +149,7 @@ lemma CSeq.exists (l : M) : ∃ s, c.CSeq v s ∧ l + 1 = lh s := by rcases this with ⟨z, hz⟩ exact ⟨s ⁀' c.succ v l z, Hs.successor hls hz, by simp [Hs.seq, hls]⟩ -lemma cSeq_result_existsUnique (l : M) : ∃! z, ∃ s, c.CSeq v s ∧ l + 1 = lh s ∧ ⟪l, z⟫ ∈ s := by +lemma cSeq_result_existsUnique (l : V) : ∃! z, ∃ s, c.CSeq v s ∧ l + 1 = lh s ∧ ⟪l, z⟫ ∈ s := by rcases CSeq.exists c v l with ⟨s, Hs, h⟩ have : ∃ z, ⟪l, z⟫ ∈ s := Hs.seq.exists (show l < lh s from by simp [←h]) rcases this with ⟨z, hz⟩ @@ -157,23 +157,23 @@ lemma cSeq_result_existsUnique (l : M) : ∃! z, ∃ s, c.CSeq v s ∧ l + 1 = l rintro z' ⟨s', Hs', h', hz'⟩ exact Eq.symm <| Hs.unique Hs' (by simp [←h, ←h']) (show l < lh s from by simp [←h]) hz hz') -def result (u : M) : M := Classical.choose! (c.cSeq_result_existsUnique v u) +def result (u : V) : V := Classical.choose! (c.cSeq_result_existsUnique v u) -lemma result_spec (u : M) : ∃ s, c.CSeq v s ∧ u + 1 = lh s ∧ ⟪u, c.result v u⟫ ∈ s := +lemma result_spec (u : V) : ∃ s, c.CSeq v s ∧ u + 1 = lh s ∧ ⟪u, c.result v u⟫ ∈ s := Classical.choose!_spec (c.cSeq_result_existsUnique v u) @[simp] theorem result_zero : c.result v 0 = c.zero v := by rcases c.result_spec v 0 with ⟨s, Hs, _, h0⟩ exact Hs.seq.isMapping.uniq h0 Hs.zero -@[simp] theorem result_succ (u : M) : c.result v (u + 1) = c.succ v u (c.result v u) := by +@[simp] theorem result_succ (u : V) : c.result v (u + 1) = c.succ v u (c.result v u) := by rcases c.result_spec v u with ⟨s, Hs, hk, h⟩ have : CSeq c v (s ⁀' c.succ v u (result c v u) ) := Hs.successor hk h exact Eq.symm <| Classical.choose_uniq (c.cSeq_result_existsUnique v (u + 1)) ⟨_, this, by simp [Hs.seq, hk], by simp [hk]⟩ -lemma result_graph (z u : M) : z = c.result v u ↔ ∃ s, c.CSeq v s ∧ ⟪u, z⟫ ∈ s := +lemma result_graph (z u : V) : z = c.result v u ↔ ∃ s, c.CSeq v s ∧ ⟪u, z⟫ ∈ s := ⟨by rintro rfl rcases c.result_spec v u with ⟨s, Hs, _, h⟩ exact ⟨s, Hs, h⟩, @@ -183,22 +183,22 @@ lemma result_graph (z u : M) : z = c.result v u ↔ ∃ s, c.CSeq v s ∧ ⟪u, (by simp [←hu, succ_le_iff_lt]; exact Hs.seq.lt_lh_iff.mpr (mem_domain_of_pair_mem h)) (by simp [←hu]) h' h⟩ -lemma result_defined : Arith.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → M) → M) p.resultDef := by +lemma result_defined : 𝚺₁.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → V) → V) p.resultDef := by intro v; simp [Blueprint.resultDef, result_graph] apply exists_congr; intro x simp [c.cseq_defined_iff]; intros; rfl -lemma result_defined_delta : Arith.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → M) → M) p.resultDeltaDef := +lemma result_defined_delta : 𝚫₁.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → V) → V) p.resultDeltaDef := c.result_defined.graph_delta @[simp] lemma result_defined_iff (v) : - Semiformula.Evalbm M v p.resultDef.val ↔ v 0 = c.result (v ·.succ.succ) (v 1) := c.result_defined.df.iff v + Semiformula.Evalbm V v p.resultDef.val ↔ v 0 = c.result (v ·.succ.succ) (v 1) := c.result_defined.df.iff v -instance result_definable : DefinableFunction ℒₒᵣ 𝚺₁ (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → M) → M) := - Defined.to_definable _ c.result_defined +instance result_definable : 𝚺₁.BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → V) → V) := + c.result_defined.to_definable -instance result_definable_delta₁ : DefinableFunction ℒₒᵣ 𝚫₁ (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → M) → M) := - Defined.to_definable _ c.result_defined_delta +instance result_definable_delta₁ : 𝚫₁.BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0) : (Fin (k + 1) → V) → V) := + c.result_defined_delta.to_definable end Construction diff --git a/Arithmetization/ISigmaOne/HFS/Seq.lean b/Arithmetization/ISigmaOne/HFS/Seq.lean index 6c5373f..b2dbded 100644 --- a/Arithmetization/ISigmaOne/HFS/Seq.lean +++ b/Arithmetization/ISigmaOne/HFS/Seq.lean @@ -12,13 +12,13 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] -def Seq (s : M) : Prop := IsMapping s ∧ ∃ l, domain s = under l +def Seq (s : V) : Prop := IsMapping s ∧ ∃ l, domain s = under l -def Seq.isMapping {s : M} (h : Seq s) : IsMapping s := h.1 +def Seq.isMapping {s : V} (h : Seq s) : IsMapping s := h.1 -private lemma seq_iff (s : M) : Seq s ↔ IsMapping s ∧ ∃ l ≤ 2 * s, ∃ d ≤ 2 * s, d = domain s ∧ d = under l := +private lemma seq_iff (s : V) : Seq s ↔ IsMapping s ∧ ∃ l ≤ 2 * s, ∃ d ≤ 2 * s, d = domain s ∧ d = under l := ⟨by rintro ⟨hs, l, h⟩ exact ⟨hs, l, (by calc @@ -29,15 +29,15 @@ private lemma seq_iff (s : M) : Seq s ↔ IsMapping s ∧ ∃ l ≤ 2 * s, ∃ d def _root_.LO.FirstOrder.Arith.seqDef : 𝚺₀.Semisentence 1 := .mkSigma “s | !isMappingDef s ∧ ∃ l <⁺ 2 * s, ∃ d <⁺ 2 * s, !domainDef d s ∧ !underDef d l” (by simp) -lemma seq_defined : 𝚺₀-Predicate (Seq : M → Prop) via seqDef := by +lemma seq_defined : 𝚺₀-Predicate (Seq : V → Prop) via seqDef := by intro v; simp [seqDef, seq_iff] @[simp] lemma seq_defined_iff (v) : - Semiformula.Evalbm M v seqDef.val ↔ Seq (v 0) := seq_defined.df.iff v + Semiformula.Evalbm V v seqDef.val ↔ Seq (v 0) := seq_defined.df.iff v -instance seq_definable : 𝚺₀-Predicate (Seq : M → Prop) := Defined.to_definable _ seq_defined +instance seq_definable : 𝚺₀-Predicate (Seq : V → Prop) := seq_defined.to_definable -@[simp, definability] instance seq_definable' (Γ) : Γ-Predicate (Seq : M → Prop) := .of_zero seq_definable _ +instance seq_definable' (ℌ) : ℌ-Predicate (Seq : V → Prop) := seq_definable.of_zero section @@ -51,7 +51,7 @@ scoped macro_rules end -lemma lh_exists_uniq (s : M) : ∃! l, (Seq s → domain s = under l) ∧ (¬Seq s → l = 0) := by +lemma lh_exists_uniq (s : V) : ∃! l, (Seq s → domain s = under l) ∧ (¬Seq s → l = 0) := by by_cases h : Seq s · rcases h with ⟨h, l, hl⟩ exact ExistsUnique.intro l @@ -59,22 +59,22 @@ lemma lh_exists_uniq (s : M) : ∃! l, (Seq s → domain s = under l) ∧ (¬Seq (by simp [show Seq s from ⟨h, l, hl⟩, hl]) · simp [h] -def lh (s : M) : M := Classical.choose! (lh_exists_uniq s) +def lh (s : V) : V := Classical.choose! (lh_exists_uniq s) -lemma lh_prop (s : M) : (Seq s → domain s = under (lh s)) ∧ (¬Seq s → lh s = 0) := Classical.choose!_spec (lh_exists_uniq s) +lemma lh_prop (s : V) : (Seq s → domain s = under (lh s)) ∧ (¬Seq s → lh s = 0) := Classical.choose!_spec (lh_exists_uniq s) -lemma lh_prop_of_not_seq {s : M} (h : ¬Seq s) : lh s = 0 := (lh_prop s).2 h +lemma lh_prop_of_not_seq {s : V} (h : ¬Seq s) : lh s = 0 := (lh_prop s).2 h -lemma Seq.domain_eq {s : M} (h : Seq s) : domain s = under (lh s) := (lh_prop s).1 h +lemma Seq.domain_eq {s : V} (h : Seq s) : domain s = under (lh s) := (lh_prop s).1 h -@[simp] lemma lh_bound (s : M) : lh s ≤ 2 * s := by +@[simp] lemma lh_bound (s : V) : lh s ≤ 2 * s := by by_cases hs : Seq s · calc lh s ≤ under (lh s) := le_under _ _ ≤ 2 * s := by simp [←hs.domain_eq] · simp [lh_prop_of_not_seq hs] -private lemma lh_graph (l s : M) : l = lh s ↔ (Seq s → ∃ d ≤ 2 * s, d = domain s ∧ d = under l) ∧ (¬Seq s → l = 0) := +private lemma lh_graph (l s : V) : l = lh s ↔ (Seq s → ∃ d ≤ 2 * s, d = domain s ∧ d = under l) ∧ (¬Seq s → l = 0) := ⟨by rintro rfl by_cases Hs : Seq s <;> simp [Hs, ←Seq.domain_eq, lh_prop_of_not_seq], by @@ -86,75 +86,75 @@ private lemma lh_graph (l s : M) : l = lh s ↔ (Seq s → ∃ d ≤ 2 * s, d = def _root_.LO.FirstOrder.Arith.lhDef : 𝚺₀.Semisentence 2 := .mkSigma “l s | (!seqDef s → ∃ d <⁺ 2 * s, !domainDef d s ∧ !underDef d l) ∧ (¬!seqDef s → l = 0)” (by simp) -lemma lh_defined : 𝚺₀-Function₁ (lh : M → M) via lhDef := by +lemma lh_defined : 𝚺₀-Function₁ (lh : V → V) via lhDef := by intro v; simp [lhDef, -exists_eq_right_right, lh_graph] @[simp] lemma lh_defined_iff (v) : - Semiformula.Evalbm M v lhDef.val ↔ v 0 = lh (v 1) := lh_defined.df.iff v + Semiformula.Evalbm V v lhDef.val ↔ v 0 = lh (v 1) := lh_defined.df.iff v -instance lh_definable : 𝚺₀-Function₁ (lh : M → M) := Defined.to_definable _ lh_defined +instance lh_definable : 𝚺₀-Function₁ (lh : V → V) := lh_defined.to_definable -instance lh_definable' (Γ) : Γ-Function₁ (lh : M → M) := .of_zero lh_definable _ +instance lh_definable' (ℌ) : ℌ-Function₁ (lh : V → V) := lh_definable.of_zero -instance : Bounded₁ (lh : M → M) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ +instance : Bounded₁ (lh : V → V) := ⟨‘x | 2 * x’, fun _ ↦ by simp⟩ -lemma Seq.exists {s : M} (h : Seq s) {x : M} (hx : x < lh s) : ∃ y, ⟪x, y⟫ ∈ s := h.isMapping x (by simpa [h.domain_eq] using hx) |>.exists +lemma Seq.exists {s : V} (h : Seq s) {x : V} (hx : x < lh s) : ∃ y, ⟪x, y⟫ ∈ s := h.isMapping x (by simpa [h.domain_eq] using hx) |>.exists -lemma Seq.nth_exists_uniq {s : M} (h : Seq s) {x : M} (hx : x < lh s) : ∃! y, ⟪x, y⟫ ∈ s := h.isMapping x (by simpa [h.domain_eq] using hx) +lemma Seq.nth_exists_uniq {s : V} (h : Seq s) {x : V} (hx : x < lh s) : ∃! y, ⟪x, y⟫ ∈ s := h.isMapping x (by simpa [h.domain_eq] using hx) -def Seq.nth {s : M} (h : Seq s) {x : M} (hx : x < lh s) : M := Classical.choose! (h.nth_exists_uniq hx) +def Seq.nth {s : V} (h : Seq s) {x : V} (hx : x < lh s) : V := Classical.choose! (h.nth_exists_uniq hx) -@[simp] lemma Seq.nth_mem {s : M} (h : Seq s) {x : M} (hx : x < lh s) : +@[simp] lemma Seq.nth_mem {s : V} (h : Seq s) {x : V} (hx : x < lh s) : ⟪x, h.nth hx⟫ ∈ s := Classical.choose!_spec (h.nth_exists_uniq hx) -lemma Seq.nth_uniq {s : M} (h : Seq s) {x y : M} (hx : x < lh s) (hy : ⟪x, y⟫ ∈ s) : y = h.nth hx := +lemma Seq.nth_uniq {s : V} (h : Seq s) {x y : V} (hx : x < lh s) (hy : ⟪x, y⟫ ∈ s) : y = h.nth hx := (h.nth_exists_uniq hx).unique hy (by simp) -@[simp] lemma Seq.nth_lt {s : M} (h : Seq s) {x} (hx : x < lh s) : h.nth hx < s := lt_of_mem_rng (h.nth_mem hx) +@[simp] lemma Seq.nth_lt {s : V} (h : Seq s) {x} (hx : x < lh s) : h.nth hx < s := lt_of_mem_rng (h.nth_mem hx) -lemma Seq.lh_eq_of {s : M} (H : Seq s) {l} (h : domain s = under l) : lh s = l := by +lemma Seq.lh_eq_of {s : V} (H : Seq s) {l} (h : domain s = under l) : lh s = l := by simpa [H.domain_eq] using h -lemma Seq.lt_lh_iff {s : M} (h : Seq s) {i} : i < lh s ↔ i ∈ domain s := by simp [h.domain_eq] +lemma Seq.lt_lh_iff {s : V} (h : Seq s) {i} : i < lh s ↔ i ∈ domain s := by simp [h.domain_eq] -lemma Seq.lt_lh_of_mem {s : M} (h : Seq s) {i x} (hix : ⟪i, x⟫ ∈ s) : i < lh s := by simp [h.lt_lh_iff, mem_domain_iff]; exact ⟨x, hix⟩ +lemma Seq.lt_lh_of_mem {s : V} (h : Seq s) {i x} (hix : ⟪i, x⟫ ∈ s) : i < lh s := by simp [h.lt_lh_iff, mem_domain_iff]; exact ⟨x, hix⟩ -def seqCons (s x : M) : M := insert ⟪lh s, x⟫ s +def seqCons (s x : V) : V := insert ⟪lh s, x⟫ s section znth -def znth_existsUnique (s i : M) : ∃! x, (Seq s ∧ i < lh s → ⟪i, x⟫ ∈ s) ∧ (¬(Seq s ∧ i < lh s) → x = 0) := by +def znth_existsUnique (s i : V) : ∃! x, (Seq s ∧ i < lh s → ⟪i, x⟫ ∈ s) ∧ (¬(Seq s ∧ i < lh s) → x = 0) := by by_cases h : Seq s ∧ i < lh s <;> simp [h] exact h.1.nth_exists_uniq h.2 -def znth (s i : M) : M := Classical.choose! (znth_existsUnique s i) +def znth (s i : V) : V := Classical.choose! (znth_existsUnique s i) -protected lemma Seq.znth {s i : M} (h : Seq s) (hi : i < lh s) : ⟪i, znth s i⟫ ∈ s := Classical.choose!_spec (znth_existsUnique s i) |>.1 ⟨h, hi⟩ +protected lemma Seq.znth {s i : V} (h : Seq s) (hi : i < lh s) : ⟪i, znth s i⟫ ∈ s := Classical.choose!_spec (znth_existsUnique s i) |>.1 ⟨h, hi⟩ -lemma Seq.znth_eq_of_mem {s i : M} (h : Seq s) (hi : ⟪i, x⟫ ∈ s) : znth s i = x := +lemma Seq.znth_eq_of_mem {s i : V} (h : Seq s) (hi : ⟪i, x⟫ ∈ s) : znth s i = x := h.isMapping.uniq (h.znth (h.lt_lh_of_mem hi)) hi -lemma znth_prop_not {s i : M} (h : ¬Seq s ∨ lh s ≤ i) : znth s i = 0 := +lemma znth_prop_not {s i : V} (h : ¬Seq s ∨ lh s ≤ i) : znth s i = 0 := Classical.choose!_spec (znth_existsUnique s i) |>.2 (by simpa [-not_and, not_and_or] using h) def _root_.LO.FirstOrder.Arith.znthDef : 𝚺₀.Semisentence 3 := .mkSigma “x s i | ∃ l <⁺ 2 * s, !lhDef l s ∧ (:Seq s ∧ i < l → i ~[s] x) ∧ (¬(:Seq s ∧ i < l) → x = 0)” (by simp) -private lemma znth_graph {x s i : M} : x = znth s i ↔ ∃ l ≤ 2 * s, l = lh s ∧ (Seq s ∧ i < l → ⟪i, x⟫ ∈ s) ∧ (¬(Seq s ∧ i < l) → x = 0) := by +private lemma znth_graph {x s i : V} : x = znth s i ↔ ∃ l ≤ 2 * s, l = lh s ∧ (Seq s ∧ i < l → ⟪i, x⟫ ∈ s) ∧ (¬(Seq s ∧ i < l) → x = 0) := by simp [znth, Classical.choose!_eq_iff]; constructor · rintro h; exact ⟨lh s, by simp, by simp, h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -lemma znth_defined : 𝚺₀-Function₂ (znth : M → M → M) via znthDef := by +lemma znth_defined : 𝚺₀-Function₂ (znth : V → V → V) via znthDef := by intro v; - simpa [znthDef, -not_and, not_and_or] using znth_graph (M := M) + simpa [znthDef, -not_and, not_and_or] using znth_graph (V := V) @[simp] lemma eval_znthDef (v) : - Semiformula.Evalbm M v znthDef.val ↔ v 0 = znth (v 1) (v 2) := znth_defined.df.iff v + Semiformula.Evalbm V v znthDef.val ↔ v 0 = znth (v 1) (v 2) := znth_defined.df.iff v -instance znth_definable : 𝚺₀-Function₂ (znth : M → M → M) := Defined.to_definable _ znth_defined +instance znth_definable : 𝚺₀-Function₂ (znth : V → V → V) := znth_defined.to_definable -instance znth_definable' (Γ) : Γ-Function₂ (znth : M → M → M) := .of_zero znth_definable _ +instance znth_definable' (ℌ) : ℌ-Function₂ (znth : V → V → V) := znth_definable.of_zero end znth @@ -162,49 +162,49 @@ end znth infixr:67 " ⁀' " => seqCons -@[simp] lemma seq_empty : Seq (∅ : M) := ⟨by simp, 0, by simp⟩ +@[simp] lemma seq_empty : Seq (∅ : V) := ⟨by simp, 0, by simp⟩ -@[simp] lemma lh_empty : lh (∅ : M) = 0 := by - have : under (lh ∅ : M) = under 0 := by simpa using Eq.symm <| Seq.domain_eq (M := M) (s := ∅) (by simp) +@[simp] lemma lh_empty : lh (∅ : V) = 0 := by + have : under (lh ∅ : V) = under 0 := by simpa using Eq.symm <| Seq.domain_eq (V := V) (s := ∅) (by simp) exact under_inj.mp this -lemma Seq.isempty_of_lh_eq_zero {s : M} (Hs : Seq s) (h : lh s = 0) : s = ∅ := by simpa [h] using Hs.domain_eq +lemma Seq.isempty_of_lh_eq_zero {s : V} (Hs : Seq s) (h : lh s = 0) : s = ∅ := by simpa [h] using Hs.domain_eq -@[simp] lemma Seq.subset_seqCons (s x : M) : s ⊆ s ⁀' x := by simp [seqCons] +@[simp] lemma Seq.subset_seqCons (s x : V) : s ⊆ s ⁀' x := by simp [seqCons] -lemma Seq.lt_seqCons {s} (hs : Seq s) (x : M) : s < s ⁀' x := +lemma Seq.lt_seqCons {s} (hs : Seq s) (x : V) : s < s ⁀' x := lt_iff_le_and_ne.mpr <| ⟨le_of_subset <| by simp, by simp [seqCons]; intro A have : ⟪lh s, x⟫ ∈ s := by simpa [←A] using mem_insert ⟪lh s, x⟫ s simpa using hs.lt_lh_of_mem this⟩ -@[simp] lemma Seq.mem_seqCons (s x : M) : ⟪lh s, x⟫ ∈ s ⁀' x := by simp [seqCons] +@[simp] lemma Seq.mem_seqCons (s x : V) : ⟪lh s, x⟫ ∈ s ⁀' x := by simp [seqCons] -protected lemma Seq.seqCons {s : M} (h : Seq s) (x : M) : Seq (s ⁀' x) := +protected lemma Seq.seqCons {s : V} (h : Seq s) (x : V) : Seq (s ⁀' x) := ⟨h.isMapping.insert (by simp [h.domain_eq]), lh s + 1, by simp [seqCons, h.domain_eq]⟩ -@[simp] lemma Seq.lh_seqCons (x : M) {s} (h : Seq s) : lh (s ⁀' x) = lh s + 1 := by +@[simp] lemma Seq.lh_seqCons (x : V) {s} (h : Seq s) : lh (s ⁀' x) = lh s + 1 := by have : under (lh s + 1) = under (lh (s ⁀' x)) := by simpa [seqCons, h.domain_eq] using (h.seqCons x).domain_eq exact Eq.symm <| under_inj.mp this -lemma mem_seqCons_iff {i x z s : M} : ⟪i, x⟫ ∈ s ⁀' z ↔ (i = lh s ∧ x = z) ∨ ⟪i, x⟫ ∈ s := by simp [seqCons] +lemma mem_seqCons_iff {i x z s : V} : ⟪i, x⟫ ∈ s ⁀' z ↔ (i = lh s ∧ x = z) ∨ ⟪i, x⟫ ∈ s := by simp [seqCons] -@[simp] lemma lh_mem_seqCons (s z : M) : ⟪lh s, z⟫ ∈ s ⁀' z := by simp [seqCons] +@[simp] lemma lh_mem_seqCons (s z : V) : ⟪lh s, z⟫ ∈ s ⁀' z := by simp [seqCons] -@[simp] lemma lh_mem_seqCons_iff {s x z : M} (H : Seq s) : ⟪lh s, x⟫ ∈ s ⁀' z ↔ x = z := by +@[simp] lemma lh_mem_seqCons_iff {s x z : V} (H : Seq s) : ⟪lh s, x⟫ ∈ s ⁀' z ↔ x = z := by simp [seqCons] intro h; have := H.lt_lh_of_mem h; simp at this -lemma Seq.mem_seqCons_iff_of_lt {s x z : M} (hi : i < lh s) : ⟪i, x⟫ ∈ s ⁀' z ↔ ⟪i, x⟫ ∈ s := by +lemma Seq.mem_seqCons_iff_of_lt {s x z : V} (hi : i < lh s) : ⟪i, x⟫ ∈ s ⁀' z ↔ ⟪i, x⟫ ∈ s := by simp [seqCons, hi] rintro rfl; simp at hi -@[simp] lemma lh_not_mem {s} (Ss : Seq s) (x : M) : ⟪lh s, x⟫ ∉ s := fun h ↦ by have := Ss.lt_lh_of_mem h; simp at this +@[simp] lemma lh_not_mem {s} (Ss : Seq s) (x : V) : ⟪lh s, x⟫ ∉ s := fun h ↦ by have := Ss.lt_lh_of_mem h; simp at this section -lemma seqCons_graph (t x s : M) : +lemma seqCons_graph (t x s : V) : t = s ⁀' x ↔ ∃ l ≤ 2 * s, l = lh s ∧ ∃ p ≤ (2 * s + x + 1)^2, p = ⟪l, x⟫ ∧ t = insert p s := ⟨by rintro rfl exact ⟨lh s, by simp[lt_succ_iff_le], rfl, ⟪lh s, x⟫, @@ -214,30 +214,30 @@ lemma seqCons_graph (t x s : M) : def _root_.LO.FirstOrder.Arith.seqConsDef : 𝚺₀.Semisentence 3 := .mkSigma “t s x | ∃ l <⁺ 2 * s, !lhDef l s ∧ ∃ p <⁺ (2 * s + x + 1)², !pairDef p l x ∧ !insertDef t p s” (by simp) -lemma seqCons_defined : 𝚺₀-Function₂ (seqCons : M → M → M) via seqConsDef := by +lemma seqCons_defined : 𝚺₀-Function₂ (seqCons : V → V → V) via seqConsDef := by intro v; simp [seqConsDef, seqCons_graph] @[simp] lemma seqCons_defined_iff (v) : - Semiformula.Evalbm M v seqConsDef.val ↔ v 0 = v 1 ⁀' v 2 := seqCons_defined.df.iff v + Semiformula.Evalbm V v seqConsDef.val ↔ v 0 = v 1 ⁀' v 2 := seqCons_defined.df.iff v -instance seqCons_definable : 𝚺₀-Function₂ (seqCons : M → M → M) := Defined.to_definable _ seqCons_defined +instance seqCons_definable : 𝚺₀-Function₂ (seqCons : V → V → V) := seqCons_defined.to_definable -instance seqCons_definable' (Γ) : Γ-Function₂ (seqCons : M → M → M) := .of_zero seqCons_definable _ +instance seqCons_definable' (ℌ) : ℌ-Function₂ (seqCons : V → V → V) := seqCons_definable.of_zero -@[simp] lemma natCast_empty : ((∅ : ℕ) : M) = ∅ := by simp [emptyset_def] +@[simp] lemma natCast_empty : ((∅ : ℕ) : V) = ∅ := by simp [emptyset_def] -lemma seqCons_absolute (s a : ℕ) : ((s ⁀' a : ℕ) : M) = (s : M) ⁀' (a : M) := by - simpa using DefinedFunction.shigmaZero_absolute_func M seqCons_defined seqCons_defined ![s, a] +lemma seqCons_absolute (s a : ℕ) : ((s ⁀' a : ℕ) : V) = (s : V) ⁀' (a : V) := by + simpa using DefinedFunction.shigmaZero_absolute_func V seqCons_defined seqCons_defined ![s, a] end -lemma Seq.restr {s : M} (H : Seq s) {i : M} (hi : i ≤ lh s) : Seq (s ↾ under i) := +lemma Seq.restr {s : V} (H : Seq s) {i : V} (hi : i ≤ lh s) : Seq (s ↾ under i) := ⟨H.isMapping.restr (under i), i, domain_restr_of_subset_domain (by simp [H.domain_eq, hi])⟩ -lemma Seq.restr_lh {s : M} (H : Seq s) {i : M} (hi : i ≤ lh s) : lh (s ↾ under i) = i := +lemma Seq.restr_lh {s : V} (H : Seq s) {i : V} (hi : i ≤ lh s) : lh (s ↾ under i) = i := (H.restr hi).lh_eq_of (domain_restr_of_subset_domain <| by simp [H.domain_eq, hi]) -lemma domain_bitRemove_of_isMapping_of_mem {x y s : M} (hs : IsMapping s) (hxy : ⟪x, y⟫ ∈ s) : +lemma domain_bitRemove_of_isMapping_of_mem {x y s : V} (hs : IsMapping s) (hxy : ⟪x, y⟫ ∈ s) : domain (bitRemove ⟪x, y⟫ s) = bitRemove x (domain s) := by apply mem_ext; simp [mem_domain_iff]; intro x₁ constructor @@ -245,7 +245,7 @@ lemma domain_bitRemove_of_isMapping_of_mem {x y s : M} (hs : IsMapping s) (hxy : · intro ⟨hx, y₁, hx₁y₁⟩ exact ⟨y₁, by intro _; contradiction, hx₁y₁⟩ -lemma Seq.eq_of_eq_of_subset {s₁ s₂ : M} (H₁ : Seq s₁) (H₂ : Seq s₂) +lemma Seq.eq_of_eq_of_subset {s₁ s₂ : V} (H₁ : Seq s₁) (H₂ : Seq s₂) (hl : lh s₁ = lh s₂) (h : s₁ ⊆ s₂) : s₁ = s₂ := by apply mem_ext; intro u constructor @@ -258,18 +258,18 @@ lemma Seq.eq_of_eq_of_subset {s₁ s₂ : M} (H₁ : Seq s₁) (H₂ : Seq s₂) rcases this with rfl simpa using hy -lemma subset_pair {s t : M} (h : ∀ i x, ⟪i, x⟫ ∈ s → ⟪i, x⟫ ∈ t) : s ⊆ t := by +lemma subset_pair {s t : V} (h : ∀ i x, ⟪i, x⟫ ∈ s → ⟪i, x⟫ ∈ t) : s ⊆ t := by intro u hu simpa using h (π₁ u) (π₂ u) (by simpa using hu) -lemma Seq.lh_ext {s₁ s₂ : M} (H₁ : Seq s₁) (H₂ : Seq s₂) (h : lh s₁ = lh s₂) +lemma Seq.lh_ext {s₁ s₂ : V} (H₁ : Seq s₁) (H₂ : Seq s₂) (h : lh s₁ = lh s₂) (H : ∀ i x₁ x₂, ⟪i, x₁⟫ ∈ s₁ → ⟪i, x₂⟫ ∈ s₂ → x₁ = x₂) : s₁ = s₂ := H₁.eq_of_eq_of_subset H₂ h <| subset_pair <| by intro i x hx have hi : i < lh s₂ := by simpa [← h] using H₁.lt_lh_of_mem hx rcases H i _ _ hx (H₂.nth_mem hi) simp -@[simp] lemma Seq.seqCons_ext {a₁ a₂ s₁ s₂ : M} (H₁ : Seq s₁) (H₂ : Seq s₂) : +@[simp] lemma Seq.seqCons_ext {a₁ a₂ s₁ s₂ : V} (H₁ : Seq s₁) (H₂ : Seq s₂) : s₁ ⁀' a₁ = s₂ ⁀' a₂ ↔ a₁ = a₂ ∧ s₁ = s₂ := ⟨by intro h have hs₁s₂ : lh s₁ = lh s₂ := by simpa [H₁, H₂] using congr_arg lh h @@ -286,9 +286,9 @@ lemma Seq.lh_ext {s₁ s₂ : M} (H₁ : Seq s₁) (H₂ : Seq s₂) (h : lh s by rintro ⟨rfl, rfl⟩; rfl⟩ /-- TODO: move to Lemmata.lean-/ -lemma ne_zero_iff_one_le {a : M} : a ≠ 0 ↔ 1 ≤ a := Iff.trans pos_iff_ne_zero.symm (pos_iff_one_le (a := a)) +lemma ne_zero_iff_one_le {a : V} : a ≠ 0 ↔ 1 ≤ a := Iff.trans pos_iff_ne_zero.symm (pos_iff_one_le (a := a)) -lemma Seq.cases_iff {s : M} : Seq s ↔ s = ∅ ∨ ∃ x s', Seq s' ∧ s = s' ⁀' x := ⟨fun h ↦ by +lemma Seq.cases_iff {s : V} : Seq s ↔ s = ∅ ∨ ∃ x s', Seq s' ∧ s = s' ⁀' x := ⟨fun h ↦ by by_cases hs : lh s = 0 · left simpa [hs] using h.domain_eq @@ -316,9 +316,9 @@ lemma Seq.cases_iff {s : M} : Seq s ↔ s = ∅ ∨ ∃ x s', Seq s' ∧ s = s' alias ⟨Seq.cases, _⟩ := Seq.cases_iff @[elab_as_elim] -theorem seq_induction (Γ) {P : M → Prop} (hP : DefinablePred ℒₒᵣ (Γ, 1) P) +theorem seq_induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) (hnil : P ∅) (hcons : ∀ s x, Seq s → P s → P (s ⁀' x)) : - ∀ {s : M}, Seq s → P s := by + ∀ {s : V}, Seq s → P s := by intro s sseq induction s using order_induction_h_sigma1 · exact Γ @@ -344,49 +344,49 @@ def vecConsUnexpander : Lean.PrettyPrinter.Unexpander | `($_ ∅ $term) => `(!⟦$term⟧) | _ => throw () -@[simp] lemma singleton_seq (x : M) : Seq !⟦x⟧ := by apply Seq.seqCons; simp +@[simp] lemma singleton_seq (x : V) : Seq !⟦x⟧ := by apply Seq.seqCons; simp -@[simp] lemma doubleton_seq (x y : M) : Seq !⟦x, y⟧ := by apply Seq.seqCons; simp +@[simp] lemma doubleton_seq (x y : V) : Seq !⟦x, y⟧ := by apply Seq.seqCons; simp -@[simp] lemma mem_singleton_seq_iff (x y : M) : ⟪0, x⟫ ∈ !⟦y⟧ ↔ x = y := by simp [mem_seqCons_iff] +@[simp] lemma mem_singleton_seq_iff (x y : V) : ⟪0, x⟫ ∈ !⟦y⟧ ↔ x = y := by simp [mem_seqCons_iff] section def _root_.LO.FirstOrder.Arith.mkSeq₁Def : 𝚺₀.Semisentence 2 := .mkSigma “s x | !seqConsDef s 0 x” (by simp) -lemma mkSeq₁_defined : 𝚺₀-Function₁ (fun x : M ↦ !⟦x⟧) via mkSeq₁Def := by +lemma mkSeq₁_defined : 𝚺₀-Function₁ (fun x : V ↦ !⟦x⟧) via mkSeq₁Def := by intro v; simp [mkSeq₁Def]; rfl @[simp] lemma eval_mkSeq₁Def (v) : - Semiformula.Evalbm M v mkSeq₁Def.val ↔ v 0 = !⟦v 1⟧ := mkSeq₁_defined.df.iff v + Semiformula.Evalbm V v mkSeq₁Def.val ↔ v 0 = !⟦v 1⟧ := mkSeq₁_defined.df.iff v -instance mkSeq₁_definable : 𝚺₀-Function₁ (fun x : M ↦ !⟦x⟧) := Defined.to_definable _ mkSeq₁_defined +instance mkSeq₁_definable : 𝚺₀-Function₁ (fun x : V ↦ !⟦x⟧) := mkSeq₁_defined.to_definable -instance mkSeq₁_definable' (Γ) : Γ-Function₁ (fun x : M ↦ !⟦x⟧) := .of_zero mkSeq₁_definable _ +instance mkSeq₁_definable' (Γ) : Γ-Function₁ (fun x : V ↦ !⟦x⟧) := mkSeq₁_definable.of_zero -def _root_.LO.FirstOrder.Arith.mkSeq₂Def : 𝚺₁-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.mkSeq₂Def : 𝚺₁.Semisentence 3 := .mkSigma “s x y | ∃ sx, !mkSeq₁Def sx x ∧ !seqConsDef s sx y” (by simp) -lemma mkSeq₂_defined : 𝚺₁-Function₂ (fun x y : M ↦ !⟦x, y⟧) via mkSeq₂Def := by +lemma mkSeq₂_defined : 𝚺₁-Function₂ (fun x y : V ↦ !⟦x, y⟧) via mkSeq₂Def := by intro v; simp [mkSeq₂Def] @[simp] lemma eval_mkSeq₂Def (v) : - Semiformula.Evalbm M v mkSeq₂Def.val ↔ v 0 = !⟦v 1, v 2⟧ := mkSeq₂_defined.df.iff v + Semiformula.Evalbm V v mkSeq₂Def.val ↔ v 0 = !⟦v 1, v 2⟧ := mkSeq₂_defined.df.iff v -instance mkSeq₂_definable : 𝚺₁-Function₂ (fun x y : M ↦ !⟦x, y⟧) := Defined.to_definable _ mkSeq₂_defined +instance mkSeq₂_definable : 𝚺₁-Function₂ (fun x y : V ↦ !⟦x, y⟧) := mkSeq₂_defined.to_definable -instance mkSeq₂_definable' (Γ) : (Γ, m + 1)-Function₂ (fun x y : M ↦ !⟦x, y⟧) := .of_sigmaOne mkSeq₂_definable _ _ +instance mkSeq₂_definable' (Γ m) : Γ-[m + 1]-Function₂ (fun x y : V ↦ !⟦x, y⟧) := .of_sigmaOne mkSeq₂_definable _ _ end -theorem sigmaOne_skolem_seq {R : M → M → Prop} (hP : 𝚺₁-Relation R) {l} +theorem sigmaOne_skolem_seq {R : V → V → Prop} (hP : 𝚺₁-Relation R) {l} (H : ∀ x < l, ∃ y, R x y) : ∃ s, Seq s ∧ lh s = l ∧ ∀ i x, ⟪i, x⟫ ∈ s → R i x := by rcases sigmaOne_skolem hP (show ∀ x ∈ under l, ∃ y, R x y by simpa using H) with ⟨s, ms, sdom, h⟩ have : Seq s := ⟨ms, l, sdom⟩ exact ⟨s, this, by simpa [this.domain_eq] using sdom, h⟩ -theorem sigmaOne_skolem_seq! {R : M → M → Prop} (hP : 𝚺₁-Relation R) {l} +theorem sigmaOne_skolem_seq! {R : V → V → Prop} (hP : 𝚺₁-Relation R) {l} (H : ∀ x < l, ∃! y, R x y) : ∃! s, Seq s ∧ lh s = l ∧ ∀ i x, ⟪i, x⟫ ∈ s → R i x := by have : ∀ x < l, ∃ y, R x y := fun x hx ↦ (H x hx).exists rcases sigmaOne_skolem_seq hP this with ⟨s, Ss, rfl, hs⟩ @@ -396,23 +396,23 @@ theorem sigmaOne_skolem_seq! {R : M → M → Prop} (hP : 𝚺₁-Relation R) {l section seqToVec -def vecToSeq : {n : ℕ} → (Fin n → M) → M +def vecToSeq : {n : ℕ} → (Fin n → V) → V | 0, _ => ∅ | n + 1, v => vecToSeq (v ·.castSucc) ⁀' v (Fin.last n) -@[simp] lemma vecToSeq_nil : vecToSeq ![] = (∅ : M) := by simp [vecToSeq] +@[simp] lemma vecToSeq_nil : vecToSeq ![] = (∅ : V) := by simp [vecToSeq] -@[simp] lemma vecToSeq_vecCons {n} (v : Fin n → M) (a : M) : +@[simp] lemma vecToSeq_vecCons {n} (v : Fin n → V) (a : V) : vecToSeq (v <: a) = vecToSeq v ⁀' a := by simp [vecToSeq] -@[simp] lemma vecToSeq_seq {n} (v : Fin n → M) : Seq (vecToSeq v) := by +@[simp] lemma vecToSeq_seq {n} (v : Fin n → V) : Seq (vecToSeq v) := by induction' n with n ih <;> simp [vecToSeq] exact (ih _).seqCons _ -@[simp] lemma lh_vecToSeq {n} (v : Fin n → M) : lh (vecToSeq v) = n := by +@[simp] lemma lh_vecToSeq {n} (v : Fin n → V) : lh (vecToSeq v) = n := by induction' n with n ih <;> simp [vecToSeq, *] -lemma mem_vectoSeq {n : ℕ} (v : Fin n → M) (i : Fin n) : ⟪(i : M), v i⟫ ∈ vecToSeq v := by +lemma mem_vectoSeq {n : ℕ} (v : Fin n → V) (i : Fin n) : ⟪(i : V), v i⟫ ∈ vecToSeq v := by induction' n with n ih · exact i.elim0 · simp [vecToSeq] @@ -423,7 +423,7 @@ lemma mem_vectoSeq {n : ℕ} (v : Fin n → M) (i : Fin n) : ⟪(i : M), v i⟫ end seqToVec -lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function₂ f) {P : M → M → Prop} (hP : 𝚺₁-Relation P) +lemma sigma₁_order_ball_induction {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; @@ -437,32 +437,38 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function ∀ l < k, ∀ m < W, ∀ m' < W, ⟪l, m⟫ ∈ W → ⟪l + 1, m'⟫ ∈ W → ∀ x' ≤ x - l, ∀ y' ≤ m, f x' y' ≤ m' := by intro k hk induction k using induction_sigma1 - · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) - apply Definable.ex - apply Definable.and (Definable.comp₁_infer (DefinableFunction.var _)) - apply Definable.and - (Definable.comp₂_infer - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) - (DefinableFunction.comp₁_infer <| DefinableFunction.var _)) - apply Definable.and - (Definable.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.const _)) (DefinableFunction.var _)) - apply Definable.ball_lt (DefinableFunction.var _) - apply Definable.ball_lt (DefinableFunction.var _) - apply Definable.ball_lt (DefinableFunction.var _) - apply Definable.imp - (Definable.comp₂_infer - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _)) - (DefinableFunction.var _)) - apply Definable.imp - (Definable.comp₂_infer - (DefinableFunction.comp₂_infer - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) - (DefinableFunction.var _)) - (DefinableFunction.var _)) - apply Definable.ball_le - (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.var _))) - apply Definable.ball_le (DefinableFunction.var _) - apply Definable.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _)) (DefinableFunction.var _) + · apply HierarchySymbol.Boldface.imp (HierarchySymbol.Boldface.comp₂ (by definability) (by definability)) + apply HierarchySymbol.Boldface.ex + apply HierarchySymbol.Boldface.and (HierarchySymbol.Boldface.comp₁ (by definability)) + apply HierarchySymbol.Boldface.and + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.const _)) + (HierarchySymbol.BoldfaceFunction.comp₁ <| HierarchySymbol.BoldfaceFunction.var _)) + apply HierarchySymbol.Boldface.and + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ + (HierarchySymbol.BoldfaceFunction.const _) (HierarchySymbol.BoldfaceFunction.const _)) (HierarchySymbol.BoldfaceFunction.var _)) + apply HierarchySymbol.Boldface.ball_lt (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.ball_lt (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.ball_lt (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.imp + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _)) + (HierarchySymbol.BoldfaceFunction.var _)) + apply HierarchySymbol.Boldface.imp + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.const _)) + (HierarchySymbol.BoldfaceFunction.var _)) + (HierarchySymbol.BoldfaceFunction.var _)) + apply HierarchySymbol.Boldface.ball_le + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.comp₂ (HierarchySymbol.BoldfaceFunction.const _) (HierarchySymbol.BoldfaceFunction.var _))) + apply HierarchySymbol.Boldface.ball_le (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ + (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _)) (HierarchySymbol.BoldfaceFunction.var _) case zero => exact ⟨!⟦y⟧, by simp⟩ case succ k ih => rcases ih (le_trans le_self_add hk) with ⟨W, SW, hkW, hW₀, hWₛ⟩ @@ -482,15 +488,18 @@ lemma sigma₁_order_ball_induction {f : M → M → M} (hf : 𝚺₁-Function have : ∀ i ≤ x, ∀ m < W, ⟪x - i, m⟫ ∈ W → ∀ x' ≤ i, ∀ y' ≤ m, P x' y' := by intro i induction i using induction_sigma1 - · apply Definable.imp (Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.const _)) - apply Definable.ball_lt (DefinableFunction.const _) - apply Definable.imp - (Definable.comp₂_infer - (DefinableFunction.comp₂_infer (DefinableFunction.comp₂_infer (DefinableFunction.const _) (DefinableFunction.var _)) (DefinableFunction.var _)) - (DefinableFunction.const _)) - apply Definable.ball_le (DefinableFunction.var _) - apply Definable.ball_le (DefinableFunction.var _) - apply Definable.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _) + · apply HierarchySymbol.Boldface.imp + (HierarchySymbol.Boldface.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.const _)) + apply HierarchySymbol.Boldface.ball_lt (HierarchySymbol.BoldfaceFunction.const _) + apply HierarchySymbol.Boldface.imp + (HierarchySymbol.Boldface.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ + (HierarchySymbol.BoldfaceFunction.comp₂ + (HierarchySymbol.BoldfaceFunction.const _) (HierarchySymbol.BoldfaceFunction.var _)) (HierarchySymbol.BoldfaceFunction.var _)) + (HierarchySymbol.BoldfaceFunction.const _)) + apply HierarchySymbol.Boldface.ball_le (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.ball_le (HierarchySymbol.BoldfaceFunction.var _) + apply HierarchySymbol.Boldface.comp₂ (HierarchySymbol.BoldfaceFunction.var _) (HierarchySymbol.BoldfaceFunction.var _) case zero => intro _ _ _ _ _ h y' _ rcases nonpos_iff_eq_zero.mp h diff --git a/Arithmetization/ISigmaOne/HFS/Supplemental.lean b/Arithmetization/ISigmaOne/HFS/Supplemental.lean index 1755d7f..e603f1f 100644 --- a/Arithmetization/ISigmaOne/HFS/Supplemental.lean +++ b/Arithmetization/ISigmaOne/HFS/Supplemental.lean @@ -12,30 +12,30 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] section seqExp -lemma seqCons_le {x y s t : M} (hxy : x ≤ y) (hst : s ≤ t) : +lemma seqCons_le {x y s t : V} (hxy : x ≤ y) (hst : s ≤ t) : s ⁀' x ≤ t + exp ((2 * t + y + 1)^2) := by have : s ⁀' x ≤ t + exp ⟪2 * t, y⟫ := by simp [seqCons]; exact insert_le_of_le_of_le (pair_le_pair (le_trans (lh_bound s) (by simp [hst])) hxy) hst exact le_trans this (by simp) -lemma seqProduct_exists_unique (s a : M) : ∃! t : M, ∀ x, x ∈ t ↔ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := by +lemma seqProduct_exists_unique (s a : V) : ∃! t : V, ∀ x, x ∈ t ↔ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := by have : 𝚺₁-Predicate fun x ↦ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := by definability exact finite_comprehension₁! this ⟨log s + exp ((2 * log s + log a + 1)^2) + 1, by rintro x ⟨v, hv, u, hu, rfl⟩ exact lt_succ_iff_le.mpr <| seqCons_le (le_log_of_mem hu) (le_log_of_mem hv)⟩ -def seqProduct (a s : M) : M := Classical.choose! (seqProduct_exists_unique a s) +def seqProduct (a s : V) : V := Classical.choose! (seqProduct_exists_unique a s) infixl:60 " ×ˢ " => seqProduct -lemma mem_seqProduct_iff {x s a : M} : x ∈ s ×ˢ a ↔ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := +lemma mem_seqProduct_iff {x s a : V} : x ∈ s ×ˢ a ↔ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := Classical.choose!_spec (seqProduct_exists_unique s a) x -lemma mem_seqProduct_iff' {u v a s : M} (Hv : Seq v) (Hs : ∀ w ∈ s, Seq w) : +lemma mem_seqProduct_iff' {u v a s : V} (Hv : Seq v) (Hs : ∀ w ∈ s, Seq w) : v ⁀' u ∈ s ×ˢ a ↔ v ∈ s ∧ u ∈ a := ⟨by intro h rcases mem_seqProduct_iff.mp h with ⟨v', hv', u', hu', e⟩ @@ -45,16 +45,16 @@ lemma mem_seqProduct_iff' {u v a s : M} (Hv : Seq v) (Hs : ∀ w ∈ s, Seq w) : by rintro ⟨hv, hu⟩ exact mem_seqProduct_iff.mpr ⟨v, hv, u, hu, rfl⟩⟩ -lemma seqCons_mem_seqProduct {u v a s : M} (hv : v ∈ s) (hu : u ∈ a) : v ⁀' u ∈ s ×ˢ a := +lemma seqCons_mem_seqProduct {u v a s : V} (hv : v ∈ s) (hu : u ∈ a) : v ⁀' u ∈ s ×ˢ a := mem_seqProduct_iff.mpr (⟨v, hv, u, hu, rfl⟩) -lemma mem_seqProduct_bound {x s a : M} (h : x ∈ s ×ˢ a) : x ≤ s + exp ((2 * s + a + 1)^2) := by +lemma mem_seqProduct_bound {x s a : V} (h : x ∈ s ×ˢ a) : x ≤ s + exp ((2 * s + a + 1)^2) := by rcases mem_seqProduct_iff.mp h with ⟨v, hv, u, hu, rfl⟩ exact seqCons_le (le_of_lt <| lt_of_mem hu) (le_of_lt <| lt_of_mem hv) section -private lemma seqProduct_graph (t s a : M) : +private lemma seqProduct_graph (t s a : V) : t = s ×ˢ a ↔ ∃ e, e = exp ((2 * s + a + 1)^2) ∧ ∀ x ≤ t + s + e, x ∈ t ↔ ∃ v ∈ s, ∃ u ∈ a, x = v ⁀' u := ⟨by rintro rfl; exact ⟨exp ((2 * s + a + 1)^2), rfl, by intro x _; simp [mem_seqProduct_iff]⟩, by rintro ⟨_, rfl, h⟩ @@ -67,17 +67,17 @@ private lemma seqProduct_graph (t s a : M) : exact h x (le_trans (mem_seqProduct_bound hx) <| by simp [add_assoc]) |>.mpr (mem_seqProduct_iff.mp hx)⟩ -def _root_.LO.FirstOrder.Arith.seqProductDef : 𝚺₁-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.seqProductDef : 𝚺₁.Semisentence 3 := .mkSigma “t s a | ∃ e, !expDef e (2 * s + a + 1)² ∧ ∀ x <⁺ t + s + e, x ∈ t ↔ ∃ v ∈' s, ∃ u ∈' a, !seqConsDef x v u” (by simp [Hierarchy.iff_iff]) -lemma seqProduct_defined : 𝚺₁-Function₂ (seqProduct : M → M → M) via seqProductDef := by +lemma seqProduct_defined : 𝚺₁-Function₂ (seqProduct : V → V → V) via seqProductDef := by intro v; simp [seqProductDef, seqProduct_graph] @[simp] lemma seqProduct_defined_iff (v) : - Semiformula.Evalbm M v seqProductDef.val ↔ v 0 = v 1 ×ˢ v 2 := seqProduct_defined.df.iff v + Semiformula.Evalbm V v seqProductDef.val ↔ v 0 = v 1 ×ˢ v 2 := seqProduct_defined.df.iff v -instance seqProduct_definable : 𝚺₁-Function₂ (seqProduct : M → M → M) := Defined.to_definable _ seqProduct_defined +instance seqProduct_definable : 𝚺₁-Function₂ (seqProduct : V → V → V) := seqProduct_defined.to_definable end @@ -85,39 +85,39 @@ def seqExp.formulae : PR.Blueprint 1 where zero := .mkSigma “y x | y = 1” (by simp) succ := .mkSigma “y ih n x | !seqProductDef y ih x” (by simp) -def seqExp.construction : PR.Construction M seqExp.formulae where +def seqExp.construction : PR.Construction V seqExp.formulae where zero := fun _ ↦ {∅} succ := fun a _ s ↦ s ×ˢ a 0 - zero_defined := by intro v; simp [formulae, one_eq_singleton (M := M)] + zero_defined := by intro v; simp [formulae, one_eq_singleton (V := V)] succ_defined := by intro v; simp [formulae]; rfl -def seqExp (a k : M) : M := seqExp.construction.result ![a] k +def seqExp (a k : V) : V := seqExp.construction.result ![a] k infix:80 " ^ˢ " => seqExp -@[simp] lemma seqExp_zero (a : M) : a ^ˢ 0 = {∅} := by simp [seqExp, seqExp.construction] +@[simp] lemma seqExp_zero (a : V) : a ^ˢ 0 = {∅} := by simp [seqExp, seqExp.construction] -@[simp] lemma seqExp_succ (a k : M) : a ^ˢ (k + 1) = (a ^ˢ k) ×ˢ a := by simp [seqExp, seqExp.construction] +@[simp] lemma seqExp_succ (a k : V) : a ^ˢ (k + 1) = (a ^ˢ k) ×ˢ a := by simp [seqExp, seqExp.construction] -def _root_.LO.FirstOrder.Arith.seqExpDef : 𝚺₁-Semisentence 3 := seqExp.formulae.resultDef |>.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.seqExpDef : 𝚺₁.Semisentence 3 := seqExp.formulae.resultDef |>.rew (Rew.substs ![#0, #2, #1]) -lemma seqExp_defined : 𝚺₁-Function₂ (seqExp : M → M → M) via seqExpDef := +lemma seqExp_defined : 𝚺₁-Function₂ (seqExp : V → V → V) via seqExpDef := fun v ↦ by simp [seqExp.construction.result_defined_iff, seqExpDef]; rfl @[simp] lemma seqExp_defined_iff (v) : - Semiformula.Evalbm M v seqExpDef.val ↔ v 0 = v 1 ^ˢ v 2 := seqExp_defined.df.iff v + Semiformula.Evalbm V v seqExpDef.val ↔ v 0 = v 1 ^ˢ v 2 := seqExp_defined.df.iff v -instance seqExp_definable : 𝚺₁-Function₂ (seqExp : M → M → M) := Defined.to_definable _ seqExp_defined +instance seqExp_definable : 𝚺₁-Function₂ (seqExp : V → V → V) := seqExp_defined.to_definable -@[simp, definability] instance seqExp_definable' (Γ) : (Γ, m + 1)-Function₂ (seqExp : M → M → M) := +@[simp, definability] instance seqExp_definable' (Γ) : Γ-[m + 1]-Function₂ (seqExp : V → V → V) := .of_sigmaOne seqExp_definable _ _ -lemma mem_seqExp_iff {s a k : M} : s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ (∀ i z, ⟪i, z⟫ ∈ s → z ∈ a) := by +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 - · suffices 𝚷₁-Predicate fun {k} => ∀ {s : M}, s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ ∀ i < s, ∀ z < s, ⟪i, z⟫ ∈ s → z ∈ a + · suffices 𝚷₁-Predicate fun {k} => ∀ {s : V}, s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ ∀ i < s, ∀ z < s, ⟪i, z⟫ ∈ s → z ∈ a by exact this.of_iff (fun k ↦ forall_congr' <| fun s ↦ iff_congr (by rfl) <| and_congr (by rfl) <| and_congr (by rfl) - ⟨fun h i hi z _ hiz ↦ h i z hiz, fun h i z hiz ↦ h i (lt_of_mem_dom hiz) z (lt_of_mem_rng hiz) hiz⟩) + ⟨fun h i _ z _ hiz ↦ h i z hiz, fun h i z hiz ↦ h i (lt_of_mem_dom hiz) z (lt_of_mem_rng hiz) hiz⟩) definability case zero => simp only [seqExp_zero, mem_singleton_iff] @@ -148,11 +148,11 @@ lemma mem_seqExp_iff {s a k : M} : s ∈ a ^ˢ k ↔ Seq s ∧ lh s = k ∧ (∀ have hs : s ∈ a ^ˢ k := @ih s |>.mpr ⟨Hs', hsk, fun i z hiz ↦ hs i z (Seq.subset_seqCons s x hiz)⟩ exact seqCons_mem_seqProduct hs hx -lemma seq_of_mem_seqExp {s a k : M} (h : s ∈ a ^ˢ k) : Seq s := (mem_seqExp_iff.mp h).1 +lemma seq_of_mem_seqExp {s a k : V} (h : s ∈ a ^ˢ k) : Seq s := (mem_seqExp_iff.mp h).1 -lemma lh_of_mem_seqExp {s a k : M} (h : s ∈ a ^ˢ k) : lh s = k := (mem_seqExp_iff.mp h).2.1 +lemma lh_of_mem_seqExp {s a k : V} (h : s ∈ a ^ˢ k) : lh s = k := (mem_seqExp_iff.mp h).2.1 -lemma pair_mem_mem_seqExp {s a k : M} (h : s ∈ a ^ˢ k) {i z} (hiz : ⟪i, z⟫ ∈ s) : z ∈ a := (mem_seqExp_iff.mp h).2.2 i z hiz +lemma pair_mem_mem_seqExp {s a k : V} (h : s ∈ a ^ˢ k) {i z} (hiz : ⟪i, z⟫ ∈ s) : z ∈ a := (mem_seqExp_iff.mp h).2.2 i z hiz end seqExp diff --git a/Arithmetization/ISigmaOne/HFS/Vec.lean b/Arithmetization/ISigmaOne/HFS/Vec.lean index 9897c92..47ee0ab 100644 --- a/Arithmetization/ISigmaOne/HFS/Vec.lean +++ b/Arithmetization/ISigmaOne/HFS/Vec.lean @@ -12,7 +12,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] section cons @@ -76,9 +76,9 @@ lemma cons_defined : 𝚺₀-Function₂ (cons : V → V → V) via consDef := b @[simp] lemma eval_cons (v) : Semiformula.Evalbm V v consDef.val ↔ v 0 = v 1 ∷ v 2 := cons_defined.df.iff v -instance cons_definable : 𝚺₀-Function₂ (cons : V → V → V) := Defined.to_definable _ cons_defined +instance cons_definable : 𝚺₀-Function₂ (cons : V → V → V) := cons_defined.to_definable -instance cons_definable' (Γ) : Γ-Function₂ (cons : V → V → V) := .of_zero cons_definable _ +instance cons_definable' (ℌ) : ℌ-Function₂ (cons : V → V → V) := cons_definable.of_zero def _root_.LO.FirstOrder.Arith.mkVec₁Def : 𝚺₀.Semisentence 2 := .mkSigma “s x | !consDef s x 0” (by simp) @@ -89,11 +89,11 @@ lemma mkVec₁_defined : 𝚺₀-Function₁ (fun x : V ↦ ?[x]) via mkVec₁De @[simp] lemma eval_mkVec₁Def (v) : Semiformula.Evalbm V v mkVec₁Def.val ↔ v 0 = ?[v 1] := mkVec₁_defined.df.iff v -instance mkVec₁_definable : 𝚺₀-Function₁ (fun x : V ↦ ?[x]) := Defined.to_definable _ mkVec₁_defined +instance mkVec₁_definable : 𝚺₀-Function₁ (fun x : V ↦ ?[x]) := mkVec₁_defined.to_definable -instance mkVec₁_definable' (Γ) : Γ-Function₁ (fun x : V ↦ ?[x]) := .of_zero mkVec₁_definable _ +instance mkVec₁_definable' (ℌ) : ℌ-Function₁ (fun x : V ↦ ?[x]) := mkVec₁_definable.of_zero -def _root_.LO.FirstOrder.Arith.mkVec₂Def : 𝚺₁-Semisentence 3 := .mkSigma +def _root_.LO.FirstOrder.Arith.mkVec₂Def : 𝚺₁.Semisentence 3 := .mkSigma “s x y | ∃ sy, !mkVec₁Def sy y ∧ !consDef s x sy” (by simp) lemma mkVec₂_defined : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) via mkVec₂Def := by @@ -102,9 +102,9 @@ lemma mkVec₂_defined : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) via mkVec @[simp] lemma eval_mkVec₂Def (v) : Semiformula.Evalbm V v mkVec₂Def.val ↔ v 0 = ?[v 1, v 2] := mkVec₂_defined.df.iff v -instance mkVec₂_definable : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) := Defined.to_definable _ mkVec₂_defined +instance mkVec₂_definable : 𝚺₁-Function₂ (fun x y : V ↦ ?[x, y]) := mkVec₂_defined.to_definable -instance mkVec₂_definable' (Γ) : (Γ, 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]) := .of_sigmaOne mkVec₂_definable _ _ end @@ -162,12 +162,12 @@ def Graph : V → Prop := construction.Fixpoint ![] section -def graphDef : 𝚺₁-Semisentence 1 := blueprint.fixpointDef +def graphDef : 𝚺₁.Semisentence 1 := blueprint.fixpointDef lemma graph_defined : 𝚺₁-Predicate (Graph : V → Prop) via graphDef := construction.fixpoint_defined -instance graph_definable : 𝚺₁-Predicate (Graph : V → Prop) := Defined.to_definable _ graph_defined +instance graph_definable : 𝚺₁-Predicate (Graph : V → Prop) := graph_defined.to_definable end @@ -201,7 +201,7 @@ lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by suffices ∀ i' ≤ i, ∀ v' ≤ v, ∃ x, Graph ⟪v', i', x⟫ from this i (by simp) v (by simp) intro i' hi' induction i' using induction_sigma1 - · definability + · { } case zero => intro v' _ exact ⟨fstIdx v', graph_case.mpr <| Or.inl ⟨v', rfl⟩⟩ @@ -209,7 +209,7 @@ lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by intro v' hv' rcases ih (le_trans le_self_add hi') (sndIdx v') (le_trans (by simp) hv') with ⟨x, hx⟩ exact ⟨x, graph_case.mpr <| Or.inr ⟨v', i', x, rfl, hx⟩⟩ - +/--/ lemma graph_unique {v i x₁ x₂ : V} : Graph ⟪v, i, x₁⟫ → Graph ⟪v, i, x₂⟫ → x₁ = x₂ := by induction i using induction_pi1 generalizing v x₁ x₂ · definability @@ -260,7 +260,7 @@ lemma cons_cases (x : V) : x = 0 ∨ ∃ y v, x = y ∷ v := by · simp · right; exact ⟨π₁ z, π₂ z, by simp [cons]⟩ -lemma cons_induction (Γ) {P : V → Prop} (hP : (Γ, 1)-Predicate P) +lemma cons_induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) (nil : P 0) (cons : ∀ x v, P v → P (x ∷ v)) : ∀ v, P v := order_induction_hh ℒₒᵣ Γ 1 hP (by intro v ih @@ -280,7 +280,7 @@ lemma cons_induction_pi₁ {P : V → Prop} (hP : 𝚷₁-Predicate P) section -def _root_.LO.FirstOrder.Arith.nthDef : 𝚺₁-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.nthDef : 𝚺₁.Semisentence 3 := .mkSigma “y v i | ∃ pr, !pair₃Def pr v i y ∧ !graphDef pr” (by simp) lemma nth_defined : 𝚺₁-Function₂ (nth : V → V → V) via nthDef := by @@ -344,8 +344,8 @@ end nth namespace VecRec structure Blueprint (arity : ℕ) where - nil : 𝚺₁-Semisentence (arity + 1) - cons : 𝚺₁-Semisentence (arity + 4) + nil : 𝚺₁.Semisentence (arity + 1) + cons : 𝚺₁.Semisentence (arity + 4) namespace Blueprint @@ -368,9 +368,9 @@ def blueprint : Fixpoint.Blueprint arity where !pairDef pr xxs cons ∧ :⟪xs, ih⟫:∈ C)” (by simp)) -def graphDef : 𝚺₁-Semisentence (arity + 1) := β.blueprint.fixpointDef +def graphDef : 𝚺₁.Semisentence (arity + 1) := β.blueprint.fixpointDef -def resultDef : 𝚺₁-Semisentence (arity + 2) := +def resultDef : 𝚺₁.Semisentence (arity + 2) := .mkSigma “y xs | ∃ pr, !pairDef pr xs y ∧ !β.graphDef pr ⋯” (by simp) end Blueprint @@ -433,7 +433,7 @@ lemma graph_defined : Arith.Defined (fun v ↦ c.Graph (v ·.succ) (v 0)) β.gra instance graph_definable : Arith.Definable ℒₒᵣ 𝚺₁ (fun v ↦ c.Graph (v ·.succ) (v 0)) := Defined.to_definable _ c.graph_defined instance graph_definable' (param) : 𝚺₁-Predicate (c.Graph param) := by - simpa using Definable.retractiont (n := 1) c.graph_definable (#0 :> fun i ↦ &(param i)) + simpa using HierarchySymbol.Boldface.retractiont (n := 1) c.graph_definable (#0 :> fun i ↦ &(param i)) end @@ -560,7 +560,7 @@ def len (v : V) : V := construction.result ![] v section -def _root_.LO.FirstOrder.Arith.lenDef : 𝚺₁-Semisentence 2 := blueprint.resultDef +def _root_.LO.FirstOrder.Arith.lenDef : 𝚺₁.Semisentence 2 := blueprint.resultDef lemma len_defined : 𝚺₁-Function₁ (len : V → V) via lenDef := construction.result_defined @@ -681,7 +681,7 @@ lemma takeLast_cons (x v : V) : section -def _root_.LO.FirstOrder.Arith.takeLastDef : 𝚺₁-Semisentence 3 := blueprint.resultDef +def _root_.LO.FirstOrder.Arith.takeLastDef : 𝚺₁.Semisentence 3 := blueprint.resultDef lemma takeLast_defined : 𝚺₁-Function₂ (takeLast : V → V → V) via takeLastDef := construction.result_defined @@ -768,7 +768,7 @@ def concat (v z : V) : V := construction.result ![z] v section -def _root_.LO.FirstOrder.Arith.concatDef : 𝚺₁-Semisentence 3 := blueprint.resultDef +def _root_.LO.FirstOrder.Arith.concatDef : 𝚺₁.Semisentence 3 := blueprint.resultDef lemma concat_defined : 𝚺₁-Function₂ (concat : V → V → V) via concatDef := construction.result_defined @@ -840,7 +840,7 @@ lemma le_of_memVec {x v : V} (h : x ∈ᵥ v) : x ≤ v := by section -def _root_.LO.FirstOrder.Arith.memVecDef : 𝚫₁-Semisentence 2 := .mkDelta +def _root_.LO.FirstOrder.Arith.memVecDef : 𝚫₁.Semisentence 2 := .mkDelta (.mkSigma “x v | ∃ l, !lenDef l v ∧ ∃ i < l, !nthDef x v i” (by simp)) (.mkPi “x v | ∀ l, !lenDef l v → ∃ i < l, ∀ vi, !nthDef vi v i → x = vi” (by simp)) @@ -876,7 +876,7 @@ scoped infix:30 " ⊆ᵥ " => SubsetVec section -def _root_.LO.FirstOrder.Arith.subsetVecDef : 𝚫₁-Semisentence 2 := .mkDelta +def _root_.LO.FirstOrder.Arith.subsetVecDef : 𝚫₁.Semisentence 2 := .mkDelta (.mkSigma “v w | ∀ x <⁺ v, !memVecDef.pi x v → !memVecDef.sigma x w” (by simp)) (.mkPi “v w | ∀ x <⁺ v, !memVecDef.sigma x v → !memVecDef.pi x w” (by simp)) @@ -926,7 +926,7 @@ def repeatVec (x k : V) : V := repeatVec.construction.result ![x] k section -def _root_.LO.FirstOrder.Arith.repeatVecDef : 𝚺₁-Semisentence 3 := repeatVec.blueprint.resultDef |>.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.repeatVecDef : 𝚺₁.Semisentence 3 := repeatVec.blueprint.resultDef |>.rew (Rew.substs ![#0, #2, #1]) lemma repeatVec_defined : 𝚺₁-Function₂ (repeatVec : V → V → V) via repeatVecDef := fun v ↦ by simp [repeatVec.construction.result_defined_iff, repeatVecDef]; rfl @@ -997,7 +997,7 @@ def vecToSet (v : V) : V := construction.result ![] v section -def _root_.LO.FirstOrder.Arith.vecToSetDef : 𝚺₁-Semisentence 2 := blueprint.resultDef +def _root_.LO.FirstOrder.Arith.vecToSetDef : 𝚺₁.Semisentence 2 := blueprint.resultDef lemma vecToSet_defined : 𝚺₁-Function₁ (vecToSet : V → V) via vecToSetDef := construction.result_defined diff --git a/Arithmetization/ISigmaOne/Ind.lean b/Arithmetization/ISigmaOne/Ind.lean index 14496dc..7b9b052 100644 --- a/Arithmetization/ISigmaOne/Ind.lean +++ b/Arithmetization/ISigmaOne/Ind.lean @@ -54,17 +54,17 @@ lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P intro x hx induction x using order_induction_sigma1 · clear hp hq ind - apply LO.FirstOrder.Arith.Definable.imp + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.Definable.comp₂_infer + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] · simp_all only [Fin.isValue] - apply LO.FirstOrder.Arith.Definable.or - · apply LO.FirstOrder.Arith.Definable.comp₂_infer + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.or + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] - · apply LO.FirstOrder.Arith.Definable.comp₂_infer + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, DefinableFunction.const] case ind z ih => diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index 58edbaf..099d1a7 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -364,7 +364,7 @@ variable (L) def Language.UFormula : V → Prop := (construction L).Fixpoint ![] -def _root_.LO.FirstOrder.Arith.LDef.uformulaDef (pL : LDef) : 𝚫₁-Semisentence 1 := +def _root_.LO.FirstOrder.Arith.LDef.uformulaDef (pL : LDef) : 𝚫₁.Semisentence 1 := (blueprint pL).fixpointDefΔ₁ lemma uformula_defined : 𝚫₁-Predicate L.UFormula via pL.uformulaDef := @@ -385,7 +385,7 @@ abbrev Language.Formula (p : V) : Prop := L.Semiformula 0 p lemma Language.UFormula.toSemiformula {p} (h : L.UFormula p) : L.Semiformula (fstIdx p) p := ⟨h, by rfl⟩ -def _root_.LO.FirstOrder.Arith.LDef.isSemiformulaDef (pL : LDef) : 𝚫₁-Semisentence 2 := .mkDelta +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)) @@ -508,7 +508,7 @@ lemma Language.Semiformula.pos {n p : V} (h : L.Semiformula n p) : 0 < p := h.1. @[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) +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]) @@ -543,8 +543,8 @@ lemma Language.Semiformula.induction (Γ) {P : V → V → Prop} (hP : (Γ, 1)-R suffices ∀ p, 𝐔 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 Definable.ball_le (by definability) - apply Definable.imp (by definability) + · apply HierarchySymbol.Boldface.ball_le (by definability) + apply HierarchySymbol.Boldface.imp (by definability) simp; exact hP · rintro n k r v hr hv _ _ rfl; simpa using hrel n k r v hr hv · rintro n k r v hr hv _ _ rfl; simpa using hnrel n k r v hr hv @@ -590,16 +590,16 @@ 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 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 namespace Blueprint @@ -637,10 +637,10 @@ def blueprint (β : Blueprint pL) : Fixpoint.Blueprint 0 := ⟨.mkDelta (∀ param', !β.exChanges param' param n → :⟪param', p₁, y₁⟫:∈ C) ∧ !qqExDef p n p₁ ∧ !β.ex.graphDelta.pi.val y param n p₁ y₁)) ” (by simp))⟩ -def graph : 𝚺₁-Semisentence 3 := .mkSigma +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 +def result : 𝚺₁.Semisentence 3 := .mkSigma “y param p | (!pL.uformulaDef.pi p → !β.graph param p y) ∧ (¬!pL.uformulaDef.sigma p → y = 0)” (by simp) end Blueprint @@ -1209,9 +1209,9 @@ lemma uformula_result_induction {P : V → V → V → Prop} (hP : 𝚺₁-Relat (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _))) (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _)))) apply sigma₁_order_ball_induction hf ?_ ?_ p param - · apply Definable.imp - (Definable.comp₁_infer (DefinableFunction.var _)) - (Definable.comp₃_infer + · apply HierarchySymbol.Boldface.imp + (HierarchySymbol.Boldface.comp₁_infer (DefinableFunction.var _)) + (HierarchySymbol.Boldface.comp₃_infer (DefinableFunction.var _) (DefinableFunction.var _) (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _))) @@ -1255,17 +1255,17 @@ lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺 intro param p hp apply c.uformula_result_induction (P := fun param p y ↦ ∀ n ≤ p, n = fstIdx p → P param n p y) ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp - · apply Definable.ball_le (DefinableFunction.var _) + · apply HierarchySymbol.Boldface.ball_le (DefinableFunction.var _) simp_all only [zero_add, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, Fin.succ_one_eq_two, Fin.succ_zero_eq_one] - apply LO.FirstOrder.Arith.Definable.imp + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.Definable.comp₂_infer + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue] apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · apply LO.FirstOrder.Arith.Definable.comp₄_infer + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₄_infer · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] · simp_all only [zero_add, Fin.isValue] apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer @@ -1299,16 +1299,16 @@ 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) + 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 @@ -1336,7 +1336,7 @@ variable {arity} (β : Blueprint pL arity) namespace Blueprint -def decomp {n : ℕ} (s : 𝚺₁-Semisentence n) : 𝚺₁-Semisentence 1 := +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) diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index 726c28e..d5ce81d 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -63,7 +63,7 @@ variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.negDef (pL : LDef) : 𝚺₁-Semisentence 2 := (blueprint pL).result.rew (Rew.substs ![#0, ‘0’, #1]) +def _root_.LO.FirstOrder.Arith.LDef.negDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result.rew (Rew.substs ![#0, ‘0’, #1]) variable (L) @@ -169,7 +169,7 @@ section imp section -def _root_.LO.FirstOrder.Arith.LDef.impDef (pL : LDef) : 𝚺₁-Semisentence 4 := .mkSigma +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) variable (L) @@ -242,7 +242,7 @@ variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.shiftDef (pL : LDef) : 𝚺₁-Semisentence 2 := (blueprint pL).result.rew (Rew.substs ![#0, ‘0’, #1]) +def _root_.LO.FirstOrder.Arith.LDef.shiftDef (pL : LDef) : 𝚺₁.Semisentence 2 := (blueprint pL).result.rew (Rew.substs ![#0, ‘0’, #1]) variable (L) @@ -333,7 +333,7 @@ section variable (L) -def _root_.LO.FirstOrder.Arith.LDef.qVecDef (pL : LDef) : 𝚺₁-Semisentence 4 := .mkSigma +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) lemma qVec_defined : 𝚺₁-Function₃ L.qVec via pL.qVecDef := by @@ -395,7 +395,7 @@ variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.substsDef (pL : LDef) : 𝚺₁-Semisentence 4 := .mkSigma +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) variable (L) @@ -459,7 +459,7 @@ lemma uformula_subst_induction {P : V → V → V → V → Prop} (hP : 𝚺₁- 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 Definable.comp₄_infer + · apply HierarchySymbol.Boldface.comp₄_infer (DefinableFunction.comp₁_infer (DefinableFunction.var _)) (DefinableFunction.comp₁_infer (DefinableFunction.var _)) (DefinableFunction.var _) @@ -498,7 +498,7 @@ lemma semiformula_subst_induction {P : V → V → V → V → V → Prop} (hP : 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 Definable.comp₅_infer + · apply HierarchySymbol.Boldface.comp₅_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _)) (DefinableFunction.comp₁_infer (DefinableFunction.var _)) @@ -605,13 +605,13 @@ 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 revert m w n v apply Language.Semiformula.induction_pi1 ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ _ _ hp - · apply Definable.all - apply Definable.all - apply Definable.all - apply Definable.all - apply Definable.imp (by definability) - apply Definable.imp (by definability) - apply Definable.comp₂_infer (by simp; definability) + · 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₂_infer (by simp; definability) apply DefinableFunction.comp₃_infer (by definability) ?_ (by definability) apply DefinableFunction₅.comp (termSubstVec_definable _) <;> definability · intro l k R ts hR hts m w n v _ hv @@ -700,7 +700,7 @@ section substs₁ section -def _root_.LO.FirstOrder.Arith.LDef.substs₁Def (pL : LDef) : 𝚺₁-Semisentence 3 := .mkSigma +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) variable (L) @@ -727,7 +727,7 @@ section free section -def _root_.LO.FirstOrder.Arith.LDef.freeDef (pL : LDef) : 𝚺₁-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.LDef.freeDef (pL : LDef) : 𝚺₁.Semisentence 2 := .mkSigma “q p | ∃ fz, !qqFvarDef fz 0 ∧ ∃ sp, !pL.shiftDef sp p ∧ !pL.substs₁Def q fz sp” (by simp) variable (L) @@ -752,7 +752,7 @@ def Language.IsFVFree (n p : V) : Prop := L.Semiformula n p ∧ L.shift p = p section -def _root_.LO.FirstOrder.Arith.LDef.isFVFreeDef (pL : LDef) : 𝚺₁-Semisentence 2 := +def _root_.LO.FirstOrder.Arith.LDef.isFVFreeDef (pL : LDef) : 𝚺₁.Semisentence 2 := .mkSigma “n p | !pL.isSemiformulaDef.sigma n p ∧ !pL.shiftDef p p” (by simp) lemma isFVFree_defined : 𝚺₁-Relation L.IsFVFree via pL.isFVFreeDef := by @@ -816,16 +816,16 @@ notation:78 x:78 " ^≮[" n "] " y:79 => qqNLT n x y notation:78 x:78 " ^≮ " y:79 => qqNLT 0 x y -def _root_.LO.FirstOrder.Arith.qqEQDef : 𝚺₁-Semisentence 4 := +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.qqNEQDef : 𝚺₁-Semisentence 4 := +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.qqLTDef : 𝚺₁-Semisentence 4 := +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.qqNLTDef : 𝚺₁-Semisentence 4 := +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) lemma qqEQ_defined : 𝚺₁-Function₃ (qqEQ : V → V → V → V) via qqEQDef := by diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean index 523857d..9e0e4b1 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean @@ -41,7 +41,7 @@ scoped notation:65 "^⋀ " ps:66 => qqConj 0 ps section -def _root_.LO.FirstOrder.Arith.qqConjDef : 𝚺₁-Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.qqConjDef : 𝚺₁.Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) 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] @@ -113,7 +113,7 @@ scoped notation:65 "^⋁ " ps:66 => qqDisj 0 ps section -def _root_.LO.FirstOrder.Arith.qqDisjDef : 𝚺₁-Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) +def _root_.LO.FirstOrder.Arith.qqDisjDef : 𝚺₁.Semisentence 3 := blueprint.resultDef.rew (Rew.substs ![#0, #2, #1]) 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] @@ -176,7 +176,7 @@ def substItr (n w p k : V) : V := construction.result ![n, w, p] k 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 5 := blueprint.resultDef |>.rew (Rew.substs ![#0, #4, #1, #2, #3]) lemma substItr_defined : 𝚺₁-Function₄ (substItr : V → V → V → V → V) via substItrDef := fun v ↦ by simp [construction.result_defined_iff, substItrDef]; rfl @@ -291,7 +291,7 @@ def qqVerums (n k : V) : V := ^⋀[n] repeatVec (^⊤[n]) k section -def _root_.LO.FirstOrder.Arith.qqVerumsDef : 𝚺₁-Semisentence 3 := .mkSigma +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) lemma qqVerums_defined : 𝚺₁-Function₂ (qqVerums : V → V → V) via qqVerumsDef := diff --git a/Arithmetization/ISigmaOne/Metamath/Language.lean b/Arithmetization/ISigmaOne/Metamath/Language.lean index 2cb3f31..38a4030 100644 --- a/Arithmetization/ISigmaOne/Metamath/Language.lean +++ b/Arithmetization/ISigmaOne/Metamath/Language.lean @@ -41,10 +41,10 @@ instance Defined.func_definable : 𝚺₀-Relation L.Func := Defined.to_definabl instance Defined.rel_definable : 𝚺₀-Relation L.Rel := Defined.to_definable _ Defined.rel @[simp, definability] instance Defined.func_definable' (Γ) : Γ-Relation L.Func := - Definable.of_zero Defined.func_definable _ + HierarchySymbol.Boldface.of_zero Defined.func_definable _ @[simp, definability] instance Defined.rel_definable' (Γ) : Γ-Relation L.Rel := - Definable.of_zero Defined.rel_definable _ + HierarchySymbol.Boldface.of_zero Defined.rel_definable _ end Language diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean index 1e8a59a..3339f36 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean @@ -22,7 +22,7 @@ variable {L} section -def _root_.LO.FirstOrder.Arith.LDef.formulaSetDef (pL : LDef) : 𝚫₁-Semisentence 1 := .mkDelta +def _root_.LO.FirstOrder.Arith.LDef.formulaSetDef (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)) @@ -115,7 +115,7 @@ private lemma setShift_graph (t s : V) : rcases mem_setShift_iff.mp hy with ⟨x, hx, rfl⟩ exact h₂ x hx -def _root_.LO.FirstOrder.Arith.LDef.setShiftDef (pL : LDef) : 𝚺₁-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.LDef.setShiftDef (pL : LDef) : 𝚺₁.Semisentence 2 := .mkSigma “t s | (∀ y ∈' t, ∃ x ∈' s, !pL.shiftDef y x) ∧ (∀ x ∈' s, ∃ y, !pL.shiftDef y x ∧ y ∈ t)” (by simp) variable (L) @@ -587,7 +587,7 @@ def Language.Theory.Provable (p : V) : Prop := T.Derivable {p} section -def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationDef {pL : LDef} (pT : pL.TDef) : 𝚫₁-Semisentence 1 := (blueprint pT).fixpointDefΔ₁ +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Δ₁ @@ -596,7 +596,7 @@ instance derivation_definable : 𝚫₁-Predicate T.Derivation := Defined.to_def @[simp] instance derivatin_definable' (Γ) : (Γ, m + 1)-Predicate T.Derivation := .of_deltaOne (derivation_definable T) _ _ -def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationOfDef {pL : LDef} (pT : pL.TDef) : 𝚫₁-Semisentence 2 := .mkDelta +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)) @@ -609,7 +609,7 @@ instance derivationOf_definable : 𝚫₁-Relation T.DerivationOf := Defined.to_ @[simp] instance derivatinOf_definable' (Γ) : (Γ, m + 1)-Relation T.DerivationOf := .of_deltaOne (derivationOf_definable T) _ _ -def _root_.LO.FirstOrder.Arith.LDef.TDef.derivableDef {pL : LDef} (pT : pL.TDef) : 𝚺₁-Semisentence 1 := .mkSigma +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) lemma Language.Theory.derivable_defined : 𝚺₁-Predicate T.Derivable via pT.derivableDef := by @@ -620,7 +620,7 @@ instance Language.Theory.derivable_definable : 𝚺₁-Predicate T.Derivable := /-- instance for definability tactic-/ @[simp] instance Language.Theory.derivable_definable' : (𝚺, 0 + 1)-Predicate T.Derivable := derivable_definable T -def _root_.LO.FirstOrder.Arith.LDef.TDef.prv {pL : LDef} (pT : pL.TDef) : 𝚺₁-Semisentence 1 := .mkSigma +def _root_.LO.FirstOrder.Arith.LDef.TDef.prv {pL : LDef} (pT : pL.TDef) : 𝚺₁.Semisentence 1 := .mkSigma “p | ∃ s, !insertDef s p 0 ∧ !pT.derivableDef s” (by simp) protected lemma Language.Theory.provable_defined : 𝚺₁-Predicate T.Provable via pT.prv := by @@ -802,10 +802,10 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( (∀ i < len ps - k, ps.[i] ∈ s') → T.Derivable (insert (^⋁ takeLast ps k) (s' ∪ s)) := by intro k hk induction k using induction_sigma1 - · apply Definable.imp (by definability) - apply Definable.ball_le (by definability) - apply Definable.imp (by definability) - apply Definable.imp (by definability) + · apply HierarchySymbol.Boldface.imp (by definability) + apply HierarchySymbol.Boldface.ball_le (by definability) + apply HierarchySymbol.Boldface.imp (by definability) + apply HierarchySymbol.Boldface.imp (by definability) definability case zero => intro s' _ ss hs' diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean index 4a9e30c..6f3fd09 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Theory.lean @@ -114,10 +114,10 @@ lemma disjDistr (ps s : V) (d : T.Derivable (vecToSet ps ∪ s)) : T.Derivable ( (∀ i < len ps - k, ps.[i] ∈ s') → T.Derivable (insert (^⋁ takeLast ps k) (s' ∪ s)) := by intro k hk induction k using induction_sigma1 - · apply Definable.imp (by definability) - apply Definable.ball_le (by definability) - apply Definable.imp (by definability) - apply Definable.imp (by definability) + · apply HierarchySymbol.Boldface.imp (by definability) + apply HierarchySymbol.Boldface.ball_le (by definability) + apply HierarchySymbol.Boldface.imp (by definability) + apply HierarchySymbol.Boldface.imp (by definability) definability case zero => intro s' _ ss hs' diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean index bbcb15e..7d87189 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean @@ -153,7 +153,7 @@ def Language.Semiterm (n : V) : V → Prop := (construction L).Fixpoint ![n] abbrev Language.Term : V → Prop := L.Semiterm 0 -def _root_.LO.FirstOrder.Arith.LDef.isSemitermDef (pL : LDef) : 𝚫₁-Semisentence 2 := (blueprint pL).fixpointDefΔ₁.rew (Rew.substs ![#1, #0]) +def _root_.LO.FirstOrder.Arith.LDef.isSemitermDef (pL : LDef) : 𝚫₁.Semisentence 2 := (blueprint pL).fixpointDefΔ₁.rew (Rew.substs ![#1, #0]) lemma isSemiterm_defined : 𝚫₁-Relation L.Semiterm via pL.isSemitermDef := ⟨HSemiformula.ProperOn.rew (construction L).fixpoint_definedΔ₁.proper _, @@ -198,7 +198,7 @@ lemma Language.SemitermVec.prop {n m w : V} (h : L.SemitermVec n m w) {i} : i < section -def _root_.LO.FirstOrder.Arith.LDef.semitermVecDef (pL : LDef) : 𝚫₁-Semisentence 3 := .mkDelta +def _root_.LO.FirstOrder.Arith.LDef.semitermVecDef (pL : LDef) : 𝚫₁.Semisentence 3 := .mkDelta (.mkSigma “n m w | !lenDef n w ∧ ∀ i < n, ∃ u, !nthDef u w i ∧ !pL.isSemitermDef.sigma m u” (by simp)) @@ -256,7 +256,7 @@ alias ⟨Language.Semiterm.case, Language.Semiterm.mk⟩ := Language.Semiterm.ca 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) +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 := @@ -271,9 +271,9 @@ 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 + 3) + fvar : 𝚺₁.Semisentence (arity + 3) + func : 𝚺₁.Semisentence (arity + 6) namespace Blueprint @@ -297,13 +297,13 @@ def blueprint : Fixpoint.Blueprint (arity + 1) := ⟨.mkDelta !qqFuncDef t k f v ∧ !β.func.graphDelta.pi y n k f v w ⋯) )” (by simp))⟩ -def graph : 𝚺₁-Semisentence (arity + 3) := .mkSigma +def graph : 𝚺₁.Semisentence (arity + 3) := .mkSigma “n t y | ∃ pr <⁺ (t + y + 1)², !pairDef pr t y ∧ !β.blueprint.fixpointDef pr n ⋯” (by simp) -def result : 𝚺₁-Semisentence (arity + 3) := .mkSigma +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 resultVec : 𝚺₁-Semisentence (arity + 4) := .mkSigma +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) @@ -472,7 +472,7 @@ instance termSubst_definable : Arith.Definable ℒₒᵣ 𝚺₁ (fun v ↦ c.Gr Defined.to_definable _ (graph_defined c) @[simp, definability] instance termSubst_definable₂ (param n) : 𝚺₁-Relation (c.Graph param n) := by - simpa using Definable.retractiont (n := 2) (termSubst_definable c) (&n :> #0 :> #1 :> fun i ↦ &(param i)) + simpa using HierarchySymbol.Boldface.retractiont (n := 2) (termSubst_definable c) (&n :> #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 @@ -524,7 +524,7 @@ lemma graph_exists {t : V} : L.Semiterm n t → ∃ y, c.Graph param n t y := by 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 apply Definable.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) ih + (by apply HierarchySymbol.Boldface.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) 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⟩ @@ -588,7 +588,7 @@ lemma graph_existsUnique_vec {k n w : V} (hw : L.SemitermVec k n w) : 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)⟩ rcases sigmaOne_skolem_vec - (by apply Definable.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) this + (by apply HierarchySymbol.Boldface.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) this with ⟨w', hw'k, hw'⟩ refine ExistsUnique.intro w' ⟨hw'k.symm, hw'⟩ ?_ intro w'' ⟨hkw'', hw''⟩ diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean index 00afeff..cd12e89 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean @@ -57,9 +57,9 @@ variable {n m w : V} 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 5 := (blueprint pL).result.rew <| Rew.substs ![#0, #1, #4, #2, #3] -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 6 := (blueprint pL).resultVec.rew <| Rew.substs ![#0, #1, #2, #5, #3, #4] variable (L) @@ -216,10 +216,10 @@ variable {n : V} section -def _root_.LO.FirstOrder.Arith.LDef.termShiftDef (pL : LDef) : 𝚺₁-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.LDef.termShiftDef (pL : LDef) : 𝚺₁.Semisentence 3 := (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 4 := (blueprint pL).resultVec variable (L) @@ -333,10 +333,10 @@ variable {n : V} section -def _root_.LO.FirstOrder.Arith.LDef.termBShiftDef (pL : LDef) : 𝚺₁-Semisentence 3 := +def _root_.LO.FirstOrder.Arith.LDef.termBShiftDef (pL : LDef) : 𝚺₁.Semisentence 3 := (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 4 := (blueprint pL).resultVec variable (L) @@ -576,7 +576,7 @@ def numeralAux (x : V) : V := construction.result ![] x section -def numeralAuxDef : 𝚺₁-Semisentence 2 := blueprint.resultDef +def numeralAuxDef : 𝚺₁.Semisentence 2 := blueprint.resultDef lemma numeralAux_defined : 𝚺₁-Function₁ (numeralAux : V → V) via numeralAuxDef := fun v ↦ by simp [construction.result_defined_iff, numeralAuxDef]; rfl @@ -618,7 +618,7 @@ lemma numeral_succ_pos (pos : 0 < n) : numeral (n + 1 : V) = numeral n ^+ 𝟏 : section -def _root_.LO.FirstOrder.Arith.numeralDef : 𝚺₁-Semisentence 2 := .mkSigma +def _root_.LO.FirstOrder.Arith.numeralDef : 𝚺₁.Semisentence 2 := .mkSigma “t x | (x = 0 → t = !!(Semiterm.Operator.numeral ℒₒᵣ Formalized.zero)) ∧ (x ≠ 0 → ∃ x', !subDef x' x 1 ∧ !numeralAuxDef t x')” diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean index 3952c77..3d1a6e4 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean @@ -46,7 +46,7 @@ class R₀Theory (T : LOR.TTheory (V := V)) extends EQTheory T where /- section -def _root_.LO.FirstOrder.Arith.eqTheory : 𝚺₁-Semisentence 0 := .mkSigma +def _root_.LO.FirstOrder.Arith.eqTheory : 𝚺₁.Semisentence 0 := .mkSigma “(∃ b0, !qqBvarDef b0 0 ∧ !qqAllDef )” (by simp) end diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean index 30d1277..952d4f0 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean @@ -33,14 +33,14 @@ variable {T V} lemma Language.SyntacticTheory.codeIn_iff : x ∈ T.codeIn V ↔ V ⊧/![x] T.tDef.ch.val := iff_of_eq rfl 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 + have := consequence_iff_add_eq.mp (sound! <| SyntacticTheory.Δ₁HierarchySymbol.Boldface.mem_iff.mp h) V inferInstance simpa [models_iff, Semiformula.syntacticformula_goedelNumber_def, numeral_eq_natCast] using this instance tDef_defined : (T.codeIn V).Defined T.tDef where defined := ⟨by intro v rw [show v = ![v 0] from Matrix.constant_eq_singleton'] - have := consequence_iff_add_eq.mp (sound! <| FirstOrder.SyntacticTheory.Δ₁Definable.isΔ₁ (T := T)) V inferInstance + have := consequence_iff_add_eq.mp (sound! <| FirstOrder.SyntacticTheory.Δ₁HierarchySymbol.Boldface.isΔ₁ (T := T)) V inferInstance simp [models_iff] at this ⊢ simp [SyntacticTheory.tDef, this], by intro v; simp [SyntacticTheory.codeIn, ←Matrix.constant_eq_singleton']⟩ diff --git a/Arithmetization/ISigmaZero/Exponential/Exp.lean b/Arithmetization/ISigmaZero/Exponential/Exp.lean index eb35f39..d0bb659 100644 --- a/Arithmetization/ISigmaZero/Exponential/Exp.lean +++ b/Arithmetization/ISigmaZero/Exponential/Exp.lean @@ -4,17 +4,17 @@ noncomputable section namespace LO.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] open FirstOrder FirstOrder.Arith section ISigma₀ -variable [M ⊧ₘ* 𝐈𝚺₀] +variable [V ⊧ₘ* 𝐈𝚺₀] -def ext (u z : M) : M := z / u % u +def ext (u z : V) : V := z / u % u -lemma ext_graph (a b c : M) : a = ext b c ↔ ∃ x ≤ c, x = c / b ∧ a = x % b := by +lemma ext_graph (a b c : V) : a = ext b c ↔ ∃ x ≤ c, x = c / b ∧ a = x % b := by simp [ext]; constructor · rintro rfl; exact ⟨c / b, by simp, rfl, by rfl⟩ · rintro ⟨_, _, rfl, rfl⟩; simp @@ -22,48 +22,48 @@ lemma ext_graph (a b c : M) : a = ext b c ↔ ∃ x ≤ c, x = c / b ∧ a = x % def _root_.LO.FirstOrder.Arith.extDef : 𝚺₀.Semisentence 3 := .mkSigma “a b c | ∃ x <⁺ c, !divDef x c b ∧ !remDef a x b” (by simp) -lemma ext_defined : 𝚺₀-Function₂ (λ a b : M ↦ ext a b) via extDef := by +lemma ext_defined : 𝚺₀-Function₂ (λ a b : V ↦ ext a b) via extDef := by intro v; simp [Matrix.vecHead, Matrix.vecTail, extDef, ext_graph, Semiformula.eval_substs, div_defined.df.iff, rem_defined.df.iff, le_iff_lt_succ] -instance ext_definable : 𝚺₀-Function₂ (ext : M → M → M) := Defined.to_definable _ ext_defined +instance ext_definable : 𝚺₀-Function₂ (ext : V → V → V) := ext_defined.to_definable -@[simp] lemma ext_le_add (u z : M) : ext u z ≤ z := +@[simp] lemma ext_le_add (u z : V) : ext u z ≤ z := le_trans (mod_le (z / u) u) (by simp [add_comm]) -instance : Bounded₂ (ext : M → M → M) := ⟨#1, by intro v; simp⟩ +instance : Bounded₂ (ext : V → V → V) := ⟨#1, by intro v; simp⟩ -@[simp] lemma ext_lt {u} (z : M) (pos : 0 < u) : ext u z < u := by simp [ext, pos] +@[simp] lemma ext_lt {u} (z : V) (pos : 0 < u) : ext u z < u := by simp [ext, pos] -lemma ext_add_of_dvd_sq_right {u z₁ z₂ : M} (pos : 0 < u) (h : u^2 ∣ z₂) : ext u (z₁ + z₂) = ext u z₁ := by +lemma ext_add_of_dvd_sq_right {u z₁ z₂ : V} (pos : 0 < u) (h : u^2 ∣ z₂) : ext u (z₁ + z₂) = ext u z₁ := by simp [ext] have : ∃ z', z₂ = z' * u * u := by rcases h with ⟨u', rfl⟩; exact ⟨u', by simp [mul_comm _ u', mul_assoc]; simp [sq]⟩ rcases this with ⟨z₂, rfl⟩ simp [div_add_mul_self, pos] -lemma ext_add_of_dvd_sq_left {u z₁ z₂ : M} (pos : 0 < u) (h : u^2 ∣ z₁) : ext u (z₁ + z₂) = ext u z₂ := by +lemma ext_add_of_dvd_sq_left {u z₁ z₂ : V} (pos : 0 < u) (h : u^2 ∣ z₁) : ext u (z₁ + z₂) = ext u z₂ := by rw [add_comm]; exact ext_add_of_dvd_sq_right pos h -lemma ext_rem {i j z : M} (ppi : PPow2 i) (ppj : PPow2 j) (hij : i < j) : ext i (z % j) = ext i z := by +lemma ext_rem {i j z : V} (ppi : PPow2 i) (ppj : PPow2 j) (hij : i < j) : ext i (z % j) = ext i z := by have := div_add_mod z j have : i^2 ∣ j := ppi.pow2.sq.dvd_of_le ppj.pow2 (PPow2.sq_le_of_lt ppi ppj hij) calc ext i (z % j) = ext i (j * (z / j) + (z % j)) := by symm; exact ext_add_of_dvd_sq_left ppi.pos (Dvd.dvd.mul_right this (z / j)) _ = ext i z := by simp [div_add_mod] -def Exponential.Seq₀ (X Y : M) : Prop := ext 4 X = 1 ∧ ext 4 Y = 2 +def Exponential.Seq₀ (X Y : V) : Prop := ext 4 X = 1 ∧ ext 4 Y = 2 -def Exponential.Seqₛ.Even (X Y u : M) : Prop := ext (u^2) X = 2 * ext u X ∧ ext (u^2) Y = (ext u Y)^2 +def Exponential.Seqₛ.Even (X Y u : V) : Prop := ext (u^2) X = 2 * ext u X ∧ ext (u^2) Y = (ext u Y)^2 -def Exponential.Seqₛ.Odd (X Y u : M) : Prop := ext (u^2) X = 2 * ext u X + 1 ∧ ext (u^2) Y = 2 * (ext u Y)^2 +def Exponential.Seqₛ.Odd (X Y u : V) : Prop := ext (u^2) X = 2 * ext u X + 1 ∧ ext (u^2) Y = 2 * (ext u Y)^2 -def Exponential.Seqₛ (y X Y : M) : Prop := ∀ u ≤ y, u ≠ 2 → PPow2 u → Seqₛ.Even X Y u ∨ Seqₛ.Odd X Y u +def Exponential.Seqₛ (y X Y : V) : Prop := ∀ u ≤ y, u ≠ 2 → PPow2 u → Seqₛ.Even X Y u ∨ Seqₛ.Odd X Y u -def Exponential.Seqₘ (x y X Y : M) : Prop := ∃ u ≤ y^2, u ≠ 2 ∧ PPow2 u ∧ ext u X = x ∧ ext u Y = y +def Exponential.Seqₘ (x y X Y : V) : Prop := ∃ u ≤ y^2, u ≠ 2 ∧ PPow2 u ∧ ext u X = x ∧ ext u Y = y -def Exponential (x y : M) : Prop := (x = 0 ∧ y = 1) ∨ ∃ X ≤ y^4, ∃ Y ≤ y^4, Exponential.Seq₀ X Y ∧ Exponential.Seqₛ y X Y ∧ Exponential.Seqₘ x y X Y +def Exponential (x y : V) : Prop := (x = 0 ∧ y = 1) ∨ ∃ X ≤ y^4, ∃ Y ≤ y^4, Exponential.Seq₀ X Y ∧ Exponential.Seqₛ y X Y ∧ Exponential.Seqₘ x y X Y -lemma Exponential.Seqₛ.iff (y X Y : M) : +lemma Exponential.Seqₛ.iff (y X Y : V) : Exponential.Seqₛ y X Y ↔ ∀ u ≤ y, u ≠ 2 → PPow2 u → ((∃ ext_u_X ≤ X, ext_u_X = ext u X ∧ 2 * ext_u_X = ext (u^2) X) ∧ (∃ ext_u_Y ≤ Y, ext_u_Y = ext u Y ∧ ext_u_Y^2 = ext (u^2) Y)) ∨ @@ -85,11 +85,11 @@ def Exponential.Seqₛ.def : 𝚺₀.Semisentence 3 := .mkSigma ( (∃ ext_u_X <⁺ X, !extDef ext_u_X u X ∧ !extDef (2 * ext_u_X + 1) u² X) ∧ (∃ ext_u_Y <⁺ Y, !extDef ext_u_Y u Y ∧ !extDef (2 * ext_u_Y²) u² Y) ) ” (by simp) -lemma Exponential.Seqₛ.defined : 𝚺₀-Relation₃ (Exponential.Seqₛ : M → M → M → Prop) via Exponential.Seqₛ.def := by +lemma Exponential.Seqₛ.defined : 𝚺₀-Relation₃ (Exponential.Seqₛ : V → V → V → Prop) via Exponential.Seqₛ.def := by intro v; simp [Exponential.Seqₛ.iff, Exponential.Seqₛ.def, ppow2_defined.df.iff, ext_defined.df.iff, ←le_iff_lt_succ, sq, numeral_eq_natCast] -lemma Exponential.graph_iff (x y : M) : +lemma Exponential.graph_iff (x y : V) : Exponential x y ↔ (x = 0 ∧ y = 1) ∨ ∃ X ≤ y^4, ∃ Y ≤ y^4, (1 = ext 4 X ∧ 2 = ext 4 Y) ∧ @@ -109,41 +109,41 @@ def _root_.LO.FirstOrder.Arith.exponentialDef : 𝚺₀.Semisentence 2 := .mkSig !Exponential.Seqₛ.def y X Y ∧ ∃ u <⁺ y², u ≠ 2 ∧ !ppow2Def u ∧ !extDef x u X ∧ !extDef y u Y” (by simp) -lemma Exponential.defined : 𝚺₀-Relation (Exponential : M → M → Prop) via exponentialDef := by +lemma Exponential.defined : 𝚺₀-Relation (Exponential : V → V → Prop) via exponentialDef := by intro v; simp [Exponential.graph_iff, exponentialDef, ppow2_defined.df.iff, ext_defined.df.iff, Exponential.Seqₛ.defined.df.iff, ←le_iff_lt_succ, pow_four, sq, numeral_eq_natCast] @[simp] lemma exponential_defined_iff (v) : - Semiformula.Evalbm M v exponentialDef.val ↔ Exponential (v 0) (v 1) := Exponential.defined.df.iff v + Semiformula.Evalbm V v exponentialDef.val ↔ Exponential (v 0) (v 1) := Exponential.defined.df.iff v -instance exponential_definable : 𝚺₀-Relation (Exponential : M → M → Prop) := Defined.to_definable _ Exponential.defined +instance exponential_definable : 𝚺₀-Relation (Exponential : V → V → Prop) := Exponential.defined.to_definable -@[simp] instance exponential_definable' (Γ) : Γ-Relation (Exponential : M → M → Prop) := Definable.of_zero exponential_definable _ +@[simp] instance exponential_definable' (Γ) : Γ-Relation (Exponential : V → V → Prop) := exponential_definable.of_zero namespace Exponential -def seqX₀ : M := 4 +def seqX₀ : V := 4 -def seqY₀ : M := 2 * 4 +def seqY₀ : V := 2 * 4 -lemma one_lt_four : (1 : M) < 4 := by +lemma one_lt_four : (1 : V) < 4 := by rw [←three_add_one_eq_four] exact lt_add_of_pos_left 1 three_pos -lemma two_lt_three : (2 : M) < 3 := by rw [←two_add_one_eq_three]; exact lt_add_one 2 +lemma two_lt_three : (2 : V) < 3 := by rw [←two_add_one_eq_three]; exact lt_add_one 2 -lemma three_lt_four : (3 : M) < 4 := by rw [←three_add_one_eq_four]; exact lt_add_one 3 +lemma three_lt_four : (3 : V) < 4 := by rw [←three_add_one_eq_four]; exact lt_add_one 3 -lemma two_lt_four : (2 : M) < 4 := lt_trans two_lt_three three_lt_four +lemma two_lt_four : (2 : V) < 4 := lt_trans two_lt_three three_lt_four -lemma seq₀_zero_two : Seq₀ (seqX₀ : M) (seqY₀ : M) := by simp [seqX₀, seqY₀, Seq₀, ext, one_lt_four, two_lt_four] +lemma seq₀_zero_two : Seq₀ (seqX₀ : V) (seqY₀ : V) := by simp [seqX₀, seqY₀, Seq₀, ext, one_lt_four, two_lt_four] -lemma Seq₀.rem {X Y i : M} (h : Seq₀ X Y) (ppi : PPow2 i) (hi : 4 < i) : +lemma Seq₀.rem {X Y i : V} (h : Seq₀ X Y) (ppi : PPow2 i) (hi : 4 < i) : Seq₀ (X % i) (Y % i) := by rw [Seq₀, ext_rem, ext_rem] <;> try simp [ppi, hi] exact h -lemma Seqₛ.rem {y y' X Y i : M} (h : Seqₛ y X Y) (ppi : PPow2 i) (hi : y'^2 < i) (hy : y' ≤ y) : +lemma Seqₛ.rem {y y' X Y i : V} (h : Seqₛ y X Y) (ppi : PPow2 i) (hi : y'^2 < i) (hy : y' ≤ y) : Seqₛ y' (X % i) (Y % i) := by intro j hj ne2 ppj have : j^2 < i := lt_of_le_of_lt (sq_le_sq.mpr hj) hi @@ -152,21 +152,21 @@ lemma Seqₛ.rem {y y' X Y i : M} (h : Seqₛ y X Y) (ppi : PPow2 i) (hi : y'^2 · left; simpa [Seqₛ.Even, ext_rem, ppj, ppj.sq, ppi, *] using H · right; simpa [Seqₛ.Odd, ext_rem, ppj, ppj.sq, ppi, *] using H -lemma seqₛ_one_zero_two : Seqₛ (1 : M) (seqX₀ : M) (seqY₀ : M) := by +lemma seqₛ_one_zero_two : Seqₛ (1 : V) (seqX₀ : V) (seqY₀ : V) := by intro u leu; rcases le_one_iff_eq_zero_or_one.mp leu with (rfl | rfl) <;> simp -def append (i X z : M) : M := X % i + z * i +def append (i X z : V) : V := X % i + z * i -lemma append_lt (i X : M) {z} (hz : z < i) : append i X z < i^2 := calc +lemma append_lt (i X : V) {z} (hz : z < i) : append i X z < i^2 := calc append i X z = (X % i) + z * i := rfl _ < (1 + z) * i := by simp [add_mul]; exact mod_lt _ (pos_of_gt hz) _ ≤ i^2 := by simp [sq, add_comm]; exact mul_le_mul_of_nonneg_right (lt_iff_succ_le.mp hz) (by simp) -lemma ext_append_last (i X : M) {z} (hz : z < i) : +lemma ext_append_last (i X : V) {z} (hz : z < i) : ext i (append i X z) = z := by simp [ext, append, div_add_mul_self, show 0 < i from pos_of_gt hz, hz] -lemma ext_append_of_lt {i j : M} (hi : PPow2 i) (hj : PPow2 j) (hij : i < j) (X z : M) : +lemma ext_append_of_lt {i j : V} (hi : PPow2 i) (hj : PPow2 j) (hij : i < j) (X z : V) : ext i (append j X z) = ext i X := by have : i^2 ∣ j := Pow2.dvd_of_le hi.pow2.sq hj.pow2 (PPow2.sq_le_of_lt hi hj hij) calc @@ -175,12 +175,12 @@ lemma ext_append_of_lt {i j : M} (hi : PPow2 i) (hj : PPow2 j) (hij : i < j) (X _ = ext i (j * (X / j) + (X % j)) := by rw [add_comm]; refine Eq.symm <| ext_add_of_dvd_sq_right hi.pos (Dvd.dvd.mul_right this _) _ = ext i X := by simp [div_add_mod] -lemma Seq₀.append {X Y i x y : M} (H : Seq₀ X Y) (ppi : PPow2 i) (hi : 4 < i) : +lemma Seq₀.append {X Y i x y : V} (H : Seq₀ X Y) (ppi : PPow2 i) (hi : 4 < i) : Seq₀ (append i X x) (append i Y y) := by rw [Seq₀, ext_append_of_lt, ext_append_of_lt] <;> try simp [ppi, hi] exact H -lemma Seqₛ.append {z x y X Y i : M} (h : Seqₛ z X Y) (ppi : PPow2 i) (hz : z < i) : +lemma Seqₛ.append {z x y X Y i : V} (h : Seqₛ z X Y) (ppi : PPow2 i) (hz : z < i) : Seqₛ z (append (i^2) X x) (append (i^2) Y y) := by intro j hj ne2 ppj have : j < i^2 := lt_of_lt_of_le (lt_of_le_of_lt hj hz) (by simp) @@ -191,17 +191,17 @@ lemma Seqₛ.append {z x y X Y i : M} (h : Seqₛ z X Y) (ppi : PPow2 i) (hz : z · right; rw [ext_append_of_lt, ext_append_of_lt, ext_append_of_lt, ext_append_of_lt] <;> try simp [ppi.sq, ppj.sq, lt_of_le_of_lt hj hz, *] exact H -@[simp] lemma exponential_zero_one : Exponential (0 : M) 1 := Or.inl (by simp) +@[simp] lemma exponential_zero_one : Exponential (0 : V) 1 := Or.inl (by simp) -@[simp] lemma exponential_one_two : Exponential (1 : M) 2 := +@[simp] lemma exponential_one_two : Exponential (1 : V) 2 := Or.inr ⟨ 4, by simp [pow_four_eq_sq_sq, two_pow_two_eq_four], - 2 * 4, by simp [pow_four_eq_sq_sq, two_pow_two_eq_four, sq (4 : M)]; exact le_of_lt two_lt_four, + 2 * 4, by simp [pow_four_eq_sq_sq, two_pow_two_eq_four, sq (4 : V)]; exact le_of_lt two_lt_four, by simp [Seq₀, ext, one_lt_four, two_lt_four], by simp [Seqₛ]; intro i hi ne2 ppi; exact False.elim <| not_le.mpr (ppi.two_lt ne2) hi, ⟨4, by simp [two_pow_two_eq_four], by simp, by simp [ext, one_lt_four, two_lt_four]⟩⟩ -lemma pow2_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma pow2_ext_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : Pow2 (ext i Y) := by induction i using order_induction_sigma0 · definability @@ -217,12 +217,12 @@ lemma pow2_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) ( · have : ext i Y = 2*(ext (√i) Y)^2 := by simpa [ppi.sq_sqrt_eq ne2] using hodd.2 simp [this, ppsq] -lemma range_pow2 {x y : M} (h : Exponential x y) : Pow2 y := by +lemma range_pow2 {x y : V} (h : Exponential x y) : Pow2 y := by rcases h with (⟨rfl, rfl⟩ | ⟨X, bX, Y, bY, H₀, Hₛ, ⟨u, hu, ne2, ppu, rfl, rfl⟩⟩) · simp · exact pow2_ext_of_seq₀_of_seqₛ H₀ Hₛ ne2 hu ppu -lemma le_sq_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma le_sq_ext_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : i ≤ (ext i Y)^2 := by induction i using order_induction_sigma0 · definability @@ -243,7 +243,7 @@ lemma le_sq_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) example {a b c : ℕ} : a * (b * c) = b * (a * c) := by exact Nat.mul_left_comm a b c -lemma two_mul_ext_le_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma two_mul_ext_le_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 2 * ext i Y ≤ i := by induction i using order_induction_sigma0 · definability @@ -266,14 +266,14 @@ lemma two_mul_ext_le_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ _ ≤ (√i)^2 := sq_le_sq.mpr IH _ = i := ppi.sq_sqrt_eq ne2 -lemma exponential_exists_sq_of_exponential_even {x y : M} : Exponential (2 * x) y → ∃ y', y = y'^2 ∧ Exponential x y' := by +lemma exponential_exists_sq_of_exponential_even {x y : V} : Exponential (2 * x) y → ∃ y', y = y'^2 ∧ Exponential x y' := by rintro (⟨hx, rfl⟩ | ⟨X, _, Y, _, hseq₀, hseqₛ, i, hi, ne2, ppi, hXx, hYy⟩) · exact ⟨1, by simp [show x = 0 from by simpa using hx]⟩ by_cases ne4 : i = 4 · rcases ne4 with rfl have ex : 1 = 2 * x := by simpa [hseq₀.1] using hXx - have : (2 : M) ∣ 1 := by rw [ex]; simp - have : ¬(2 : M) ∣ 1 := not_dvd_of_lt (by simp) one_lt_two + have : (2 : V) ∣ 1 := by rw [ex]; simp + have : ¬(2 : V) ∣ 1 := not_dvd_of_lt (by simp) one_lt_two contradiction have : Seqₛ.Even X Y (√i) ∨ Seqₛ.Odd X Y (√i) := hseqₛ (√i) (sqrt_le_of_le_sq hi) (ppi.sqrt_ne_two ne2 ne4) (ppi.sqrt ne2) @@ -299,7 +299,7 @@ lemma exponential_exists_sq_of_exponential_even {x y : M} : Exponential (2 * x) ←mod_eq_zero_iff_dvd, one_lt_two] contradiction -lemma bit_zero {x y : M} : Exponential x y → Exponential (2 * x) (y ^ 2) := by +lemma bit_zero {x y : V} : Exponential x y → Exponential (2 * x) (y ^ 2) := by rintro (⟨hx, rfl⟩ | ⟨X, _, Y, _, hseq₀, hseqₛ, i, hi, ne2, ppi, hXx, hYy⟩) · rcases hx with rfl; simp have hxsqi : 2 * x < i ^ 2 := lt_of_lt_of_le (by simp [←hXx, ppi.pos]) (two_mul_le_sq ppi.two_le) @@ -332,16 +332,16 @@ lemma bit_zero {x y : M} : Exponential x y → Exponential (2 * x) (y ^ 2) := by by simp at *; rw [ext_append_last, ext_append_last] <;> try simp [hxsqi, hysqi]⟩ exact Or.inr <| ⟨X', bX', Y', bY', hseq₀', hseqₛ', hseqₘ'⟩ -lemma exponential_even {x y : M} : Exponential (2 * x) y ↔ ∃ y', y = y'^2 ∧ Exponential x y' := +lemma exponential_even {x y : V} : Exponential (2 * x) y ↔ ∃ y', y = y'^2 ∧ Exponential x y' := ⟨exponential_exists_sq_of_exponential_even, by rintro ⟨y, rfl, h⟩; exact bit_zero h⟩ -lemma exponential_even_sq {x y : M} : Exponential (2 * x) (y ^ 2) ↔ Exponential x y := +lemma exponential_even_sq {x y : V} : Exponential (2 * x) (y ^ 2) ↔ Exponential x y := ⟨by intro h rcases exponential_exists_sq_of_exponential_even h with ⟨y', e, h⟩ simpa [show y = y' from by simpa using e] using h, bit_zero⟩ -lemma exponential_exists_sq_of_exponential_odd {x y : M} : Exponential (2 * x + 1) y → ∃ y', y = 2 * y'^2 ∧ Exponential x y' := by +lemma exponential_exists_sq_of_exponential_odd {x y : V} : Exponential (2 * x + 1) y → ∃ y', y = 2 * y'^2 ∧ Exponential x y' := by rintro (⟨hx, rfl⟩ | ⟨X, _, Y, _, hseq₀, hseqₛ, i, hi, ne2, ppi, hXx, hYy⟩) · simp at hx by_cases ne4 : i = 4 @@ -375,7 +375,7 @@ lemma exponential_exists_sq_of_exponential_odd {x y : M} : Exponential (2 * x + Or.inr ⟨X', bX', Y', bY', hseq₀.rem ppi (ppi.four_lt ne2 ne4), hseqₛ', hseqₘ'⟩ exact ⟨ext (√i) Y, hYy, this⟩ -lemma bit_one {x y : M} : Exponential x y → Exponential (2 * x + 1) (2 * y ^ 2) := by +lemma bit_one {x y : V} : Exponential x y → Exponential (2 * x + 1) (2 * y ^ 2) := by rintro (⟨hx, rfl⟩ | ⟨X, _, Y, _, hseq₀, hseqₛ, i, hi, ne2, ppi, hXx, hYy⟩) · rcases hx with rfl; simp have hxsqi : 2 * x + 1 < i ^ 2 := calc @@ -416,16 +416,16 @@ lemma bit_one {x y : M} : Exponential x y → Exponential (2 * x + 1) (2 * y ^ 2 by simp [X', Y']; rw [ext_append_last, ext_append_last] <;> try simp [hxsqi, hysqi]⟩ exact Or.inr <| ⟨X', bX', Y', bY', hseq₀', hseqₛ', hseqₘ'⟩ -lemma exponential_odd {x y : M} : Exponential (2 * x + 1) y ↔ ∃ y', y = 2 * y' ^ 2 ∧ Exponential x y' := +lemma exponential_odd {x y : V} : Exponential (2 * x + 1) y ↔ ∃ y', y = 2 * y' ^ 2 ∧ Exponential x y' := ⟨exponential_exists_sq_of_exponential_odd, by rintro ⟨y, rfl, h⟩; exact bit_one h⟩ -lemma exponential_odd_two_mul_sq {x y : M} : Exponential (2 * x + 1) (2 * y ^ 2) ↔ Exponential x y := +lemma exponential_odd_two_mul_sq {x y : V} : Exponential (2 * x + 1) (2 * y ^ 2) ↔ Exponential x y := ⟨by intro h rcases exponential_exists_sq_of_exponential_odd h with ⟨y', e, h⟩ simpa [show y = y' from by simpa using e] using h, bit_one⟩ -lemma two_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma two_le_ext_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 2 ≤ ext i Y := by induction i using order_induction_sigma0 · definability @@ -446,7 +446,7 @@ lemma two_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) _ ≤ 2 * (ext (√i) Y)^2 := by simp _ = ext i Y := by simpa [ppi.sq_sqrt_eq ne2] using Eq.symm hodd.2 -lemma ext_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma ext_le_ext_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : ext i X < ext i Y := by induction i using order_induction_sigma0 · definability @@ -470,30 +470,30 @@ lemma ext_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) simp [one_add_one_eq_two]; exact (two_le_ext_of_seq₀_of_seqₛ h₀ hₛ (ppi.sqrt_ne_two ne2 ne4) (le_trans (by simp) hi) (ppi.sqrt ne2)))) _ = ext i Y := by simpa [ppi.sq_sqrt_eq ne2] using Eq.symm hodd.2 -lemma range_pos {x y : M} (h : Exponential x y) : 0 < y := by +lemma range_pos {x y : V} (h : Exponential x y) : 0 < y := by rcases h with (⟨rfl, rfl⟩ | ⟨X, bX, Y, bY, H₀, Hₛ, ⟨u, hu, ne2, ppu, rfl, rfl⟩⟩) · simp · have : 2 ≤ ext u Y := two_le_ext_of_seq₀_of_seqₛ H₀ Hₛ ne2 hu ppu exact lt_of_lt_of_le (by simp) this -lemma lt {x y : M} (h : Exponential x y) : x < y := by +lemma lt {x y : V} (h : Exponential x y) : x < y := by rcases h with (⟨rfl, rfl⟩ | ⟨X, bX, Y, bY, H₀, Hₛ, ⟨u, hu, ne2, ppu, rfl, rfl⟩⟩) · simp · exact ext_le_ext_of_seq₀_of_seqₛ H₀ Hₛ ne2 hu ppu -lemma not_exponential_of_le {x y : M} (h : x ≤ y) : ¬Exponential y x := by +lemma not_exponential_of_le {x y : V} (h : x ≤ y) : ¬Exponential y x := by intro hxy; exact not_le.mpr (lt hxy) h -@[simp] lemma one_not_even (a : M) : 1 ≠ 2 * a := by +@[simp] lemma one_not_even (a : V) : 1 ≠ 2 * a := by intro h - have : (2 : M) ∣ 1 := by rw [h]; simp - have : ¬(2 : M) ∣ 1 := not_dvd_of_lt (by simp) one_lt_two + have : (2 : V) ∣ 1 := by rw [h]; simp + have : ¬(2 : V) ∣ 1 := not_dvd_of_lt (by simp) one_lt_two contradiction -@[simp] lemma exponential_two_four : Exponential (2 : M) 4 := by - simpa [two_pow_two_eq_four] using (show Exponential (1 : M) 2 from by simp).bit_zero +@[simp] lemma exponential_two_four : Exponential (2 : V) 4 := by + simpa [two_pow_two_eq_four] using (show Exponential (1 : V) 2 from by simp).bit_zero -lemma exponential_succ {x y : M} : Exponential (x + 1) y ↔ ∃ z, y = 2 * z ∧ Exponential x z := by +lemma exponential_succ {x y : V} : Exponential (x + 1) y ↔ ∃ z, y = 2 * z ∧ Exponential x z := by suffices x < y → (Exponential (x + 1) y ↔ ∃ z ≤ y, y = 2 * z ∧ Exponential x z) by by_cases hxy : x < y · simp [this hxy] @@ -539,13 +539,13 @@ lemma exponential_succ {x y : M} : Exponential (x + 1) y ↔ ∃ z, y = 2 * z have : Exponential (x + 1) (2 * y) := this.mpr ⟨y, by simp, rfl, H'⟩ simpa [sq, mul_add, add_assoc, mul_assoc, one_add_one_eq_two, mul_left_comm y 2] using this.bit_zero -lemma exponential_succ_mul_two {x y : M} : Exponential (x + 1) (2 * y) ↔ Exponential x y := +lemma exponential_succ_mul_two {x y : V} : Exponential (x + 1) (2 * y) ↔ Exponential x y := ⟨by intro h; rcases exponential_succ.mp h with ⟨y', e, h⟩; simpa [show y = y' from by simpa using e] using h, by intro h; exact exponential_succ.mpr ⟨y, rfl, h⟩⟩ alias ⟨of_succ_two_mul, succ⟩ := exponential_succ_mul_two -lemma one_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) +lemma one_le_ext_of_seq₀_of_seqₛ {y X Y : V} (h₀ : Exponential.Seq₀ X Y) (hₛ : Exponential.Seqₛ y X Y) {i} (ne2 : i ≠ 2) (hi : i ≤ y^2) (ppi : PPow2 i) : 1 ≤ ext i X := by induction i using order_induction_sigma0 · definability @@ -561,21 +561,21 @@ lemma one_le_ext_of_seq₀_of_seqₛ {y X Y : M} (h₀ : Exponential.Seq₀ X Y) · have : ext i X = 2 * ext (√i) X + 1 := by simpa [ppi.sq_sqrt_eq ne2] using hodd.1 simp [this] -lemma zero_uniq {y : M} (h : Exponential 0 y) : y = 1 := by +lemma zero_uniq {y : V} (h : Exponential 0 y) : y = 1 := by rcases h with (⟨_, rfl⟩ | ⟨X, _, Y, _, H₀, Hₛ, ⟨u, hu, ne2, ppu, hX, _⟩⟩) · rfl · have : 1 ≤ ext u X := one_le_ext_of_seq₀_of_seqₛ H₀ Hₛ ne2 hu ppu simp [hX] at this -@[simp] lemma zero_uniq_iff {y : M} : Exponential 0 y ↔ y = 1 := +@[simp] lemma zero_uniq_iff {y : V} : Exponential 0 y ↔ y = 1 := ⟨zero_uniq, by rintro rfl; simp⟩ -lemma succ_lt_s {y : M} (h : Exponential (x + 1) y) : 2 ≤ y := by +lemma succ_lt_s {y : V} (h : Exponential (x + 1) y) : 2 ≤ y := by rcases h with (⟨h, rfl⟩ | ⟨X, _, Y, _, H₀, Hₛ, ⟨u, hu, ne2, ppu, _, hY⟩⟩) · simp at h · simpa [hY] using two_le_ext_of_seq₀_of_seqₛ H₀ Hₛ ne2 hu ppu -protected lemma uniq {x y₁ y₂ : M} : Exponential x y₁ → Exponential x y₂ → y₁ = y₂ := by +protected lemma uniq {x y₁ y₂ : V} : Exponential x y₁ → Exponential x y₂ → y₁ = y₂ := by intro h₁ h₂ wlog h : y₁ ≤ y₂ · exact Eq.symm <| this h₂ h₁ (show y₂ ≤ y₁ from le_of_not_ge h) @@ -594,7 +594,7 @@ protected lemma uniq {x y₁ y₂ : M} : Exponential x y₁ → Exponential x y x h₂'.lt y₁ (by simpa using h) h₁' h₂' simp [this] -protected lemma inj {x₁ x₂ y : M} : Exponential x₁ y → Exponential x₂ y → x₁ = x₂ := by +protected lemma inj {x₁ x₂ y : V} : Exponential x₁ y → Exponential x₂ y → x₁ = x₂ := by intro h₁ h₂ revert x₁ x₂ h₁ h₂ suffices ∀ x₁ < y, ∀ x₂ < y, Exponential x₁ y → Exponential x₂ y → x₁ = x₂ by @@ -618,7 +618,7 @@ protected lemma inj {x₁ x₂ y : M} : Exponential x₁ y → Exponential x₂ x₁ hy₁.lt x₂ hy₂.lt hy₁ hy₂ simp [this] -lemma exponential_elim {x y : M} : Exponential x y ↔ (x = 0 ∧ y = 1) ∨ ∃ x', ∃ y', x = x' + 1 ∧ y = 2 * y' ∧ Exponential x' y' := +lemma exponential_elim {x y : V} : Exponential x y ↔ (x = 0 ∧ y = 1) ∨ ∃ x', ∃ y', x = x' + 1 ∧ y = 2 * y' ∧ Exponential x' y' := ⟨by intro h rcases zero_or_succ x with (rfl | ⟨x', rfl⟩) · simp [h.zero_uniq] @@ -628,7 +628,7 @@ lemma exponential_elim {x y : M} : Exponential x y ↔ (x = 0 ∧ y = 1) ∨ ∃ · simp · exact h.succ⟩ -lemma monotone {x₁ x₂ y₁ y₂ : M} : Exponential x₁ y₁ → Exponential x₂ y₂ → x₁ < x₂ → y₁ < y₂ := by +lemma monotone {x₁ x₂ y₁ y₂ : V} : Exponential x₁ y₁ → Exponential x₂ y₂ → x₁ < x₂ → y₁ < y₂ := by suffices ∀ x₁ < y₁, ∀ y₂ ≤ y₁, ∀ x₂ < y₂, Exponential x₁ y₁ → Exponential x₂ y₂ → x₂ ≤ x₁ by intro h₁ h₂; contrapose; simp intro hy @@ -650,22 +650,22 @@ lemma monotone {x₁ x₂ y₁ y₂ : M} : Exponential x₁ y₁ → Exponential x₁ h₁'.lt y₂ (le_of_mul_le_mul_left hy (by simp)) x₂ h₂'.lt h₁' h₂' simpa using this -lemma monotone_le {x₁ x₂ y₁ y₂ : M} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ ≤ x₂ → y₁ ≤ y₂ := by +lemma monotone_le {x₁ x₂ y₁ y₂ : V} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ ≤ x₂ → y₁ ≤ y₂ := by rintro (rfl | h) · exact (h₁.uniq h₂).le · exact le_of_lt (monotone h₁ h₂ h) -lemma monotone_iff {x₁ x₂ y₁ y₂ : M} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ < x₂ ↔ y₁ < y₂ := by +lemma monotone_iff {x₁ x₂ y₁ y₂ : V} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ < x₂ ↔ y₁ < y₂ := by constructor · exact monotone h₁ h₂ · contrapose; simp; exact monotone_le h₂ h₁ -lemma monotone_le_iff {x₁ x₂ y₁ y₂ : M} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ ≤ x₂ ↔ y₁ ≤ y₂ := by +lemma monotone_le_iff {x₁ x₂ y₁ y₂ : V} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : x₁ ≤ x₂ ↔ y₁ ≤ y₂ := by constructor · exact monotone_le h₁ h₂ · contrapose; simp; exact monotone h₂ h₁ -lemma add_mul {x₁ x₂ y₁ y₂ : M} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : Exponential (x₁ + x₂) (y₁ * y₂) := by +lemma add_mul {x₁ x₂ y₁ y₂ : V} (h₁ : Exponential x₁ y₁) (h₂ : Exponential x₂ y₂) : Exponential (x₁ + x₂) (y₁ * y₂) := by wlog hy : y₁ ≥ y₂ · simpa [add_comm, mul_comm] using this h₂ h₁ (le_of_not_ge hy) revert y₂ @@ -688,77 +688,76 @@ end ISigma₀ section ISigma₁ -variable [M ⊧ₘ* 𝐈𝚺₁] +variable [V ⊧ₘ* 𝐈𝚺₁] namespace Exponential -lemma range_exists (x : M) : ∃ y, Exponential x y := by +lemma range_exists (x : V) : ∃ y, Exponential x y := by induction x using induction_sigma1 - · apply Definable.ex - definability + · definability case zero => exact ⟨1, by simp⟩ case succ x IH => rcases IH with ⟨y, IH⟩ exact ⟨2 * y, IH.succ⟩ -lemma range_exists_unique (x : M) : ∃! y, Exponential x y := by +lemma range_exists_unique (x : V) : ∃! y, Exponential x y := by rcases range_exists x with ⟨y, h⟩ exact ExistsUnique.intro y h (by intro y' h'; exact h'.uniq h) end Exponential -instance : Exp M := ⟨fun a ↦ Classical.choose! (Exponential.range_exists_unique a)⟩ +instance : Exp V := ⟨fun a ↦ Classical.choose! (Exponential.range_exists_unique a)⟩ section exponential -lemma exponential_exp (a : M) : Exponential a (exp a) := Classical.choose!_spec (Exponential.range_exists_unique a) +lemma exponential_exp (a : V) : Exponential a (exp a) := Classical.choose!_spec (Exponential.range_exists_unique a) -lemma exponential_graph {a b : M} : a = exp b ↔ Exponential b a := Classical.choose!_eq_iff _ +lemma exponential_graph {a b : V} : a = exp b ↔ Exponential b a := Classical.choose!_eq_iff _ def _root_.LO.FirstOrder.Arith.expDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | !exponentialDef.val y x” (by simp) -lemma exp_defined_deltaZero : 𝚺₀-Function₁ (Exp.exp : M → M) via expDef := by +lemma exp_defined_deltaZero : 𝚺₀-Function₁ (Exp.exp : V → V) via expDef := by intro v; simp [expDef, exponential_graph] @[simp] lemma exp_defined_iff (v) : - Semiformula.Evalbm M v expDef.val ↔ v 0 = Exp.exp (v 1) := exp_defined_deltaZero.df.iff v + Semiformula.Evalbm V v expDef.val ↔ v 0 = Exp.exp (v 1) := exp_defined_deltaZero.df.iff v -instance exp_definable_deltaZero : 𝚺₀-Function₁ (Exp.exp : M → M) := Defined.to_definable _ exp_defined_deltaZero +instance exp_definable_deltaZero : 𝚺₀-Function₁ (Exp.exp : V → V) := exp_defined_deltaZero.to_definable -lemma exp_of_exponential {a b : M} (h : Exponential a b) : exp a = b := +lemma exp_of_exponential {a b : V} (h : Exponential a b) : exp a = b := Eq.symm <| exponential_graph.mpr h -lemma exp_inj : Function.Injective (Exp.exp : M → M) := λ a _ H ↦ +lemma exp_inj : Function.Injective (Exp.exp : V → V) := λ a _ H ↦ (exponential_exp a).inj (exponential_graph.mp H) -@[simp] lemma exp_zero : exp (0 : M) = 1 := exp_of_exponential (by simp) +@[simp] lemma exp_zero : exp (0 : V) = 1 := exp_of_exponential (by simp) -@[simp] lemma exp_one : exp (1 : M) = 2 := exp_of_exponential (by simp) +@[simp] lemma exp_one : exp (1 : V) = 2 := exp_of_exponential (by simp) -lemma exp_succ (a : M) : exp (a + 1) = 2 * exp a := +lemma exp_succ (a : V) : exp (a + 1) = 2 * exp a := exp_of_exponential <| Exponential.exponential_succ_mul_two.mpr <| exponential_exp a -instance models_exponential_of_models_iSigmaOne : M ⊧ₘ* 𝐄𝐗𝐏 := +instance models_exponential_of_models_iSigmaOne : V ⊧ₘ* 𝐄𝐗𝐏 := ⟨by intro f hf; rcases hf <;> simp [models_iff, exp_succ]⟩ -lemma exp_even (a : M) : exp (2 * a) = (exp a)^2 := +lemma exp_even (a : V) : exp (2 * a) = (exp a)^2 := exp_of_exponential <| Exponential.exponential_even_sq.mpr <| exponential_exp a -@[simp] lemma lt_exp (a : M) : a < exp a := (exponential_exp a).lt +@[simp] lemma lt_exp (a : V) : a < exp a := (exponential_exp a).lt -@[simp] lemma exp_pos (a : M) : 0 < exp a := (exponential_exp a).range_pos +@[simp] lemma exp_pos (a : V) : 0 < exp a := (exponential_exp a).range_pos -@[simp] lemma one_le_exp (a : M) : 1 ≤ exp a := pos_iff_one_le.mp (by simp) +@[simp] lemma one_le_exp (a : V) : 1 ≤ exp a := pos_iff_one_le.mp (by simp) -@[simp] lemma exp_pow2 (a : M) : Pow2 (exp a) := (exponential_exp a).range_pow2 +@[simp] lemma exp_pow2 (a : V) : Pow2 (exp a) := (exponential_exp a).range_pow2 -@[simp] lemma exp_monotone {a b : M} : exp a < exp b ↔ a < b := +@[simp] lemma exp_monotone {a b : V} : exp a < exp b ↔ a < b := Iff.symm <| Exponential.monotone_iff (exponential_exp a) (exponential_exp b) -@[simp] lemma exp_monotone_le {a b : M} : exp a ≤ exp b ↔ a ≤ b := +@[simp] lemma exp_monotone_le {a b : V} : exp a ≤ exp b ↔ a ≤ b := Iff.symm <| Exponential.monotone_le_iff (exponential_exp a) (exponential_exp b) -lemma nat_cast_exp (n : ℕ) : (exp n : ℕ) = exp (n : M) := by +lemma nat_cast_exp (n : ℕ) : (exp n : ℕ) = exp (n : V) := by induction' n with n ih · simp · simp [exp_succ, ih] diff --git a/Arithmetization/ISigmaZero/Exponential/Log.lean b/Arithmetization/ISigmaZero/Exponential/Log.lean index ab4fc4d..734a74c 100644 --- a/Arithmetization/ISigmaZero/Exponential/Log.lean +++ b/Arithmetization/ISigmaZero/Exponential/Log.lean @@ -2,17 +2,17 @@ import Arithmetization.ISigmaZero.Exponential.Exp noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] - namespace LO.Arith +variable {V : Type*} [ORingStruc V] + open FirstOrder FirstOrder.Arith section ISigma₀ -variable [M ⊧ₘ* 𝐈𝚺₀] +variable [V ⊧ₘ* 𝐈𝚺₀] -lemma log_exists_unique_pos {y : M} (hy : 0 < y) : ∃! x, x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y' := by +lemma log_exists_unique_pos {y : V} (hy : 0 < y) : ∃! x, x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y' := by have : ∃ x < y, ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y' := by induction y using hierarchy_polynomial_induction_oRing_sigma₀ · definability @@ -42,62 +42,62 @@ lemma log_exists_unique_pos {y : M} (hy : 0 < y) : ∃! x, x < y ∧ ∃ y' ≤ _ ≤ y := hzy' simp at this) -lemma log_exists_unique (y : M) : ∃! x, (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y') := by +lemma log_exists_unique (y : V) : ∃! x, (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y') := by by_cases hy : y = 0 · rcases hy; simp · simp [hy, pos_iff_ne_zero.mpr hy, log_exists_unique_pos] -def log (a : M) : M := Classical.choose! (log_exists_unique a) +def log (a : V) : V := Classical.choose! (log_exists_unique a) -@[simp] lemma log_zero : log (0 : M) = 0 := - (Classical.choose!_spec (log_exists_unique (0 : M))).1 rfl +@[simp] lemma log_zero : log (0 : V) = 0 := + (Classical.choose!_spec (log_exists_unique (0 : V))).1 rfl -lemma log_pos {y : M} (pos : 0 < y) : ∃ y' ≤ y, Exponential (log y) y' ∧ y < 2 * y' := +lemma log_pos {y : V} (pos : 0 < y) : ∃ y' ≤ y, Exponential (log y) y' ∧ y < 2 * y' := ((Classical.choose!_spec (log_exists_unique y)).2 pos).2 -lemma log_lt_self_of_pos {y : M} (pos : 0 < y) : log y < y := +lemma log_lt_self_of_pos {y : V} (pos : 0 < y) : log y < y := ((Classical.choose!_spec (log_exists_unique y)).2 pos).1 -@[simp] lemma log_le_self (a : M) : log a ≤ a := by +@[simp] lemma log_le_self (a : V) : log a ≤ a := by rcases zero_le a with (rfl | pos) · simp · exact le_of_lt <| log_lt_self_of_pos pos -lemma log_graph {x y : M} : x = log y ↔ (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y') := Classical.choose!_eq_iff _ +lemma log_graph {x y : V} : x = log y ↔ (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' ≤ y, Exponential x y' ∧ y < 2 * y') := Classical.choose!_eq_iff _ def _root_.LO.FirstOrder.Arith.logDef : 𝚺₀.Semisentence 2 := .mkSigma “x y | (y = 0 → x = 0) ∧ (0 < y → x < y ∧ ∃ y' <⁺ y, !exponentialDef x y' ∧ y < 2 * y')” (by simp) -lemma log_defined : 𝚺₀-Function₁ (log : M → M) via logDef := by +lemma log_defined : 𝚺₀-Function₁ (log : V → V) via logDef := by intro v; simp [logDef, log_graph, ←le_iff_lt_succ, numeral_eq_natCast] @[simp] lemma log_defined_iff (v) : - Semiformula.Evalbm M v logDef.val ↔ v 0 = log (v 1) := log_defined.df.iff v + Semiformula.Evalbm V v logDef.val ↔ v 0 = log (v 1) := log_defined.df.iff v -instance log_definable : 𝚺₀-Function₁ (log : M → M) := Defined.to_definable _ log_defined +instance log_definable : 𝚺₀-Function₁ (log : V → V) := log_defined.to_definable -instance : Bounded₁ (log : M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₁ (log : V → V) := ⟨#0, λ _ ↦ by simp⟩ -lemma log_eq_of_pos {x y : M} (pos : 0 < y) {y'} (H : Exponential x y') (hy' : y' ≤ y) (hy : y < 2 * y') : log y = x := +lemma log_eq_of_pos {x y : V} (pos : 0 < y) {y'} (H : Exponential x y') (hy' : y' ≤ y) (hy : y < 2 * y') : log y = x := (log_exists_unique_pos pos).unique ⟨log_lt_self_of_pos pos, log_pos pos⟩ ⟨lt_of_lt_of_le H.lt hy', y', hy', H, hy⟩ -@[simp] lemma log_one : log (1 : M) = 0 := log_eq_of_pos (by simp) (y' := 1) (by simp) (by rfl) (by simp [one_lt_two]) +@[simp] lemma log_one : log (1 : V) = 0 := log_eq_of_pos (by simp) (y' := 1) (by simp) (by rfl) (by simp [one_lt_two]) -@[simp] lemma log_two : log (2 : M) = 1 := log_eq_of_pos (by simp) (y' := 2) (by simp) (by rfl) (by simp [one_lt_two]) +@[simp] lemma log_two : log (2 : V) = 1 := log_eq_of_pos (by simp) (y' := 2) (by simp) (by rfl) (by simp [one_lt_two]) -lemma log_two_mul_of_pos {y : M} (pos : 0 < y) : log (2 * y) = log y + 1 := by +lemma log_two_mul_of_pos {y : V} (pos : 0 < y) : log (2 * y) = log y + 1 := by rcases log_pos pos with ⟨y', hy', H, hy⟩ exact log_eq_of_pos (by simpa using pos) (Exponential.exponential_succ_mul_two.mpr H) (by simpa using hy') (by simpa using hy) -lemma log_two_mul_add_one_of_pos {y : M} (pos : 0 < y) : log (2 * y + 1) = log y + 1 := by +lemma log_two_mul_add_one_of_pos {y : V} (pos : 0 < y) : log (2 * y + 1) = log y + 1 := by rcases log_pos pos with ⟨y', hy', H, hy⟩ exact log_eq_of_pos (by simp) (Exponential.exponential_succ_mul_two.mpr H) (le_trans (by simpa using hy') le_self_add) (two_mul_add_one_lt_two_mul_of_lt hy) -lemma Exponential.log_eq_of_exp {x y : M} (H : Exponential x y) : log y = x := +lemma Exponential.log_eq_of_exp {x y : V} (H : Exponential x y) : log y = x := log_eq_of_pos H.range_pos H (by { rfl }) (lt_mul_of_pos_of_one_lt_left H.range_pos one_lt_two) -lemma exponential_of_pow2 {p : M} (pp : Pow2 p) : Exponential (log p) p := by +lemma exponential_of_pow2 {p : V} (pp : Pow2 p) : Exponential (log p) p := by rcases log_pos pp.pos with ⟨q, hq, H, hp⟩ suffices p = q by simpa [this] using H by_contra ne @@ -107,7 +107,7 @@ lemma exponential_of_pow2 {p : M} (pp : Pow2 p) : Exponential (log p) p := by _ < 2 * q := hp simp at this -lemma log_mul_pow2_add_of_lt {a p b : M} (pos : 0 < a) (pp : Pow2 p) (hb : b < p) : log (a * p + b) = log a + log p := by +lemma log_mul_pow2_add_of_lt {a p b : V} (pos : 0 < a) (pp : Pow2 p) (hb : b < p) : log (a * p + b) = log a + log p := by rcases log_pos pos with ⟨a', ha', Ha, ha⟩ rcases log_pos pp.pos with ⟨p', hp', Hp, hp⟩ exact log_eq_of_pos (lt_of_lt_of_le (mul_pos pos pp.pos) le_self_add) @@ -118,10 +118,10 @@ lemma log_mul_pow2_add_of_lt {a p b : M} (pos : 0 < a) (pp : Pow2 p) (hb : b < p _ = (a + 1) * p := by simp [add_mul] _ ≤ 2 * (a' * p) := by simp [←mul_assoc]; exact mul_le_mul_right (lt_iff_succ_le.mp ha)) -lemma log_mul_pow2 {a p : M} (pos : 0 < a) (pp : Pow2 p) : log (a * p) = log a + log p := by +lemma log_mul_pow2 {a p : V} (pos : 0 < a) (pp : Pow2 p) : log (a * p) = log a + log p := by simpa using log_mul_pow2_add_of_lt pos pp pp.pos -lemma log_monotone {a b : M} (h : a ≤ b) : log a ≤ log b := by +lemma log_monotone {a b : V} (h : a ≤ b) : log a ≤ log b := by rcases zero_le a with (rfl | posa) · simp rcases zero_le b with (rfl | posb) @@ -137,22 +137,22 @@ lemma log_monotone {a b : M} (h : a ≤ b) : log a ≤ log b := by _ ≤ b := h simp_all -def binaryLength (a : M) : M := if 0 < a then log a + 1 else 0 +def binaryLength (a : V) : V := if 0 < a then log a + 1 else 0 -scoped instance : Length M := ⟨binaryLength⟩ +scoped instance : Length V := ⟨binaryLength⟩ -lemma length_eq_binaryLength (a : M) : ‖a‖ = if 0 < a then log a + 1 else 0 := rfl +lemma length_eq_binaryLength (a : V) : ‖a‖ = if 0 < a then log a + 1 else 0 := rfl -@[simp] lemma length_zero : ‖(0 : M)‖ = 0 := by simp [length_eq_binaryLength] +@[simp] lemma length_zero : ‖(0 : V)‖ = 0 := by simp [length_eq_binaryLength] -lemma length_of_pos {a : M} (pos : 0 < a) : ‖a‖ = log a + 1 := by simp [length_eq_binaryLength, pos] +lemma length_of_pos {a : V} (pos : 0 < a) : ‖a‖ = log a + 1 := by simp [length_eq_binaryLength, pos] -@[simp] lemma length_le (a : M) : ‖a‖ ≤ a := by +@[simp] lemma length_le (a : V) : ‖a‖ ≤ a := by rcases zero_le a with (rfl | pos) · simp · simp [pos, length_of_pos, ←lt_iff_succ_le, log_lt_self_of_pos] -lemma length_graph {i a : M} : i = ‖a‖ ↔ (0 < a → ∃ k ≤ a, k = log a ∧ i = k + 1) ∧ (a = 0 → i = 0) := by +lemma length_graph {i a : V} : i = ‖a‖ ↔ (0 < a → ∃ k ≤ a, k = log a ∧ i = k + 1) ∧ (a = 0 → i = 0) := by rcases zero_le a with (rfl | pos) · simp · simp [length_of_pos, pos, pos_iff_ne_zero.mp pos] @@ -163,100 +163,100 @@ lemma length_graph {i a : M} : i = ‖a‖ ↔ (0 < a → ∃ k ≤ a, k = log a def _root_.LO.FirstOrder.Arith.lengthDef : 𝚺₀.Semisentence 2 := .mkSigma “i a | (0 < a → ∃ k <⁺ a, !logDef k a ∧ i = k + 1) ∧ (a = 0 → i = 0)” (by simp) -lemma length_defined : 𝚺₀-Function₁ (‖·‖ : M → M) via lengthDef := by +lemma length_defined : 𝚺₀-Function₁ (‖·‖ : V → V) via lengthDef := by intro v; simp [lengthDef, length_graph, ←le_iff_lt_succ] @[simp] lemma length_defined_iff (v) : - Semiformula.Evalbm M v lengthDef.val ↔ v 0 = ‖v 1‖ := length_defined.df.iff v + Semiformula.Evalbm V v lengthDef.val ↔ v 0 = ‖v 1‖ := length_defined.df.iff v -instance length_definable : 𝚺₀-Function₁ (‖·‖ : M → M) := Defined.to_definable _ length_defined +instance length_definable : 𝚺₀-Function₁ (‖·‖ : V → V) := length_defined.to_definable -instance : Bounded₁ (‖·‖ : M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₁ (‖·‖ : V → V) := ⟨#0, λ _ ↦ by simp⟩ -@[simp] lemma length_one : ‖(1 : M)‖ = 1 := by simp [length_eq_binaryLength] +@[simp] lemma length_one : ‖(1 : V)‖ = 1 := by simp [length_eq_binaryLength] -lemma Exponential.length_eq {x y : M} (H : Exponential x y) : ‖y‖ = x + 1 := by +lemma Exponential.length_eq {x y : V} (H : Exponential x y) : ‖y‖ = x + 1 := by simp [length_of_pos H.range_pos]; exact H.log_eq_of_exp -lemma length_two_mul_of_pos {a : M} (pos : 0 < a) : ‖2 * a‖ = ‖a‖ + 1 := by +lemma length_two_mul_of_pos {a : V} (pos : 0 < a) : ‖2 * a‖ = ‖a‖ + 1 := by simp [pos, length_of_pos, log_two_mul_of_pos] -lemma length_two_mul_add_one (a : M) : ‖2 * a + 1‖ = ‖a‖ + 1 := by +lemma length_two_mul_add_one (a : V) : ‖2 * a + 1‖ = ‖a‖ + 1 := by rcases zero_le a with (rfl | pos) · simp · simp [pos, length_of_pos, log_two_mul_add_one_of_pos] -lemma length_mul_pow2_add_of_lt {a p b : M} (pos : 0 < a) (pp : Pow2 p) (hb : b < p) : ‖a * p + b‖ = ‖a‖ + log p := by +lemma length_mul_pow2_add_of_lt {a p b : V} (pos : 0 < a) (pp : Pow2 p) (hb : b < p) : ‖a * p + b‖ = ‖a‖ + log p := by simp [length_of_pos, pos, pp.pos, log_mul_pow2_add_of_lt pos pp hb, add_right_comm (log a) (log p) 1] -lemma length_mul_pow2 {a p : M} (pos : 0 < a) (pp : Pow2 p) : ‖a * p‖ = ‖a‖ + log p := by +lemma length_mul_pow2 {a p : V} (pos : 0 < a) (pp : Pow2 p) : ‖a * p‖ = ‖a‖ + log p := by simp [length_of_pos, pos, pp.pos, log_mul_pow2 pos pp, add_right_comm (log a) (log p) 1] -lemma length_monotone {a b : M} (h : a ≤ b) : ‖a‖ ≤ ‖b‖ := by +lemma length_monotone {a b : V} (h : a ≤ b) : ‖a‖ ≤ ‖b‖ := by rcases zero_le a with (rfl | posa) · simp · simp [length_of_pos posa, length_of_pos (lt_of_lt_of_le posa h)] exact log_monotone h -lemma pos_of_lt_length {a b : M} (h : a < ‖b‖) : 0 < b := by +lemma pos_of_lt_length {a b : V} (h : a < ‖b‖) : 0 < b := by by_contra A; rcases (show b = 0 from by simpa using A); simp_all -@[simp] lemma length_pos_iff {a : M} : 0 < ‖a‖ ↔ 0 < a := +@[simp] lemma length_pos_iff {a : V} : 0 < ‖a‖ ↔ 0 < a := ⟨by intro h; by_contra A; rcases (show a = 0 from by simpa using A); simp_all, by intro h; exact pos_iff_one_le.mpr (by simpa using length_monotone (pos_iff_one_le.mp h))⟩ -@[simp] lemma length_eq_zero_iff {a : M} : ‖a‖ = 0 ↔ a = 0 := not_iff_not.mp (by simp [←pos_iff_ne_zero]) +@[simp] lemma length_eq_zero_iff {a : V} : ‖a‖ = 0 ↔ a = 0 := not_iff_not.mp (by simp [←pos_iff_ne_zero]) -lemma le_log_of_lt_length {a b : M} (h : a < ‖b‖) : a ≤ log b := by +lemma le_log_of_lt_length {a b : V} (h : a < ‖b‖) : a ≤ log b := by have : 0 < b := pos_of_lt_length h exact le_iff_lt_succ.mpr (by simpa [length_of_pos this] using h) -lemma exponential_log_le_self {a b : M} (pos : 0 < a) (h : Exponential (log a) b) : b ≤ a := by +lemma exponential_log_le_self {a b : V} (pos : 0 < a) (h : Exponential (log a) b) : b ≤ a := by rcases log_pos pos with ⟨_, _, H, _⟩; rcases H.uniq h assumption -lemma lt_exponential_log_self {a b : M} (h : Exponential (log a) b) : a < 2 * b := by +lemma lt_exponential_log_self {a b : V} (h : Exponential (log a) b) : a < 2 * b := by rcases zero_le a with (rfl | pos) · simp at h; simp [h] rcases log_pos pos with ⟨_, _, H, _⟩; rcases H.uniq h assumption -lemma lt_exp_len_self {a b : M} (h : Exponential ‖a‖ b) : a < b := by +lemma lt_exp_len_self {a b : V} (h : Exponential ‖a‖ b) : a < b := by rcases zero_le a with (rfl | pos) · simp at h; simp [h] simp [length_of_pos pos] at h rcases Exponential.exponential_succ.mp h with ⟨b, rfl, H⟩ exact lt_exponential_log_self H -lemma le_iff_le_log_of_exp {x y a : M} (H : Exponential x y) (pos : 0 < a) : y ≤ a ↔ x ≤ log a := +lemma le_iff_le_log_of_exp {x y a : V} (H : Exponential x y) (pos : 0 < a) : y ≤ a ↔ x ≤ log a := ⟨by rcases H.log_eq_of_exp; exact log_monotone, fun h ↦ by rcases log_pos pos with ⟨a', ha', Haa', _⟩; exact le_trans (Exponential.monotone_le H Haa' h) ha'⟩ -lemma le_iff_lt_length_of_exp {x y a : M} (H : Exponential x y) : y ≤ a ↔ x < ‖a‖ := by +lemma le_iff_lt_length_of_exp {x y a : V} (H : Exponential x y) : y ≤ a ↔ x < ‖a‖ := by rcases zero_le a with (rfl | pos) · simp; exact pos_iff_ne_zero.mp H.range_pos simp [le_iff_le_log_of_exp H pos, length_of_pos pos, ←le_iff_lt_succ] -lemma Exponential.lt_iff_log_lt {x y a : M} (H : Exponential x y) (pos : 0 < a) : a < y ↔ log a < x := +lemma Exponential.lt_iff_log_lt {x y a : V} (H : Exponential x y) (pos : 0 < a) : a < y ↔ log a < x := not_iff_not.mp (by simpa using le_iff_le_log_of_exp H pos) -lemma Exponential.lt_iff_len_le {x y a : M} (H : Exponential x y) : a < y ↔ ‖a‖ ≤ x := not_iff_not.mp (by simpa using le_iff_lt_length_of_exp H) +lemma Exponential.lt_iff_len_le {x y a : V} (H : Exponential x y) : a < y ↔ ‖a‖ ≤ x := not_iff_not.mp (by simpa using le_iff_lt_length_of_exp H) -lemma Exponential.le_of_lt_length {x y a : M} (H : Exponential x y) : x < ‖a‖ → y ≤ a := fun h ↦ +lemma Exponential.le_of_lt_length {x y a : V} (H : Exponential x y) : x < ‖a‖ → y ≤ a := fun h ↦ (le_iff_lt_length_of_exp H).mpr h -lemma Exponential.le_log {x y : M} (H : Exponential x y) : x ≤ log y := (le_iff_le_log_of_exp H H.range_pos).mp (by rfl) +lemma Exponential.le_log {x y : V} (H : Exponential x y) : x ≤ log y := (le_iff_le_log_of_exp H H.range_pos).mp (by rfl) -lemma Exponential.lt_length {x y : M} (H : Exponential x y) : x < ‖y‖ := (le_iff_lt_length_of_exp H).mp (by rfl) +lemma Exponential.lt_length {x y : V} (H : Exponential x y) : x < ‖y‖ := (le_iff_lt_length_of_exp H).mp (by rfl) -lemma lt_exponential_length {a b : M} (h : Exponential ‖a‖ b) : a < b := by +lemma lt_exponential_length {a b : V} (h : Exponential ‖a‖ b) : a < b := by rcases zero_le a with (rfl | pos) · simp at h; simp [h] simp [length_of_pos pos] at h rcases Exponential.exponential_succ.mp h with ⟨b, rfl, H⟩ exact lt_exponential_log_self H -lemma sq_len_le_three_mul (a : M) : ‖a‖ ^ 2 ≤ 3 * a := by +lemma sq_len_le_three_mul (a : V) : ‖a‖ ^ 2 ≤ 3 * a := by induction a using hierarchy_polynomial_induction_oRing_sigma₀ · definability case zero => simp @@ -280,7 +280,7 @@ lemma sq_len_le_three_mul (a : M) : ‖a‖ ^ 2 ≤ 3 * a := by _ = 3 * (2 * a) := by simp_all only [←two_add_one_eq_three, two_mul, add_mul, add_assoc, one_mul] _ ≤ 3 * (2 * a + 1) := by simp -lemma brange_exists_unique (a : M) : ∀ x < ‖a‖, ∃! y, Exponential x y := by +lemma brange_exists_unique (a : V) : ∀ x < ‖a‖, ∃! y, Exponential x y := by suffices ∀ x < ‖a‖, ∃ y ≤ a, Exponential x y by intro x hx; rcases this x hx with ⟨_, _, H⟩ exact ExistsUnique.intro _ H (fun y' H' ↦ H'.uniq H) @@ -298,37 +298,37 @@ lemma brange_exists_unique (a : M) : ∀ x < ‖a‖, ∃! y, Exponential x y := have : 2 * y ≤ a := (le_iff_le_log_of_exp H.succ this).mpr (le_log_of_lt_length hx) exact ⟨2 * y, this, H.succ⟩ -lemma bexp_exists_unique (a x : M) : ∃! y, (x < ‖a‖ → Exponential x y) ∧ (‖a‖ ≤ x → y = 0) := by +lemma bexp_exists_unique (a x : V) : ∃! y, (x < ‖a‖ → Exponential x y) ∧ (‖a‖ ≤ x → y = 0) := by by_cases hx : x < ‖a‖ · simp [hx, show ¬‖a‖ ≤ x from by simpa using hx, log_exists_unique_pos] exact brange_exists_unique a x hx · simp [hx, show ‖a‖ ≤ x from by simpa using hx] /-- `bexp a x = exp x` if `x < ‖a‖`; `= 0` o.w.-/ -def bexp (a x : M) : M := Classical.choose! (bexp_exists_unique a x) +def bexp (a x : V) : V := Classical.choose! (bexp_exists_unique a x) -lemma exp_bexp_of_lt {a x : M} (h : x < ‖a‖) : Exponential x (bexp a x) := +lemma exp_bexp_of_lt {a x : V} (h : x < ‖a‖) : Exponential x (bexp a x) := (Classical.choose!_spec (bexp_exists_unique a x)).1 h -lemma bexp_eq_zero_of_le {a x : M} (h : ‖a‖ ≤ x) : bexp a x = 0 := +lemma bexp_eq_zero_of_le {a x : V} (h : ‖a‖ ≤ x) : bexp a x = 0 := (Classical.choose!_spec (bexp_exists_unique a x)).2 h -@[simp] lemma bexp_zero (x : M): bexp 0 x = 0 := bexp_eq_zero_of_le (by simp) +@[simp] lemma bexp_zero (x : V): bexp 0 x = 0 := bexp_eq_zero_of_le (by simp) -@[simp] lemma exp_bexp_of_lt_iff {a x : M} : Exponential x (bexp a x) ↔ x < ‖a‖ := +@[simp] lemma exp_bexp_of_lt_iff {a x : V} : Exponential x (bexp a x) ↔ x < ‖a‖ := ⟨by intro h; by_contra A have : bexp a x = 0 := bexp_eq_zero_of_le (not_lt.mp A) simp [this] at h have := h.range_pos; simp_all, exp_bexp_of_lt⟩ -@[simp] lemma bexp_le_self (a x : M) : bexp a x ≤ a := by +@[simp] lemma bexp_le_self (a x : V) : bexp a x ≤ a := by rcases show x < ‖a‖ ∨ ‖a‖ ≤ x from lt_or_ge _ _ with (lt | le) · have : 0 < a := pos_of_lt_length lt exact (le_iff_le_log_of_exp (exp_bexp_of_lt lt) this).mpr (le_log_of_lt_length lt) · simp [bexp_eq_zero_of_le le] -lemma bexp_graph {y a x : M} : y = bexp a x ↔ ∃ l ≤ a, l = ‖a‖ ∧ (x < l → Exponential x y) ∧ (l ≤ x → y = 0) := +lemma bexp_graph {y a x : V} : y = bexp a x ↔ ∃ l ≤ a, l = ‖a‖ ∧ (x < l → Exponential x y) ∧ (l ≤ x → y = 0) := ⟨by rintro rfl; exact ⟨‖a‖, by simp, rfl, exp_bexp_of_lt, bexp_eq_zero_of_le⟩, by rintro ⟨_, _, rfl, hlt, hle⟩ rcases show x < ‖a‖ ∨ ‖a‖ ≤ x from lt_or_ge _ _ with (lt | le) @@ -338,152 +338,152 @@ lemma bexp_graph {y a x : M} : y = bexp a x ↔ ∃ l ≤ a, l = ‖a‖ ∧ (x def _root_.LO.FirstOrder.Arith.bexpDef : 𝚺₀.Semisentence 3 := .mkSigma “y a x | ∃ l <⁺ a, !lengthDef l a ∧ (x < l → !exponentialDef x y) ∧ (l ≤ x → y = 0)” (by simp) -lemma bexp_defined : 𝚺₀-Function₂ (bexp : M → M → M) via bexpDef := by +lemma bexp_defined : 𝚺₀-Function₂ (bexp : V → V → V) via bexpDef := by intro v; simp [bexpDef, bexp_graph, ←le_iff_lt_succ] @[simp] lemma bexp_defined_iff (v) : - Semiformula.Evalbm M v bexpDef.val ↔ v 0 = bexp (v 1) (v 2) := bexp_defined.df.iff v + Semiformula.Evalbm V v bexpDef.val ↔ v 0 = bexp (v 1) (v 2) := bexp_defined.df.iff v -instance bexp_definable : 𝚺₀-Function₂ (bexp : M → M → M) := Defined.to_definable _ bexp_defined +instance bexp_definable : 𝚺₀-Function₂ (bexp : V → V → V) := bexp_defined.to_definable -instance : Bounded₂ (bexp : M → M → M) := ⟨#0, λ _ ↦ by simp⟩ +instance : Bounded₂ (bexp : V → V → V) := ⟨#0, λ _ ↦ by simp⟩ -lemma bexp_monotone_iff {a i j : M} (hi : i < ‖a‖) (hj : j < ‖a‖) : bexp a i < bexp a j ↔ i < j := +lemma bexp_monotone_iff {a i j : V} (hi : i < ‖a‖) (hj : j < ‖a‖) : bexp a i < bexp a j ↔ i < j := Iff.symm <| Exponential.monotone_iff (by simp [hi]) (by simp [hj]) -lemma bexp_monotone_le_iff {a i j : M} (hi : i < ‖a‖) (hj : j < ‖a‖) : bexp a i ≤ bexp a j ↔ i ≤ j := +lemma bexp_monotone_le_iff {a i j : V} (hi : i < ‖a‖) (hj : j < ‖a‖) : bexp a i ≤ bexp a j ↔ i ≤ j := Iff.symm <| Exponential.monotone_le_iff (by simp [hi]) (by simp [hj]) -lemma bexp_eq_of_lt_length {i a a' : M} (ha : i < ‖a‖) (ha' : i < ‖a'‖) : bexp a i = bexp a' i := by +lemma bexp_eq_of_lt_length {i a a' : V} (ha : i < ‖a‖) (ha' : i < ‖a'‖) : bexp a i = bexp a' i := by have H : Exponential i (bexp a i) := by simp [ha] have H' : Exponential i (bexp a' i) := by simp [ha'] exact H.uniq H' -@[simp] lemma bexp_pow2 {a x : M} (h : x < ‖a‖) : Pow2 (bexp a x) := (exp_bexp_of_lt h).range_pow2 +@[simp] lemma bexp_pow2 {a x : V} (h : x < ‖a‖) : Pow2 (bexp a x) := (exp_bexp_of_lt h).range_pow2 -@[simp] lemma lt_bexp {a x : M} (h : x < ‖a‖) : x < bexp a x := (exp_bexp_of_lt h).lt +@[simp] lemma lt_bexp {a x : V} (h : x < ‖a‖) : x < bexp a x := (exp_bexp_of_lt h).lt -@[simp] lemma bexp_pos {a x : M} (h : x < ‖a‖) : 0 < bexp a x := (exp_bexp_of_lt h).range_pos +@[simp] lemma bexp_pos {a x : V} (h : x < ‖a‖) : 0 < bexp a x := (exp_bexp_of_lt h).range_pos -lemma lt_bexp_len {a x : M} (h : ‖x‖ < ‖a‖) : x < bexp a ‖x‖ := lt_exp_len_self (exp_bexp_of_lt h) +lemma lt_bexp_len {a x : V} (h : ‖x‖ < ‖a‖) : x < bexp a ‖x‖ := lt_exp_len_self (exp_bexp_of_lt h) -lemma bexp_eq_of_exp {a x : M} (h : x < ‖a‖) (H : Exponential x y) : bexp a x = y := (exp_bexp_of_lt h).uniq H +lemma bexp_eq_of_exp {a x : V} (h : x < ‖a‖) (H : Exponential x y) : bexp a x = y := (exp_bexp_of_lt h).uniq H -lemma log_bexp {a x : M} (h : x < ‖a‖) : log (bexp a x) = x := Exponential.log_eq_of_exp (exp_bexp_of_lt h) +lemma log_bexp {a x : V} (h : x < ‖a‖) : log (bexp a x) = x := Exponential.log_eq_of_exp (exp_bexp_of_lt h) -lemma len_bexp {a x : M} (h : x < ‖a‖) : ‖bexp a x‖ = x + 1 := by rw [length_of_pos (bexp_pos h), log_bexp h] +lemma len_bexp {a x : V} (h : x < ‖a‖) : ‖bexp a x‖ = x + 1 := by rw [length_of_pos (bexp_pos h), log_bexp h] -@[simp] lemma bexp_zero_zero : bexp (0 : M) 0 = 0 := bexp_eq_zero_of_le (by simp) +@[simp] lemma bexp_zero_zero : bexp (0 : V) 0 = 0 := bexp_eq_zero_of_le (by simp) -@[simp] lemma bexp_pos_zero {a : M} (h : 0 < a) : bexp a 0 = 1 := bexp_eq_of_exp (by simpa) (by simp) +@[simp] lemma bexp_pos_zero {a : V} (h : 0 < a) : bexp a 0 = 1 := bexp_eq_of_exp (by simpa) (by simp) -lemma bexp_monotone {a₁ x₁ a₂ x₂ : M} (h₁ : x₁ < ‖a₁‖) (h₂ : x₂ < ‖a₂‖) : +lemma bexp_monotone {a₁ x₁ a₂ x₂ : V} (h₁ : x₁ < ‖a₁‖) (h₂ : x₂ < ‖a₂‖) : bexp a₁ x₁ < bexp a₂ x₂ ↔ x₁ < x₂ := Iff.symm <| (exp_bexp_of_lt h₁).monotone_iff (exp_bexp_of_lt h₂) -lemma bexp_monotone_le {a₁ x₁ a₂ x₂ : M} (h₁ : x₁ < ‖a₁‖) (h₂ : x₂ < ‖a₂‖) : +lemma bexp_monotone_le {a₁ x₁ a₂ x₂ : V} (h₁ : x₁ < ‖a₁‖) (h₂ : x₂ < ‖a₂‖) : bexp a₁ x₁ ≤ bexp a₂ x₂ ↔ x₁ ≤ x₂ := Iff.symm <| (exp_bexp_of_lt h₁).monotone_le_iff (exp_bexp_of_lt h₂) -lemma bexp_add {x₁ x₂ a : M} (h : x₁ + x₂ < ‖a‖) : +lemma bexp_add {x₁ x₂ a : V} (h : x₁ + x₂ < ‖a‖) : bexp a (x₁ + x₂) = bexp a x₁ * bexp a x₂ := (exp_bexp_of_lt h).uniq ((exp_bexp_of_lt (lt_of_le_of_lt le_self_add h)).add_mul (exp_bexp_of_lt (lt_of_le_of_lt le_add_self h))) -lemma bexp_two_mul {a a' x : M} (hx : 2 * x < ‖a‖) (hx' : x < ‖a'‖) : +lemma bexp_two_mul {a a' x : V} (hx : 2 * x < ‖a‖) (hx' : x < ‖a'‖) : bexp a (2 * x) = (bexp a' x) ^ 2 := bexp_eq_of_exp hx (exp_bexp_of_lt hx').bit_zero -lemma bexp_two_mul_succ {a i : M} : bexp (2 * a) (i + 1) = 2 * bexp a i := by +lemma bexp_two_mul_succ {a i : V} : bexp (2 * a) (i + 1) = 2 * bexp a i := by rcases zero_le a with (rfl | pos) · simp rcases show i ≥ ‖a‖ ∨ i < ‖a‖ from le_or_lt ‖a‖ i with (h | h) · simp [bexp_eq_zero_of_le, h, show ‖2 * a‖ ≤ i + 1 from by simp [length_two_mul_of_pos pos, h]] · exact bexp_eq_of_exp (by simp [length_two_mul_of_pos pos, h]) (exp_bexp_of_lt h).succ -lemma bexp_two_mul_add_one_succ {a i : M} : bexp (2 * a + 1) (i + 1) = 2 * bexp a i := by +lemma bexp_two_mul_add_one_succ {a i : V} : bexp (2 * a + 1) (i + 1) = 2 * bexp a i := by rcases show i ≥ ‖a‖ ∨ i < ‖a‖ from le_or_lt ‖a‖ i with (h | h) · simp [bexp_eq_zero_of_le, h, show ‖2 * a + 1‖ ≤ i + 1 from by simp [length_two_mul_add_one, h]] · exact bexp_eq_of_exp (by simp [length_two_mul_add_one, h]) (exp_bexp_of_lt h).succ -def fbit (a i : M) : M := (a / bexp a i) % 2 +def fbit (a i : V) : V := (a / bexp a i) % 2 -@[simp] lemma fbit_lt_two (a i : M) : fbit a i < 2 := by simp [fbit] +@[simp] lemma fbit_lt_two (a i : V) : fbit a i < 2 := by simp [fbit] -@[simp] lemma fbit_le_one (a i : M) : fbit a i ≤ 1 := lt_two_iff_le_one.mp (by simp [fbit]) +@[simp] lemma fbit_le_one (a i : V) : fbit a i ≤ 1 := lt_two_iff_le_one.mp (by simp [fbit]) -lemma fbit_eq_one_iff {a i : M} : fbit a i = 1 ↔ LenBit (bexp a i) a := by simp [fbit, LenBit.iff_rem] +lemma fbit_eq_one_iff {a i : V} : fbit a i = 1 ↔ LenBit (bexp a i) a := by simp [fbit, LenBit.iff_rem] -lemma fbit_eq_zero_iff {a i : M} : fbit a i = 0 ↔ ¬LenBit (bexp a i) a := by simp [fbit, LenBit.iff_rem] +lemma fbit_eq_zero_iff {a i : V} : fbit a i = 0 ↔ ¬LenBit (bexp a i) a := by simp [fbit, LenBit.iff_rem] -lemma fbit_eq_zero_of_le {a i : M} (hi : ‖a‖ ≤ i) : fbit a i = 0 := by simp [fbit, bexp_eq_zero_of_le hi] +lemma fbit_eq_zero_of_le {a i : V} (hi : ‖a‖ ≤ i) : fbit a i = 0 := by simp [fbit, bexp_eq_zero_of_le hi] def _root_.LO.FirstOrder.Arith.fbitDef : 𝚺₀.Semisentence 3 := .mkSigma “b a i | ∃ x <⁺ a, !bexpDef x a i ∧ ∃ y <⁺ a, !divDef y a x ∧ !remDef b y 2” (by simp) -lemma fbit_defined : 𝚺₀-Function₂ (fbit : M → M → M) via fbitDef := by +lemma fbit_defined : 𝚺₀-Function₂ (fbit : V → V → V) via fbitDef := by intro v; simp [fbitDef, ←le_iff_lt_succ, fbit, numeral_eq_natCast] constructor · intro h; exact ⟨bexp (v 1) (v 2), by simp, rfl, _, by simp, rfl, h⟩ · rintro ⟨_, _, rfl, _, _, rfl, h⟩; exact h @[simp] lemma fbit_defined_iff (v) : - Semiformula.Evalbm M v fbitDef.val ↔ v 0 = fbit (v 1) (v 2) := fbit_defined.df.iff v + Semiformula.Evalbm V v fbitDef.val ↔ v 0 = fbit (v 1) (v 2) := fbit_defined.df.iff v -instance fbit_definable : 𝚺₀-Function₂ (fbit : M → M → M) := Defined.to_definable _ fbit_defined +instance fbit_definable : 𝚺₀-Function₂ (fbit : V → V → V) := fbit_defined.to_definable -instance : Bounded₂ (fbit : M → M → M) := ⟨‘1’, λ _ ↦ by simp⟩ +instance : Bounded₂ (fbit : V → V → V) := ⟨‘1’, λ _ ↦ by simp⟩ -@[simp] lemma fbit_zero (i : M) : fbit 0 i = 0 := by simp [fbit] +@[simp] lemma fbit_zero (i : V) : fbit 0 i = 0 := by simp [fbit] -@[simp] lemma fbit_mul_two_mul (a i : M) : fbit (2 * a) (i + 1) = fbit a i := by +@[simp] lemma fbit_mul_two_mul (a i : V) : fbit (2 * a) (i + 1) = fbit a i := by simp [fbit, bexp_two_mul_succ, div_cancel_left] -@[simp] lemma fbit_mul_two_add_one_mul (a i : M) : fbit (2 * a + 1) (i + 1) = fbit a i := by +@[simp] lemma fbit_mul_two_add_one_mul (a i : V) : fbit (2 * a + 1) (i + 1) = fbit a i := by simp [fbit, bexp_two_mul_add_one_succ, div_cancel_left, div_mul] -@[simp] lemma fbit_two_mul_zero_eq_zero (a : M) : fbit (2 * a) 0 = 0 := by +@[simp] lemma fbit_two_mul_zero_eq_zero (a : V) : fbit (2 * a) 0 = 0 := by rcases zero_le a with (rfl | pos) · simp · have : bexp (2 * a) 0 = 1 := bexp_eq_of_exp (by simp [pos]) (by simp) simp [fbit, this] -@[simp] lemma fbit_two_mul_add_one_zero_eq_one (a : M) : fbit (2 * a + 1) 0 = 1 := by simp [fbit, one_lt_two] +@[simp] lemma fbit_two_mul_add_one_zero_eq_one (a : V) : fbit (2 * a + 1) 0 = 1 := by simp [fbit, one_lt_two] end ISigma₀ section ISigma₁ -variable [M ⊧ₘ* 𝐈𝚺₁] +variable [V ⊧ₘ* 𝐈𝚺₁] -@[simp] lemma log_exponential (a : M) : log (exp a) = a := (exponential_exp a).log_eq_of_exp +@[simp] lemma log_exponential (a : V) : log (exp a) = a := (exponential_exp a).log_eq_of_exp -lemma exp_log_le_self {a : M} (pos : 0 < a) : exp (log a) ≤ a := by +lemma exp_log_le_self {a : V} (pos : 0 < a) : exp (log a) ≤ a := by rcases log_pos pos with ⟨_, _, H, _⟩ rcases H.uniq (exponential_exp (log a)) assumption -lemma lt_two_mul_exponential_log {a : M} (pos : 0 < a) : a < 2 * exp (log a) := by +lemma lt_two_mul_exponential_log {a : V} (pos : 0 < a) : a < 2 * exp (log a) := by rcases log_pos pos with ⟨_, _, H, _⟩ rcases H.uniq (exponential_exp (log a)) assumption -@[simp] lemma length_exponential (a : M) : ‖exp a‖ = a + 1 := by +@[simp] lemma length_exponential (a : V) : ‖exp a‖ = a + 1 := by simp [length_of_pos (exp_pos a)] -lemma exp_add (a b : M) : exp (a + b) = exp a * exp b := +lemma exp_add (a b : V) : exp (a + b) = exp a * exp b := exp_of_exponential (Exponential.add_mul (exponential_exp a) (exponential_exp b)) -lemma log_mul_exp_add_of_lt {a b : M} (pos : 0 < a) (i : M) (hb : b < exp i) : log (a * exp i + b) = log a + i := by +lemma log_mul_exp_add_of_lt {a b : V} (pos : 0 < a) (i : V) (hb : b < exp i) : log (a * exp i + b) = log a + i := by simp [log_mul_pow2_add_of_lt pos (exp_pow2 i) hb] -lemma log_mul_exp {a : M} (pos : 0 < a) (i : M) : log (a * exp i) = log a + i := by +lemma log_mul_exp {a : V} (pos : 0 < a) (i : V) : log (a * exp i) = log a + i := by simp [log_mul_pow2 pos (exp_pow2 i)] -lemma length_mul_exp_add_of_lt {a b : M} (pos : 0 < a) (i : M) (hb : b < exp i) : ‖a * exp i + b‖ = ‖a‖ + i := by +lemma length_mul_exp_add_of_lt {a b : V} (pos : 0 < a) (i : V) (hb : b < exp i) : ‖a * exp i + b‖ = ‖a‖ + i := by simp [length_mul_pow2_add_of_lt pos (exp_pow2 i) hb] -lemma length_mul_exp {a : M} (pos : 0 < a) (i : M) : ‖a * exp i‖ = ‖a‖ + i := by +lemma length_mul_exp {a : V} (pos : 0 < a) (i : V) : ‖a * exp i‖ = ‖a‖ + i := by simp [length_mul_pow2 pos (exp_pow2 i)] -lemma exp_le_iff_le_log {i a : M} (pos : 0 < a) : exp i ≤ a ↔ i ≤ log a := +lemma exp_le_iff_le_log {i a : V} (pos : 0 < a) : exp i ≤ a ↔ i ≤ log a := ⟨by intro h; simpa using log_monotone h, fun h ↦ le_trans (exp_monotone_le.mpr h) (exp_log_le_self pos)⟩ end ISigma₁ diff --git a/Arithmetization/ISigmaZero/Exponential/PPow2.lean b/Arithmetization/ISigmaZero/Exponential/PPow2.lean index 7da1d6e..bbd5d47 100644 --- a/Arithmetization/ISigmaZero/Exponential/PPow2.lean +++ b/Arithmetization/ISigmaZero/Exponential/PPow2.lean @@ -4,15 +4,15 @@ noncomputable section namespace LO.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] open FirstOrder FirstOrder.Arith section ISigma₀ -variable [M ⊧ₘ* 𝐈𝚺₀] +variable [V ⊧ₘ* 𝐈𝚺₀] -def SPPow2 (m : M) : Prop := ¬LenBit 1 m ∧ LenBit 2 m ∧ ∀ i ≤ m, Pow2 i → 2 < i → (LenBit i m ↔ (√i)^2 = i ∧ LenBit (√i) m) +def SPPow2 (m : V) : Prop := ¬LenBit 1 m ∧ LenBit 2 m ∧ ∀ i ≤ m, Pow2 i → 2 < i → (LenBit i m ↔ (√i)^2 = i ∧ LenBit (√i) m) def _root_.LO.FirstOrder.Arith.sppow2Def : 𝚺₀.Semisentence 1 := .mkSigma @@ -20,7 +20,7 @@ def _root_.LO.FirstOrder.Arith.sppow2Def : 𝚺₀.Semisentence 1 := ∀ i <⁺ m, !pow2Def i → 2 < i → (!lenbitDef i m ↔ ∃ s <⁺ i, !sqrtDef s i ∧ s * s = i ∧ !lenbitDef s m) ” (by simp) -lemma sppow2_defined : 𝚺₀-Predicate (SPPow2 : M → Prop) via sppow2Def := by +lemma sppow2_defined : 𝚺₀-Predicate (SPPow2 : V → Prop) via sppow2Def := by intro v simp [SPPow2, sppow2Def, Matrix.vecHead, Matrix.vecTail, lenbit_defined.df.iff, pow2_defined.df.iff, sqrt_defined.df.iff, ←le_iff_lt_succ, sq, numeral_eq_natCast] @@ -30,44 +30,44 @@ lemma sppow2_defined : 𝚺₀-Predicate (SPPow2 : M → Prop) via sppow2Def := · intro h; exact ⟨√x, by simpa using h⟩ · rintro ⟨_, _, rfl, h⟩; exact h -def PPow2 (i : M) : Prop := Pow2 i ∧ ∃ m < 2 * i, SPPow2 m ∧ LenBit i m +def PPow2 (i : V) : Prop := Pow2 i ∧ ∃ m < 2 * i, SPPow2 m ∧ LenBit i m def _root_.LO.FirstOrder.Arith.ppow2Def : 𝚺₀.Semisentence 1 := .mkSigma “i | !pow2Def i ∧ ∃ m < 2 * i, !sppow2Def m ∧ !lenbitDef i m” (by simp) -lemma ppow2_defined : 𝚺₀-Predicate (PPow2 : M → Prop) via ppow2Def := by +lemma ppow2_defined : 𝚺₀-Predicate (PPow2 : V → Prop) via ppow2Def := by intro v; simp[PPow2, ppow2Def, Matrix.vecHead, Matrix.vecTail, lenbit_defined.df.iff, pow2_defined.df.iff, sppow2_defined.df.iff, numeral_eq_natCast] -instance ppow2_definable : DefinablePred ℒₒᵣ 𝚺₀ (PPow2 : M → Prop) := Defined.to_definable₀ _ ppow2_defined +instance ppow2_definable : 𝚺₀-Predicate (PPow2 : V → Prop) := ppow2_defined.to_definable namespace SPPow2 -variable {m : M} (hm : SPPow2 m) +variable {m : V} (hm : SPPow2 m) lemma not_lenbit_one : ¬LenBit 1 m := hm.1 lemma lenbit_two : LenBit 2 m := hm.2.1 -lemma lenbit_iff {i : M} (hi : i ≤ m) (pi : Pow2 i) (lt2 : 2 < i) : +lemma lenbit_iff {i : V} (hi : i ≤ m) (pi : Pow2 i) (lt2 : 2 < i) : LenBit i m ↔ (√i)^2 = i ∧ LenBit (√i) m := hm.2.2 i hi pi lt2 -lemma one_lt {i : M} (hi : LenBit i m) : 1 < i := by +lemma one_lt {i : V} (hi : LenBit i m) : 1 < i := by by_contra A rcases (le_one_iff_eq_zero_or_one.mp (show i ≤ 1 from by simpa using A)) with (rfl | rfl) · simp at hi · exact hm.1 hi -lemma two_lt {i : M} (hi : LenBit i m) (ne2 : i ≠ 2) : 2 < i := +lemma two_lt {i : V} (hi : LenBit i m) (ne2 : i ≠ 2) : 2 < i := lt_of_le_of_ne (one_lt_iff_two_le.mp $ hm.one_lt hi) (Ne.symm ne2) -lemma sqrt {i : M} (hi : LenBit i m) (pi : Pow2 i) (ne2 : i ≠ 2) : +lemma sqrt {i : V} (hi : LenBit i m) (pi : Pow2 i) (ne2 : i ≠ 2) : LenBit (√i) m := ((hm.lenbit_iff hi.le pi (hm.two_lt hi ne2)).mp hi).2 -lemma sq_sqrt_eq {i : M} (hi : LenBit i m) (pi : Pow2 i) (ne2 : i ≠ 2) : +lemma sq_sqrt_eq {i : V} (hi : LenBit i m) (pi : Pow2 i) (ne2 : i ≠ 2) : (√i)^2 = i := ((hm.lenbit_iff hi.le pi (hm.two_lt hi ne2)).mp hi).1 -lemma of_sqrt {i : M} (pi : Pow2 i) (him : i ≤ m) (hsqi : (√i)^2 = i) (hi : LenBit (√i) m) : +lemma of_sqrt {i : V} (pi : Pow2 i) (him : i ≤ m) (hsqi : (√i)^2 = i) (hi : LenBit (√i) m) : LenBit i m := by by_cases ne1 : i = 1 · rcases ne1; simpa using hi @@ -77,18 +77,18 @@ lemma of_sqrt {i : M} (pi : Pow2 i) (him : i ≤ m) (hsqi : (√i)^2 = i) (hi : (one_lt_iff_two_le.mp <| lt_of_le_of_ne (pos_iff_one_le.mp pi.pos) <| Ne.symm ne1) (Ne.symm ne2) exact (hm.lenbit_iff him pi this).mpr ⟨hsqi, hi⟩ -@[simp] lemma two : SPPow2 (2 : M) := +@[simp] lemma two : SPPow2 (2 : V) := ⟨by simp[LenBit.one], by simp, by intro i hi pi rcases le_two_iff_eq_zero_or_one_or_two.mp hi with (rfl | rfl | rfl) <;> simp⟩ -@[simp] lemma not_zero : ¬SPPow2 (0 : M) := by +@[simp] lemma not_zero : ¬SPPow2 (0 : V) := by rintro ⟨_, h, _⟩; simp at h -@[simp] lemma not_one : ¬SPPow2 (1 : M) := by +@[simp] lemma not_one : ¬SPPow2 (1 : V) := by rintro ⟨_, h, _⟩; simp [LenBit.iff_rem, one_lt_two] at h -lemma sq_le_of_lt {i j : M} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : LenBit j m) : i < j → i^2 ≤ j := by +lemma sq_le_of_lt {i j : V} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : LenBit j m) : i < j → i^2 ≤ j := by intro hij suffices ∀ i < j, Pow2 i → Pow2 j → LenBit i m → LenBit j m → i^2 ≤ j from this i hij pi pj hi hj clear i pi hi hij pj hj @@ -115,7 +115,7 @@ lemma sq_le_of_lt {i j : M} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : (pi.sqrt (hm.sq_sqrt_eq hi pi ine2)) (pj.sqrt (hm.sq_sqrt_eq hj pj jne2)) (hm.sqrt hi pi ine2) (hm.sqrt hj pj jne2) simpa [hm.sq_sqrt_eq hj pj jne2] using sq_le_sq.mpr this -lemma last_uniq {i j : M} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : LenBit j m) +lemma last_uniq {i j : V} (pi : Pow2 i) (pj : Pow2 j) (hi : LenBit i m) (hj : LenBit j m) (hsqi : m < i^2) (hsqj : m < j^2) : i = j := by by_contra ne wlog hij : i < j @@ -131,18 +131,18 @@ end SPPow2 namespace PPow2 -lemma pow2 {i : M} (h : PPow2 i) : Pow2 i := h.1 +lemma pow2 {i : V} (h : PPow2 i) : Pow2 i := h.1 -lemma pos {i : M} (ppi : PPow2 i) : 0 < i := ppi.pow2.pos +lemma pos {i : V} (ppi : PPow2 i) : 0 < i := ppi.pow2.pos -lemma one_lt {i : M} (ppi : PPow2 i) : 1 < i := by +lemma one_lt {i : V} (ppi : PPow2 i) : 1 < i := by rcases ppi with ⟨_, m, _, sppm, lb⟩; exact sppm.one_lt lb -lemma sq_sqrt_eq {i : M} (ppi : PPow2 i) (ne2 : i ≠ 2) : (√i)^2 = i := by +lemma sq_sqrt_eq {i : V} (ppi : PPow2 i) (ne2 : i ≠ 2) : (√i)^2 = i := by rcases ppi with ⟨pi, m, _, sppm, lb⟩ exact ((sppm.lenbit_iff lb.le pi (lt_of_le_of_ne (one_lt_iff_two_le.mp $ sppm.one_lt lb) (Ne.symm ne2))).mp lb).1 -lemma sqrt {i : M} (ppi : PPow2 i) (ne2 : i ≠ 2) : PPow2 (√i) := by +lemma sqrt {i : V} (ppi : PPow2 i) (ne2 : i ≠ 2) : PPow2 (√i) := by rcases ppi with ⟨pi, m, _, sppm, him⟩ have : LenBit i m ↔ (√i)^2 = i ∧ LenBit (√i) m := sppm.lenbit_iff him.le pi (lt_of_le_of_ne (one_lt_iff_two_le.mp $ sppm.one_lt him) (Ne.symm ne2)) @@ -162,9 +162,9 @@ lemma sqrt {i : M} (ppi : PPow2 i) (ne2 : i ≠ 2) : PPow2 (√i) := by · exact lt_of_le_of_lt (by simp) hjsi⟩ exact ⟨psqi, m % (2 * √i), mod_lt _ (by simp [psqi.pos]), this, by simp [H]⟩ -lemma exists_spp {i : M} (h : PPow2 i) : ∃ m < 2 * i, SPPow2 m ∧ LenBit i m := h.2 +lemma exists_spp {i : V} (h : PPow2 i) : ∃ m < 2 * i, SPPow2 m ∧ LenBit i m := h.2 -protected lemma sq {i : M} (ppi : PPow2 i) : PPow2 (i^2) := by +protected lemma sq {i : V} (ppi : PPow2 i) : PPow2 (i^2) := by rcases ppi.exists_spp with ⟨m, hm, sppm, hi⟩ have sppm' : SPPow2 (m + i^2) := ⟨by rw [LenBit.add_pow2] <;> try simp [ppi.pow2, sppm.not_lenbit_one, sppm.one_lt hi], @@ -202,16 +202,16 @@ protected lemma sq {i : M} (ppi : PPow2 i) : PPow2 (i^2) := by by simp [two_mul, hm, this], sppm', LenBit.add_self this⟩ -@[simp] lemma two : PPow2 (2 : M) := ⟨by simp, 2, by simp [one_lt_two]⟩ +@[simp] lemma two : PPow2 (2 : V) := ⟨by simp, 2, by simp [one_lt_two]⟩ -@[simp] lemma not_zero : ¬PPow2 (0 : M) := by intro h; simpa using h.pow2 +@[simp] lemma not_zero : ¬PPow2 (0 : V) := by intro h; simpa using h.pow2 -@[simp] lemma not_one : ¬PPow2 (1 : M) := by +@[simp] lemma not_one : ¬PPow2 (1 : V) := by rintro ⟨_, m, hm, H, _⟩ have : m ≤ 1 := lt_two_iff_le_one.mp (by simpa using hm) rcases le_one_iff_eq_zero_or_one.mp this with (rfl | rfl) <;> simp at H -lemma elim {i : M} : PPow2 i ↔ i = 2 ∨ ∃ b, i = b^2 ∧ PPow2 b := by +lemma elim {i : V} : PPow2 i ↔ i = 2 ∨ ∃ b, i = b^2 ∧ PPow2 b := by by_cases ei : i = 2 · rcases ei with rfl; simp · simp [ei]; constructor @@ -220,46 +220,46 @@ lemma elim {i : M} : PPow2 i ↔ i = 2 ∨ ∃ b, i = b^2 ∧ PPow2 b := by · rintro ⟨j, rfl, ppj⟩ exact ppj.sq -lemma elim' {i : M} : PPow2 i ↔ i = 2 ∨ 2 < i ∧ ∃ j, i = j^2 ∧ PPow2 j := by +lemma elim' {i : V} : PPow2 i ↔ i = 2 ∨ 2 < i ∧ ∃ j, i = j^2 ∧ PPow2 j := by by_cases ha : 2 < i <;> simp [ha, ←elim] have : i = 0 ∨ i = 1 ∨ i = 2 := by simpa [le_two_iff_eq_zero_or_one_or_two] using ha rcases this with (rfl | rfl | rfl) <;> simp -@[simp] lemma four : PPow2 (4 : M) := elim.mpr (Or.inr <| ⟨2, by simp [two_pow_two_eq_four]⟩) +@[simp] lemma four : PPow2 (4 : V) := elim.mpr (Or.inr <| ⟨2, by simp [two_pow_two_eq_four]⟩) -lemma two_le {i : M} (hi : PPow2 i) : 2 ≤ i := by +lemma two_le {i : V} (hi : PPow2 i) : 2 ≤ i := by simp [←one_add_one_eq_two, ←lt_iff_succ_le, hi.one_lt] -lemma not_three : ¬PPow2 (3 : M) := by +lemma not_three : ¬PPow2 (3 : V) := by intro h; simpa [sqrt_three] using h.sqrt (by simp) -lemma two_lt {i : M} (hi : PPow2 i) (ne : i ≠ 2) : 2 < i := by +lemma two_lt {i : V} (hi : PPow2 i) (ne : i ≠ 2) : 2 < i := by by_contra A; simp [ne, le_iff_lt_or_eq, lt_two_iff_le_one] at A rcases A with (rfl | rfl) <;> simp at hi -lemma four_le {i : M} (hi : PPow2 i) (ne : i ≠ 2) : 4 ≤ i := by +lemma four_le {i : V} (hi : PPow2 i) (ne : i ≠ 2) : 4 ≤ i := by by_contra A have : i ≤ 3 := by simpa [←three_add_one_eq_four, ←le_iff_lt_succ] using A rcases le_three_iff_eq_zero_or_one_or_two_or_three.mp this with (rfl | rfl | rfl | rfl) <;> simp at ne hi - · have : PPow2 (1 : M) := by simpa [sqrt_three] using hi.sqrt (by simp) + · have : PPow2 (1 : V) := by simpa [sqrt_three] using hi.sqrt (by simp) simp at this -lemma four_lt {i : M} (hi : PPow2 i) (ne2 : i ≠ 2) (ne4 : i ≠ 4) : 4 < i := +lemma four_lt {i : V} (hi : PPow2 i) (ne2 : i ≠ 2) (ne4 : i ≠ 4) : 4 < i := Ne.lt_of_le (Ne.symm ne4) (hi.four_le ne2) -lemma sq_ne_two {i : M} (hi : PPow2 i) : i^2 ≠ 2 := by +lemma sq_ne_two {i : V} (hi : PPow2 i) : i^2 ≠ 2 := by intro e; have : i < 2 := by simpa [←e] using lt_square_of_lt hi.one_lt exact not_le.mpr this hi.two_le -lemma sqrt_ne_two {i : M} (hi : PPow2 i) (ne2 : i ≠ 2) (ne4 : i ≠ 4) : √i ≠ 2 := by +lemma sqrt_ne_two {i : V} (hi : PPow2 i) (ne2 : i ≠ 2) (ne4 : i ≠ 4) : √i ≠ 2 := by intro e have : i = 4 := by simpa [e, two_pow_two_eq_four] using Eq.symm <| hi.sq_sqrt_eq ne2 contradiction -lemma sq_ne_four {i : M} (hi : PPow2 i) (ne2 : i ≠ 2) : i^2 ≠ 4 := by +lemma sq_ne_four {i : V} (hi : PPow2 i) (ne2 : i ≠ 2) : i^2 ≠ 4 := by simpa [two_pow_two_eq_four] using ne_of_gt (sq_lt_sq.mpr (hi.two_lt ne2)) -lemma sq_le_of_lt {i j : M} (hi : PPow2 i) (hj : PPow2 j) : i < j → i^2 ≤ j := by +lemma sq_le_of_lt {i j : V} (hi : PPow2 i) (hj : PPow2 j) : i < j → i^2 ≤ j := by intro hij suffices ∀ i < j, PPow2 i → PPow2 j → i^2 ≤ j from this i hij hi hj clear hi hij hj @@ -282,7 +282,7 @@ lemma sq_le_of_lt {i j : M} (hi : PPow2 i) (hj : PPow2 j) : i < j → i^2 ≤ j IH (√j) (sqrt_lt_self_of_one_lt hj.one_lt) (√i) this (hi.sqrt ei) (hj.sqrt ej) simpa [hj.sq_sqrt_eq ej] using sq_le_sq.mpr this -lemma sq_uniq {y i j : M} (py : Pow2 y) (ppi : PPow2 i) (ppj : PPow2 j) +lemma sq_uniq {y i j : V} (py : Pow2 y) (ppi : PPow2 i) (ppj : PPow2 j) (hi : y < i ∧ i ≤ y^2) (hj : y < j ∧ j ≤ y^2) : i = j := by by_contra ne wlog hij : i < j @@ -293,7 +293,7 @@ lemma sq_uniq {y i j : M} (py : Pow2 y) (ppi : PPow2 i) (ppj : PPow2 j) _ ≤ y^2 := hj.2 simp_all -lemma two_mul_sq_uniq {y i j : M} (py : Pow2 y) (ppi : PPow2 i) (ppj : PPow2 j) +lemma two_mul_sq_uniq {y i j : V} (py : Pow2 y) (ppi : PPow2 i) (ppj : PPow2 j) (hi : y < i ∧ i ≤ 2 * y^2) (hj : y < j ∧ j ≤ 2 * y^2) : i = j := by by_contra ne wlog hij : i < j diff --git a/Arithmetization/ISigmaZero/Exponential/Pow2.lean b/Arithmetization/ISigmaZero/Exponential/Pow2.lean index 25c4f38..81cadb7 100644 --- a/Arithmetization/ISigmaZero/Exponential/Pow2.lean +++ b/Arithmetization/ISigmaZero/Exponential/Pow2.lean @@ -4,48 +4,48 @@ noncomputable section namespace LO.Arith -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] open FirstOrder FirstOrder.Arith section IOpen -variable [M ⊧ₘ* 𝐈open] +variable [V ⊧ₘ* 𝐈open] -def Pow2 (a : M) : Prop := 0 < a ∧ ∀ r ≤ a, 1 < r → r ∣ a → 2 ∣ r +def Pow2 (a : V) : Prop := 0 < a ∧ ∀ r ≤ a, 1 < r → r ∣ a → 2 ∣ r def _root_.LO.FirstOrder.Arith.pow2Def : 𝚺₀.Semisentence 1 := .mkSigma “a | 0 < a ∧ ∀ r <⁺ a, 1 < r → r ∣ a → 2 ∣ r” (by simp [Hierarchy.pi_zero_iff_sigma_zero]) -lemma pow2_defined : 𝚺₀-Predicate (Pow2 : M → Prop) via pow2Def := by +lemma pow2_defined : 𝚺₀-Predicate (Pow2 : V → Prop) via pow2Def := by intro v simp [Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.vecHead, Matrix.constant_eq_singleton, Pow2, pow2Def, le_iff_lt_succ, dvd_defined.df.iff, numeral_eq_natCast] -instance pow2_definable : DefinablePred ℒₒᵣ 𝚺₀ (Pow2 : M → Prop) := Defined.to_definable _ pow2_defined +instance pow2_definable : 𝚺₀-Predicate (Pow2 : V → Prop) := pow2_defined.to_definable -lemma Pow2.pos {a : M} (h : Pow2 a) : 0 < a := h.1 +lemma Pow2.pos {a : V} (h : Pow2 a) : 0 < a := h.1 -lemma Pow2.dvd {a : M} (h : Pow2 a) {r} (hr : r ≤ a) : 1 < r → r ∣ a → 2 ∣ r := h.2 r hr +lemma Pow2.dvd {a : V} (h : Pow2 a) {r} (hr : r ≤ a) : 1 < r → r ∣ a → 2 ∣ r := h.2 r hr -@[simp] lemma pow2_one : Pow2 (1 : M) := ⟨by simp, by +@[simp] lemma pow2_one : Pow2 (1 : V) := ⟨by simp, by intro r hr hhr hd rcases show r = 0 ∨ r = 1 from le_one_iff_eq_zero_or_one.mp hr with (rfl | rfl) · simp · simp at hhr⟩ -@[simp] lemma not_pow2_zero : ¬Pow2 (0 : M) := by +@[simp] lemma not_pow2_zero : ¬Pow2 (0 : V) := by intro h; have := h.pos; simp at this -lemma Pow2.two_dvd {a : M} (h : Pow2 a) (lt : 1 < a) : 2 ∣ a := h.dvd (le_refl _) lt (by simp) +lemma Pow2.two_dvd {a : V} (h : Pow2 a) (lt : 1 < a) : 2 ∣ a := h.dvd (le_refl _) lt (by simp) -lemma Pow2.two_dvd' {a : M} (h : Pow2 a) (lt : a ≠ 1) : 2 ∣ a := +lemma Pow2.two_dvd' {a : V} (h : Pow2 a) (lt : a ≠ 1) : 2 ∣ a := h.dvd (le_refl _) (by by_contra A; simp [le_one_iff_eq_zero_or_one] at A rcases A with (rfl | rfl) <;> simp at h lt) (by simp) -lemma Pow2.of_dvd {a b : M} (h : b ∣ a) : Pow2 a → Pow2 b := by +lemma Pow2.of_dvd {a b : V} (h : b ∣ a) : Pow2 a → Pow2 b := by intro ha have : 0 < b := by by_contra e @@ -55,7 +55,7 @@ lemma Pow2.of_dvd {a b : M} (h : b ∣ a) : Pow2 a → Pow2 b := by exact ⟨this, fun r hr ltr hb ↦ ha.dvd (show r ≤ a from le_trans hr (le_of_dvd ha.pos h)) ltr (dvd_trans hb h)⟩ -lemma pow2_mul_two {a : M} : Pow2 (2 * a) ↔ Pow2 a := +lemma pow2_mul_two {a : V} : Pow2 (2 * a) ↔ Pow2 a := ⟨by intro H have : ∀ r ≤ a, 1 < r → r ∣ a → 2 ∣ r := by intro r hr ltr dvd @@ -68,10 +68,10 @@ lemma pow2_mul_two {a : M} : Pow2 (2 * a) ↔ Pow2 a := · exact hd · exact H.dvd (show r ≤ a from le_of_dvd H.pos hd) hr hd⟩⟩ -lemma pow2_mul_four {a : M} : Pow2 (4 * a) ↔ Pow2 a := by +lemma pow2_mul_four {a : V} : Pow2 (4 * a) ↔ Pow2 a := by simp [←two_mul_two_eq_four, mul_assoc, pow2_mul_two] -lemma Pow2.elim {p : M} : Pow2 p ↔ p = 1 ∨ ∃ q, p = 2 * q ∧ Pow2 q := +lemma Pow2.elim {p : V} : Pow2 p ↔ p = 1 ∨ ∃ q, p = 2 * q ∧ Pow2 q := ⟨by intro H by_cases hp : 1 < p · have : 2 ∣ p := H.two_dvd hp @@ -81,22 +81,22 @@ lemma Pow2.elim {p : M} : Pow2 p ↔ p = 1 ∨ ∃ q, p = 2 * q ∧ Pow2 q := left; exact this, by rintro (rfl | ⟨q, rfl, hq⟩) <;> simp [pow2_one, pow2_mul_two, *]⟩ -@[simp] lemma pow2_two : Pow2 (2 : M) := Pow2.elim.mpr (Or.inr ⟨1, by simp⟩) +@[simp] lemma pow2_two : Pow2 (2 : V) := Pow2.elim.mpr (Or.inr ⟨1, by simp⟩) -lemma Pow2.div_two {p : M} (h : Pow2 p) (ne : p ≠ 1) : Pow2 (p / 2) := by +lemma Pow2.div_two {p : V} (h : Pow2 p) (ne : p ≠ 1) : Pow2 (p / 2) := by rcases Pow2.elim.mp h with (rfl | ⟨q, rfl, pq⟩) · simp at ne simpa -lemma Pow2.two_mul_div_two {p : M} (h : Pow2 p) (ne : p ≠ 1) : 2 * (p / 2) = p := by +lemma Pow2.two_mul_div_two {p : V} (h : Pow2 p) (ne : p ≠ 1) : 2 * (p / 2) = p := by rcases Pow2.elim.mp h with (rfl | ⟨q, rfl, _⟩) · simp at ne simp -lemma Pow2.div_two_mul_two {p : M} (h : Pow2 p) (ne : p ≠ 1) : (p / 2) * 2 = p := by +lemma Pow2.div_two_mul_two {p : V} (h : Pow2 p) (ne : p ≠ 1) : (p / 2) * 2 = p := by simp [mul_comm, h.two_mul_div_two ne] -lemma Pow2.elim' {p : M} : Pow2 p ↔ p = 1 ∨ 1 < p ∧ ∃ q, p = 2 * q ∧ Pow2 q := by +lemma Pow2.elim' {p : V} : Pow2 p ↔ p = 1 ∨ 1 < p ∧ ∃ q, p = 2 * q ∧ Pow2 q := by by_cases hp : 1 < p <;> simp [hp] · exact Pow2.elim · have : p = 0 ∨ p = 1 := le_one_iff_eq_zero_or_one.mp (show p ≤ 1 from by simpa using hp) @@ -106,43 +106,43 @@ lemma Pow2.elim' {p : M} : Pow2 p ↔ p = 1 ∨ 1 < p ∧ ∃ q, p = 2 * q ∧ P section LenBit /-- $\mathrm{LenBit} (2^i, a) \iff \text{$i$th-bit of $a$ is $1$}$. -/ -def LenBit (i a : M) : Prop := ¬2 ∣ (a / i) +def LenBit (i a : V) : Prop := ¬2 ∣ (a / i) def _root_.LO.FirstOrder.Arith.lenbitDef : 𝚺₀.Semisentence 2 := .mkSigma “i a | ∃ z <⁺ a, !divDef.val z a i ∧ ¬2 ∣ z” (by simp) -lemma lenbit_defined : 𝚺₀-Relation (LenBit : M → M → Prop) via lenbitDef := by +lemma lenbit_defined : 𝚺₀-Relation (LenBit : V → V → Prop) via lenbitDef := by intro v; simp[sqrt_graph, lenbitDef, Matrix.vecHead, Matrix.vecTail, LenBit, ←le_iff_lt_succ, numeral_eq_natCast] constructor · intro h; exact ⟨v 1 / v 0, by simp, rfl, h⟩ · rintro ⟨z, hz, rfl, h⟩; exact h @[simp] lemma lenbit_defined_iff (v) : - Semiformula.Evalbm M v lenbitDef.val ↔ LenBit (v 0) (v 1) := lenbit_defined.df.iff v + Semiformula.Evalbm V v lenbitDef.val ↔ LenBit (v 0) (v 1) := lenbit_defined.df.iff v -instance lenbit_definable : DefinableRel ℒₒᵣ 𝚺₀ (LenBit : M → M → Prop) := Defined.to_definable _ lenbit_defined +instance lenbit_definable : 𝚺₀-Relation (LenBit : V → V → Prop) := lenbit_defined.to_definable -lemma LenBit.le {i a : M} (h : LenBit i a) : i ≤ a := by +lemma LenBit.le {i a : V} (h : LenBit i a) : i ≤ a := by by_contra A; simp [LenBit, show a < i from by simpa using A] at h -lemma not_lenbit_of_lt {i a : M} (h : a < i) : ¬LenBit i a := by +lemma not_lenbit_of_lt {i a : V} (h : a < i) : ¬LenBit i a := by intro A; exact not_le.mpr h A.le -@[simp] lemma LenBit.zero (a : M) : ¬LenBit 0 a := by simp [LenBit] +@[simp] lemma LenBit.zero (a : V) : ¬LenBit 0 a := by simp [LenBit] -@[simp] lemma LenBit.on_zero (a : M) : ¬LenBit a 0 := by simp [LenBit] +@[simp] lemma LenBit.on_zero (a : V) : ¬LenBit a 0 := by simp [LenBit] -lemma LenBit.one (a : M) : LenBit 1 a ↔ ¬2 ∣ a := by simp [LenBit] +lemma LenBit.one (a : V) : LenBit 1 a ↔ ¬2 ∣ a := by simp [LenBit] -lemma LenBit.iff_rem {i a : M} : LenBit i a ↔ (a / i) % 2 = 1 := by +lemma LenBit.iff_rem {i a : V} : LenBit i a ↔ (a / i) % 2 = 1 := by simp [LenBit]; rcases mod_two (a / i) with (h | h) <;> simp [h, ←mod_eq_zero_iff_dvd] -lemma not_lenbit_iff_rem {i a : M} : ¬LenBit i a ↔ (a / i) % 2 = 0 := by +lemma not_lenbit_iff_rem {i a : V} : ¬LenBit i a ↔ (a / i) % 2 = 0 := by simp [LenBit, ←mod_eq_zero_iff_dvd] -@[simp] lemma LenBit.self {a : M} (pos : 0 < a) : LenBit a a := by simp [LenBit.iff_rem, pos, one_lt_two] +@[simp] lemma LenBit.self {a : V} (pos : 0 < a) : LenBit a a := by simp [LenBit.iff_rem, pos, one_lt_two] -lemma LenBit.mod {i a k : M} (h : 2 * i ∣ k) : LenBit i (a % k) ↔ LenBit i a := by +lemma LenBit.mod {i a k : V} (h : 2 * i ∣ k) : LenBit i (a % k) ↔ LenBit i a := by have : 0 ≤ i := zero_le i rcases (eq_or_lt_of_le this) with (rfl | pos) · simp @@ -156,9 +156,9 @@ lemma LenBit.mod {i a k : M} (h : 2 * i ∣ k) : LenBit i (a % k) ↔ LenBit i a simp [mul_right_comm _ (a / k), mul_right_comm 2 k' i, ←hk']) _ ↔ LenBit i a := by simp [div_add_mod a k, LenBit.iff_rem] -@[simp] lemma LenBit.mod_two_mul_self {a i : M} : LenBit i (a % (2 * i)) ↔ LenBit i a := LenBit.mod (by simp) +@[simp] lemma LenBit.mod_two_mul_self {a i : V} : LenBit i (a % (2 * i)) ↔ LenBit i a := LenBit.mod (by simp) -lemma LenBit.add {i a b : M} (h : 2 * i ∣ b) : LenBit i (a + b) ↔ LenBit i a := by +lemma LenBit.add {i a b : V} (h : 2 * i ∣ b) : LenBit i (a + b) ↔ LenBit i a := by have : 0 ≤ i := zero_le i rcases (eq_or_lt_of_le this) with (rfl | pos) · simp @@ -169,21 +169,21 @@ lemma LenBit.add {i a b : M} (h : 2 * i ∣ b) : LenBit i (a + b) ↔ LenBit i a _ ↔ (a / i + 2 * b') % 2 = 1 := by rw [hb', div_add_mul_self _ _ pos] _ ↔ LenBit i a := by simp [LenBit.iff_rem] -lemma LenBit.add_self {i a : M} (h : a < i) : LenBit i (a + i) := by +lemma LenBit.add_self {i a : V} (h : a < i) : LenBit i (a + i) := by have pos : 0 < i := by exact pos_of_gt h simp [LenBit.iff_rem, div_add_self_right _ pos, h, one_lt_two] -lemma LenBit.add_self_of_not_lenbit {a i : M} (pos : 0 < i) (h : ¬LenBit i a) : LenBit i (a + i) := by +lemma LenBit.add_self_of_not_lenbit {a i : V} (pos : 0 < i) (h : ¬LenBit i a) : LenBit i (a + i) := by have : a / i % 2 = 0 := by simpa [LenBit.iff_rem] using h simp [LenBit.iff_rem, div_add_self_right, pos] rw [mod_add] <;> simp [this, one_lt_two] -lemma LenBit.add_self_of_lenbit {a i : M} (pos : 0 < i) (h : LenBit i a) : ¬LenBit i (a + i) := by +lemma LenBit.add_self_of_lenbit {a i : V} (pos : 0 < i) (h : LenBit i a) : ¬LenBit i (a + i) := by have : a / i % 2 = 1 := by simpa [LenBit.iff_rem] using h simp [LenBit.iff_rem, div_add_self_right, pos] rw [mod_add] <;> simp [this, one_add_one_eq_two] -lemma LenBit.sub_self_of_lenbit {a i : M} (pos : 0 < i) (h : LenBit i a) : ¬LenBit i (a - i) := by +lemma LenBit.sub_self_of_lenbit {a i : V} (pos : 0 < i) (h : LenBit i a) : ¬LenBit i (a - i) := by intro h' have : ¬LenBit i a := by simpa [sub_add_self_of_le h.le] using LenBit.add_self_of_lenbit pos h' contradiction @@ -194,14 +194,14 @@ end IOpen section ISigma₀ -variable [M ⊧ₘ* 𝐈𝚺₀] +variable [V ⊧ₘ* 𝐈𝚺₀] namespace Pow2 -lemma mul {a b : M} (ha : Pow2 a) (hb : Pow2 b) : Pow2 (a * b) := by +lemma mul {a b : V} (ha : Pow2 a) (hb : Pow2 b) : Pow2 (a * b) := by wlog hab : a ≤ b · simpa [mul_comm] using this hb ha (by simp at hab; exact LT.lt.le hab) - suffices ∀ b : M, ∀ a ≤ b, Pow2 a → Pow2 b → Pow2 (a * b) by + suffices ∀ b : V, ∀ a ≤ b, Pow2 a → Pow2 b → Pow2 (a * b) by exact this b a hab ha hb intro b induction b using order_induction_sigma0 @@ -222,17 +222,17 @@ lemma mul {a b : M} (ha : Pow2 a) (hb : Pow2 b) : Pow2 (a * b) := by simpa [this] simpa [mul_assoc, pow2_mul_four] using this -@[simp] lemma mul_iff {a b : M} : Pow2 (a * b) ↔ Pow2 a ∧ Pow2 b := +@[simp] lemma mul_iff {a b : V} : Pow2 (a * b) ↔ Pow2 a ∧ Pow2 b := ⟨fun h ↦ ⟨h.of_dvd (by simp), h.of_dvd (by simp)⟩, by rintro ⟨ha, hb⟩; exact ha.mul hb⟩ -@[simp] lemma sq_iff {a : M} : Pow2 (a^2) ↔ Pow2 a := by +@[simp] lemma sq_iff {a : V} : Pow2 (a^2) ↔ Pow2 a := by simp [_root_.sq] -lemma sq {a : M} : Pow2 a → Pow2 (a^2) := by +lemma sq {a : V} : Pow2 a → Pow2 (a^2) := by simp [_root_.sq] -lemma dvd_of_le {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b → a ∣ b := by - suffices ∀ b : M, ∀ a ≤ b, Pow2 a → Pow2 b → a ∣ b by +lemma dvd_of_le {a b : V} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b → a ∣ b := by + suffices ∀ b : V, ∀ a ≤ b, Pow2 a → Pow2 b → a ∣ b by intro hab; exact this b a hab ha hb intro b; induction b using order_induction_sigma0 · definability @@ -249,13 +249,13 @@ lemma dvd_of_le {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b → a ∣ b := b have hab : a ≤ b := le_of_mul_le_mul_left hab (by simp) exact mul_dvd_mul_left 2 <| IH b ltb a hab (by assumption) (by assumption) -lemma le_iff_dvd {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b ↔ a ∣ b := +lemma le_iff_dvd {a b : V} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b ↔ a ∣ b := ⟨Pow2.dvd_of_le ha hb, le_of_dvd hb.pos⟩ -lemma two_le {a : M} (pa : Pow2 a) (ne1 : a ≠ 1) : 2 ≤ a := +lemma two_le {a : V} (pa : Pow2 a) (ne1 : a ≠ 1) : 2 ≤ a := le_of_dvd pa.pos (pa.two_dvd' ne1) -lemma le_iff_lt_two {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b ↔ a < 2 * b := by +lemma le_iff_lt_two {a b : V} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b ↔ a < 2 * b := by constructor · intro h; exact lt_of_le_of_lt h (lt_two_mul_self hb.pos) · intro h @@ -273,12 +273,12 @@ lemma le_iff_lt_two {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a ≤ b ↔ a < 2 * rcases this with ⟨b'', rfl⟩ simp [←mul_assoc, ha.div_two_mul_two ea] -lemma lt_iff_two_mul_le {a b : M} (ha : Pow2 a) (hb : Pow2 b) : a < b ↔ 2 * a ≤ b := by +lemma lt_iff_two_mul_le {a b : V} (ha : Pow2 a) (hb : Pow2 b) : a < b ↔ 2 * a ≤ b := by by_cases eb : b = 1 · simp [eb, ←lt_two_iff_le_one] · rw [←hb.two_mul_div_two eb]; simp [le_iff_lt_two ha (hb.div_two eb)] -lemma sq_or_dsq {a : M} (pa : Pow2 a) : ∃ b, a = b^2 ∨ a = 2 * b^2 := by +lemma sq_or_dsq {a : V} (pa : Pow2 a) : ∃ b, a = b^2 ∨ a = 2 * b^2 := by suffices ∃ b ≤ a, a = b^2 ∨ a = 2 * b^2 by rcases this with ⟨b, _, h⟩ exact ⟨b, h⟩ @@ -293,20 +293,20 @@ lemma sq_or_dsq {a : M} (pa : Pow2 a) : ∃ b, a = b^2 ∨ a = 2 * b^2 := by · exact ⟨2 * b, by simp; exact le_trans (by simp) le_two_mul_left, by left; simp [_root_.sq, mul_assoc, mul_left_comm]⟩ -lemma sqrt {a : M} (h : Pow2 a) (hsq : (√a)^2 = a) : Pow2 (√a) := by +lemma sqrt {a : V} (h : Pow2 a) (hsq : (√a)^2 = a) : Pow2 (√a) := by rw [←hsq] at h; simpa using h -@[simp] lemma not_three : ¬Pow2 (3 : M) := by +@[simp] lemma not_three : ¬Pow2 (3 : V) := by intro h - have : (2 : M) ∣ 3 := h.two_dvd (by simp [←two_add_one_eq_three]) + have : (2 : V) ∣ 3 := h.two_dvd (by simp [←two_add_one_eq_three]) simp [←two_add_one_eq_three, ←mod_eq_zero_iff_dvd, one_lt_two] at this -lemma four_le {i : M} (hi : Pow2 i) (lt : 2 < i) : 4 ≤ i := by +lemma four_le {i : V} (hi : Pow2 i) (lt : 2 < i) : 4 ≤ i := by by_contra A have : i ≤ 3 := by simpa [←three_add_one_eq_four, ←le_iff_lt_succ] using A rcases le_three_iff_eq_zero_or_one_or_two_or_three.mp this with (rfl | rfl | rfl | rfl) <;> simp at lt hi -lemma mul_add_lt_of_mul_lt_of_pos {a b p q : M} (hp : Pow2 p) (hq : Pow2 q) +lemma mul_add_lt_of_mul_lt_of_pos {a b p q : V} (hp : Pow2 p) (hq : Pow2 q) (h : a * p < q) (hb : b < p) (hbq : b < q) : a * p + b < q := by rcases zero_le a with (rfl | pos) · simp [hbq] @@ -321,13 +321,13 @@ lemma mul_add_lt_of_mul_lt_of_pos {a b p q : M} (hp : Pow2 p) (hq : Pow2 q) end Pow2 -lemma LenBit.mod_pow2 {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : i < j) : LenBit i (a % j) ↔ LenBit i a := +lemma LenBit.mod_pow2 {a i j : V} (pi : Pow2 i) (pj : Pow2 j) (h : i < j) : LenBit i (a % j) ↔ LenBit i a := LenBit.mod (by rw [←Pow2.le_iff_dvd] <;> simp [pi, pj, ←Pow2.lt_iff_two_mul_le, h]) -lemma LenBit.add_pow2 {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : i < j) : LenBit i (a + j) ↔ LenBit i a := +lemma LenBit.add_pow2 {a i j : V} (pi : Pow2 i) (pj : Pow2 j) (h : i < j) : LenBit i (a + j) ↔ LenBit i a := LenBit.add (by rw [←Pow2.le_iff_dvd] <;> simp [pi, pj, ←Pow2.lt_iff_two_mul_le, h]) -lemma LenBit.add_pow2_iff_of_lt {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : a < j) : LenBit i (a + j) ↔ i = j ∨ LenBit i a := by +lemma LenBit.add_pow2_iff_of_lt {a i j : V} (pi : Pow2 i) (pj : Pow2 j) (h : a < j) : LenBit i (a + j) ↔ i = j ∨ LenBit i a := by rcases show i < j ∨ i = j ∨ i > j from lt_trichotomy i j with (hij | rfl | hij) · simp [LenBit.add_pow2 pi pj hij, hij.ne] · simp [LenBit.add_self h] @@ -336,7 +336,7 @@ lemma LenBit.add_pow2_iff_of_lt {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : a < _ ≤ i := (pj.lt_iff_two_mul_le pi).mp hij simp [not_lenbit_of_lt this, not_lenbit_of_lt (show a < i from lt_trans h hij), hij.ne.symm] -lemma lenbit_iff_add_mul {i a : M} (hi : Pow2 i) : +lemma lenbit_iff_add_mul {i a : V} (hi : Pow2 i) : LenBit i a ↔ ∃ k, ∃ r < i, a = k * (2 * i) + i + r := by constructor · intro h @@ -349,7 +349,7 @@ lemma lenbit_iff_add_mul {i a : M} (hi : Pow2 i) : · rintro ⟨k, r, h, rfl⟩ simp [LenBit.iff_rem, ←mul_assoc, add_assoc, div_mul_add_self, hi.pos, h, one_lt_two] -lemma not_lenbit_iff_add_mul {i a : M} (hi : Pow2 i) : +lemma not_lenbit_iff_add_mul {i a : V} (hi : Pow2 i) : ¬LenBit i a ↔ ∃ k, ∃ r < i, a = k * (2 * i) + r := by constructor · intro h @@ -362,7 +362,7 @@ lemma not_lenbit_iff_add_mul {i a : M} (hi : Pow2 i) : · rintro ⟨k, r, h, rfl⟩ simp [not_lenbit_iff_rem, ←mul_assoc, add_assoc, div_mul_add_self, hi.pos, h] -lemma lenbit_mul_add {i j a r : M} (pi : Pow2 i) (pj : Pow2 j) (hr : r < j) : +lemma lenbit_mul_add {i j a r : V} (pi : Pow2 i) (pj : Pow2 j) (hr : r < j) : LenBit (i * j) (a * j + r) ↔ LenBit i a := by by_cases h : LenBit i a <;> simp [h] · rcases (lenbit_iff_add_mul pi).mp h with ⟨a, b, hb, rfl⟩ @@ -374,7 +374,7 @@ lemma lenbit_mul_add {i j a r : M} (pi : Pow2 i) (pj : Pow2 j) (hr : r < j) : pj.mul_add_lt_of_mul_lt_of_pos (by simp[pi, pj]) ((mul_lt_mul_right pj.pos).mpr hb) hr (lt_of_lt_of_le hr $ le_mul_of_pos_left $ pi.pos) exact (not_lenbit_iff_add_mul (by simp [pi, pj])).mpr ⟨a, b * j + r, this, by simp [add_mul, add_assoc, mul_assoc]⟩ -lemma lenbit_add_pow2_iff_of_not_lenbit {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : ¬LenBit j a) : +lemma lenbit_add_pow2_iff_of_not_lenbit {a i j : V} (pi : Pow2 i) (pj : Pow2 j) (h : ¬LenBit j a) : LenBit i (a + j) ↔ i = j ∨ LenBit i a := by rcases show i < j ∨ i = j ∨ i > j from lt_trichotomy i j with (hij | rfl | hij) · simp [LenBit.add_pow2 pi pj hij, hij.ne] @@ -396,7 +396,7 @@ lemma lenbit_add_pow2_iff_of_not_lenbit {a i j : M} (pi : Pow2 i) (pj : Pow2 j) _ ↔ LenBit (2 * j * i) (a * (2 * j) + r) := by simp [mul_comm] -lemma lenbit_sub_pow2_iff_of_lenbit {a i j : M} (pi : Pow2 i) (pj : Pow2 j) (h : LenBit j a) : +lemma lenbit_sub_pow2_iff_of_lenbit {a i j : V} (pi : Pow2 i) (pj : Pow2 j) (h : LenBit j a) : LenBit i (a - j) ↔ i ≠ j ∧ LenBit i a := by generalize ha' : a - j = a' have h' : ¬LenBit j a' := by simpa [←ha'] using LenBit.sub_self_of_lenbit pj.pos h diff --git a/Arithmetization/OmegaOne/Basic.lean b/Arithmetization/OmegaOne/Basic.lean index 0f052dc..86e7bba 100644 --- a/Arithmetization/OmegaOne/Basic.lean +++ b/Arithmetization/OmegaOne/Basic.lean @@ -17,9 +17,9 @@ notation "𝛀₁" => Theory.omegaOne noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] +variable {V : Type*} [ORingStruc V] -lemma models_Omega₁_iff [M ⊧ₘ* 𝐈𝚺₀] : M ⊧ₘ omegaOneAxiom ↔ ∀ x : M, ∃ y, Exponential (‖x‖^2) y := by +lemma models_Omega₁_iff [V ⊧ₘ* 𝐈𝚺₀] : V ⊧ₘ omegaOneAxiom ↔ ∀ x : V, ∃ y, Exponential (‖x‖^2) y := by simp [models_def, omegaOneAxiom, length_defined.df.iff, Exponential.defined.df.iff, sq, ←le_iff_lt_succ] constructor · intro h x @@ -28,25 +28,25 @@ lemma models_Omega₁_iff [M ⊧ₘ* 𝐈𝚺₀] : M ⊧ₘ omegaOneAxiom ↔ rcases h x with ⟨y, h⟩ exact ⟨y, ‖x‖, by simp, rfl, h⟩ -lemma sigma₁_omega₁ [M ⊧ₘ* 𝐈𝚺₁] : M ⊧ₘ omegaOneAxiom := models_Omega₁_iff.mpr (fun x ↦ Exponential.range_exists (‖x‖^2)) +lemma sigma₁_omega₁ [V ⊧ₘ* 𝐈𝚺₁] : V ⊧ₘ omegaOneAxiom := models_Omega₁_iff.mpr (fun x ↦ Exponential.range_exists (‖x‖^2)) -instance [M ⊧ₘ* 𝐈𝚺₁] : M ⊧ₘ* 𝐈𝚺₀ + 𝛀₁ := +instance [V ⊧ₘ* 𝐈𝚺₁] : V ⊧ₘ* 𝐈𝚺₀ + 𝛀₁ := ModelsTheory.add_iff.mpr ⟨inferInstance, ⟨by intro _; simp; rintro rfl; exact sigma₁_omega₁⟩⟩ -variable [M ⊧ₘ* 𝐈𝚺₀ + 𝛀₁] +variable [V ⊧ₘ* 𝐈𝚺₀ + 𝛀₁] -instance : M ⊧ₘ* 𝐈𝚺₀ := ModelsTheory.of_add_left M 𝐈𝚺₀ 𝛀₁ +instance : V ⊧ₘ* 𝐈𝚺₀ := ModelsTheory.of_add_left V 𝐈𝚺₀ 𝛀₁ -instance : M ⊧ₘ* 𝛀₁ := ModelsTheory.of_add_right M 𝐈𝚺₀ 𝛀₁ +instance : V ⊧ₘ* 𝛀₁ := ModelsTheory.of_add_right V 𝐈𝚺₀ 𝛀₁ -lemma exists_exponential_sq_length (x : M) : ∃ y, Exponential (‖x‖^2) y := - models_Omega₁_iff.mp (ModelsTheory.models M Theory.omegaOne.omega) x +lemma exists_exponential_sq_length (x : V) : ∃ y, Exponential (‖x‖^2) y := + models_Omega₁_iff.mp (ModelsTheory.models V Theory.omegaOne.omega) x -lemma exists_unique_exponential_sq_length (x : M) : ∃! y, Exponential (‖x‖^2) y := by +lemma exists_unique_exponential_sq_length (x : V) : ∃! y, Exponential (‖x‖^2) y := by rcases exists_exponential_sq_length x with ⟨y, h⟩ exact ExistsUnique.intro y h (fun y' h' ↦ h'.uniq h) -lemma hash_exists_unique (x y : M) : ∃! z, Exponential (‖x‖ * ‖y‖) z := by +lemma hash_exists_unique (x y : V) : ∃! z, Exponential (‖x‖ * ‖y‖) z := by wlog le : x ≤ y · simpa [mul_comm] using this y x (le_of_not_ge le) rcases exists_exponential_sq_length y with ⟨z, h⟩ @@ -55,42 +55,42 @@ lemma hash_exists_unique (x y : M) : ∃! z, Exponential (‖x‖ * ‖y‖) z : have : Exponential (‖x‖ * ‖y‖) (bexp z (‖x‖ * ‖y‖)) := exp_bexp_of_lt (a := z) (x := ‖x‖ * ‖y‖) this exact ExistsUnique.intro (bexp z (‖x‖ * ‖y‖)) this (fun z' H' ↦ H'.uniq this) -instance : Hash M := ⟨fun a b ↦ Classical.choose! (hash_exists_unique a b)⟩ +instance : Hash V := ⟨fun a b ↦ Classical.choose! (hash_exists_unique a b)⟩ -lemma exponential_hash (a b : M) : Exponential (‖a‖ * ‖b‖) (a # b) := Classical.choose!_spec (hash_exists_unique a b) +lemma exponential_hash (a b : V) : Exponential (‖a‖ * ‖b‖) (a # b) := Classical.choose!_spec (hash_exists_unique a b) -lemma exponential_hash_one (a : M) : Exponential ‖a‖ (a # 1) := by simpa using exponential_hash a 1 +lemma exponential_hash_one (a : V) : Exponential ‖a‖ (a # 1) := by simpa using exponential_hash a 1 def hashDef : 𝚺₀.Semisentence 3 := .mkSigma “z x y | ∃ lx <⁺ x, ∃ ly <⁺ y, !lengthDef lx x ∧ !lengthDef ly y ∧ !exponentialDef (lx * ly) z” (by simp) -lemma hash_defined : 𝚺₀-Function₂ (Hash.hash : M → M → M) via hashDef := by +lemma hash_defined : 𝚺₀-Function₂ (Hash.hash : V → V → V) via hashDef := by intro v; simp[hashDef, length_defined.df.iff, Exponential.defined.df.iff, ←le_iff_lt_succ] constructor · intro h; exact ⟨‖v 1‖, by simp, ‖v 2‖, by simp, rfl, rfl, by rw [h]; exact exponential_hash _ _⟩ · rintro ⟨_, _, _, _, rfl, rfl, h⟩; exact h.uniq (exponential_hash (v 1) (v 2)) -instance hash_definable : 𝚺₀-Function₂ (Hash.hash : M → M → M) := Defined.to_definable _ hash_defined +instance hash_definable : 𝚺₀-Function₂ (Hash.hash : V → V → V) := hash_defined.to_definable -@[simp] lemma hash_pow2 (a b : M) : Pow2 (a # b) := (exponential_hash a b).range_pow2 +@[simp] lemma hash_pow2 (a b : V) : Pow2 (a # b) := (exponential_hash a b).range_pow2 -@[simp] lemma hash_pos (a b : M) : 0 < a # b := (exponential_hash a b).range_pos +@[simp] lemma hash_pos (a b : V) : 0 < a # b := (exponential_hash a b).range_pos -@[simp] lemma hash_lt (a b : M) : ‖a‖ * ‖b‖ < a # b := (exponential_hash a b).lt +@[simp] lemma hash_lt (a b : V) : ‖a‖ * ‖b‖ < a # b := (exponential_hash a b).lt -lemma length_hash (a b : M) : ‖a # b‖ = ‖a‖ * ‖b‖ + 1 := (exponential_hash a b).length_eq +lemma length_hash (a b : V) : ‖a # b‖ = ‖a‖ * ‖b‖ + 1 := (exponential_hash a b).length_eq -@[simp] lemma hash_zero_left (a : M) : 0 # a = 1 := (exponential_hash 0 a).uniq (by simp) +@[simp] lemma hash_zero_left (a : V) : 0 # a = 1 := (exponential_hash 0 a).uniq (by simp) -@[simp] lemma hash_zero_right (a : M) : a # 0 = 1 := (exponential_hash a 0).uniq (by simp) +@[simp] lemma hash_zero_right (a : V) : a # 0 = 1 := (exponential_hash a 0).uniq (by simp) -lemma hash_comm (a b : M) : a # b = b # a := (exponential_hash a b).uniq (by simpa [mul_comm] using exponential_hash b a) +lemma hash_comm (a b : V) : a # b = b # a := (exponential_hash a b).uniq (by simpa [mul_comm] using exponential_hash b a) -@[simp] lemma lt_hash_one_right (a : M) : a < a # 1 := by +@[simp] lemma lt_hash_one_right (a : V) : a < a # 1 := by have : Exponential ‖a‖ (a # 1) := by simpa using (exponential_hash a 1) exact lt_exponential_length this -@[simp] lemma lt_hash_one_righs (a : M) : a # 1 ≤ 2 * a + 1 := by +@[simp] lemma lt_hash_one_righs (a : V) : a # 1 ≤ 2 * a + 1 := by rcases zero_le a with (rfl | pos) · simp · exact (le_iff_lt_length_of_exp (exponential_hash a 1)).mpr (by @@ -99,25 +99,25 @@ lemma hash_comm (a b : M) : a # b = b # a := (exponential_hash a b).uniq (by sim simpa using length_mul_pow2_add_of_lt pos (show Pow2 2 from by simp) one_lt_two simp [this]) -lemma lt_hash_iff {a b c : M} : a < b # c ↔ ‖a‖ ≤ ‖b‖ * ‖c‖ := (exponential_hash b c).lt_iff_len_le +lemma lt_hash_iff {a b c : V} : a < b # c ↔ ‖a‖ ≤ ‖b‖ * ‖c‖ := (exponential_hash b c).lt_iff_len_le -lemma hash_le_iff {a b c : M} : b # c ≤ a ↔ ‖b‖ * ‖c‖ < ‖a‖ := +lemma hash_le_iff {a b c : V} : b # c ≤ a ↔ ‖b‖ * ‖c‖ < ‖a‖ := not_iff_not.mp <| by simp [lt_hash_iff] -lemma lt_hash_one_iff {a b : M} : a < b # 1 ↔ ‖a‖ ≤ ‖b‖ := by simpa using lt_hash_iff (a := a) (b := b) (c := 1) +lemma lt_hash_one_iff {a b : V} : a < b # 1 ↔ ‖a‖ ≤ ‖b‖ := by simpa using lt_hash_iff (a := a) (b := b) (c := 1) -lemma hash_monotone {a₁ a₂ b₁ b₂ : M} (h₁ : a₁ ≤ b₁) (h₂ : a₂ ≤ b₂) : a₁ # a₂ ≤ b₁ # b₂ := +lemma hash_monotone {a₁ a₂ b₁ b₂ : V} (h₁ : a₁ ≤ b₁) (h₂ : a₂ ≤ b₂) : a₁ # a₂ ≤ b₁ # b₂ := (exponential_hash a₁ a₂).monotone_le (exponential_hash b₁ b₂) (mul_le_mul (length_monotone h₁) (length_monotone h₂) (by simp) (by simp)) -lemma bexp_eq_hash (a b : M) : bexp (a # b) (‖a‖ * ‖b‖) = a # b := bexp_eq_of_exp (by simp [length_hash]) (exponential_hash a b) +lemma bexp_eq_hash (a b : V) : bexp (a # b) (‖a‖ * ‖b‖) = a # b := bexp_eq_of_exp (by simp [length_hash]) (exponential_hash a b) -lemma hash_two_mul (a : M) {b} (pos : 0 < b) : a # (2 * b) = (a # b) * (a # 1) := by +lemma hash_two_mul (a : V) {b} (pos : 0 < b) : a # (2 * b) = (a # b) * (a # 1) := by have h₁ : Exponential (‖a‖ * ‖b‖ + ‖a‖) (a # (2 * b)) := by simpa [length_two_mul_of_pos pos, mul_add] using exponential_hash a (2 * b) have h₂ : Exponential (‖a‖ * ‖b‖ + ‖a‖) (a # b * a # 1) := (exponential_hash a b).add_mul (exponential_hash_one a) exact h₁.uniq h₂ -lemma hash_two_mul_le_sq_hash (a b : M) : a # (2 * b) ≤ (a # b) ^ 2 := by +lemma hash_two_mul_le_sq_hash (a b : V) : a # (2 * b) ≤ (a # b) ^ 2 := by rcases zero_le b with (rfl | pos) · simp · simp [hash_two_mul a pos, sq] diff --git a/Arithmetization/OmegaOne/Nuon.lean b/Arithmetization/OmegaOne/Nuon.lean index f877319..a3541b5 100644 --- a/Arithmetization/OmegaOne/Nuon.lean +++ b/Arithmetization/OmegaOne/Nuon.lean @@ -6,31 +6,31 @@ open FirstOrder FirstOrder.Arith noncomputable section -variable {M : Type*} [Zero M] [One M] [Add M] [Mul M] [LT M] [M ⊧ₘ* 𝐈𝚺₀ + 𝛀₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₀ + 𝛀₁] namespace Nuon -@[simp] lemma llen_lt_len_hash_len (K : M) : ‖‖K‖‖ < ‖K # ‖K‖‖ := by +@[simp] lemma llen_lt_len_hash_len (K : V) : ‖‖K‖‖ < ‖K # ‖K‖‖ := by simp [length_hash, lt_succ_iff_le] rcases zero_le ‖K‖ with (hK | pos) · simp [←hK] · exact le_mul_of_pos_left pos -lemma mul_len_lt_len_hash {i I L : M} (hi : i ≤ ‖I‖) : i * ‖L‖ < ‖I # L‖ := by +lemma mul_len_lt_len_hash {i I L : V} (hi : i ≤ ‖I‖) : i * ‖L‖ < ‖I # L‖ := by simp [length_hash, lt_succ_iff_le]; exact mul_le_mul_right' hi ‖L‖ -lemma mul_len_lt_len_hash' {i K z : M} (hi : i ≤ ‖z‖) : i * ‖‖K‖‖ < ‖z # ‖K‖‖ := by +lemma mul_len_lt_len_hash' {i K z : V} (hi : i ≤ ‖z‖) : i * ‖‖K‖‖ < ‖z # ‖K‖‖ := by simp [length_hash, lt_succ_iff_le]; exact mul_le_mul_right' hi ‖‖K‖‖ -def ext (L S i : M) : M := S / bexp S (i * ‖L‖) % (L # 1) +def ext (L S i : V) : V := S / bexp S (i * ‖L‖) % (L # 1) local notation S "{" L "}[" i "]" => ext L S i -lemma ext_eq_zero_of_lt {L S i : M} (h : ‖S‖ ≤ i * ‖L‖) : S{L}[i] = 0 := by simp [ext, bexp_eq_zero_of_le h] +lemma ext_eq_zero_of_lt {L S i : V} (h : ‖S‖ ≤ i * ‖L‖) : S{L}[i] = 0 := by simp [ext, bexp_eq_zero_of_le h] -@[simp] lemma ext_le_self (L S i : M) : S{L}[i] ≤ S := le_trans (mod_le _ _) (by simp [ext]) +@[simp] lemma ext_le_self (L S i : V) : S{L}[i] ≤ S := le_trans (mod_le _ _) (by simp [ext]) -lemma ext_graph_aux (z S L i : M) : z = S{L}[i] ↔ (‖S‖ ≤ i * ‖L‖ → z = 0) ∧ (i * ‖L‖ < ‖S‖ → ∃ b ≤ S, Exponential (i * ‖L‖) b ∧ z = S / b % (L # 1)) := by +lemma ext_graph_aux (z S L i : V) : z = S{L}[i] ↔ (‖S‖ ≤ i * ‖L‖ → z = 0) ∧ (i * ‖L‖ < ‖S‖ → ∃ b ≤ S, Exponential (i * ‖L‖) b ∧ z = S / b % (L # 1)) := by rcases show ‖S‖ ≤ i * ‖L‖ ∨ i * ‖L‖ < ‖S‖ from le_or_lt _ _ with (le | lt) · simp [ext_eq_zero_of_lt le, le, not_lt.mpr le] · simp [lt, not_le.mpr lt, ext] @@ -40,7 +40,7 @@ lemma ext_graph_aux (z S L i : M) : z = S{L}[i] ↔ (‖S‖ ≤ i * ‖L‖ → · rintro ⟨b, _, H, rfl⟩ rcases H.uniq (exp_bexp_of_lt lt); rfl -lemma ext_graph (z S L i : M) : z = S{L}[i] ↔ +lemma ext_graph (z S L i : V) : z = S{L}[i] ↔ ∃ lS ≤ S, lS = ‖S‖ ∧ ∃ lL ≤ L, lL = ‖L‖ ∧ (lS ≤ i * lL → z = 0) ∧ (i * lL < lS → @@ -76,24 +76,24 @@ def extDef : 𝚺₀.Semisentence 4 := .mkSigma @[simp] lemma cons_app_nine {n : ℕ} (a : α) (s : Fin n.succ.succ.succ.succ.succ.succ.succ.succ.succ → α) : (a :> s) 9 = s 8 := rfl -lemma ext_defined : 𝚺₀-Function₃ (ext : M → M → M → M) via extDef := by +lemma ext_defined : 𝚺₀-Function₃ (ext : V → V → V → V) via extDef := by intro v; simp [extDef, length_defined.df.iff, Exponential.defined.df.iff, div_defined.df.iff, rem_defined.df.iff, lt_succ_iff_le, ext_graph, numeral_eq_natCast] -instance ext_Definable : DefinableFunction₃ ℒₒᵣ 𝚺₀ (ext : M → M → M → M) := Defined.to_definable₀ _ ext_defined +instance ext_Definable : 𝚺₀-Function₃ (ext : V → V → V → V) := ext_defined.to_definable -instance : Bounded₃ ℒₒᵣ (ext : M → M → M → M) := ⟨#1, λ _ ↦ by simp⟩ +instance : Bounded₃ (ext : V → V → V → V) := ⟨#1, λ _ ↦ by simp⟩ -@[simp] lemma ext_zero (L i : M) : 0{L}[i] = 0 := by simp [ext] +@[simp] lemma ext_zero (L i : V) : 0{L}[i] = 0 := by simp [ext] -lemma ext_zero_eq_self_of_le {L S : M} (h : ‖S‖ ≤ ‖L‖) : S{L}[0] = S := by +lemma ext_zero_eq_self_of_le {L S : V} (h : ‖S‖ ≤ ‖L‖) : S{L}[0] = S := by rcases zero_le S with (rfl | pos) · simp [ext] · simp [ext] have : bexp S 0 = 1 := (exp_bexp_of_lt (show 0 < ‖S‖ from by simp [pos])).zero_uniq simp [this, lt_hash_one_iff.mpr h] -lemma ext_eq_of_ge {L S S' i : M} (h : S ≤ S') : S / bexp S' (i * ‖L‖) % (L # 1) = S{L}[i] := by +lemma ext_eq_of_ge {L S S' i : V} (h : S ≤ S') : S / bexp S' (i * ‖L‖) % (L # 1) = S{L}[i] := by rcases show i * ‖L‖ < ‖S‖ ∨ ‖S‖ ≤ i * ‖L‖ from lt_or_ge (i * ‖L‖) ‖S‖ with (lt | le) · unfold ext; congr 2; exact bexp_eq_of_lt_length (lt_of_lt_of_le lt $ length_monotone h) lt · simp [ext_eq_zero_of_lt le] @@ -102,17 +102,17 @@ lemma ext_eq_of_ge {L S S' i : M} (h : S ≤ S') : S / bexp S' (i * ‖L‖) % ( simp [this] · simp [bexp_eq_zero_of_le le'] -lemma ext_eq_of_gt {L S S' i : M} (h : i * ‖L‖ < ‖S'‖) : S / bexp S' (i * ‖L‖) % (L # 1) = S{L}[i] := by +lemma ext_eq_of_gt {L S S' i : V} (h : i * ‖L‖ < ‖S'‖) : S / bexp S' (i * ‖L‖) % (L # 1) = S{L}[i] := by rcases show i * ‖L‖ < ‖S‖ ∨ ‖S‖ ≤ i * ‖L‖ from lt_or_ge (i * ‖L‖) ‖S‖ with (lt | le) · unfold ext; congr 2; exact bexp_eq_of_lt_length h lt · simp [ext_eq_zero_of_lt le] have : S < bexp S' (i * ‖L‖) := ((exp_bexp_of_lt h).lt_iff_len_le).mpr le simp [this] -lemma ext_eq_hash_of_le {L S i : M} (h : i ≤ ‖I‖) : S / bexp (I # L) (i * ‖L‖) % (L # 1) = S{L}[i] := +lemma ext_eq_hash_of_le {L S i : V} (h : i ≤ ‖I‖) : S / bexp (I # L) (i * ‖L‖) % (L # 1) = S{L}[i] := ext_eq_of_gt (mul_len_lt_len_hash h) -lemma ext_add₁_pow2 {L i S₁ S₂ p : M} (pp : Pow2 p) (h : (i + 1) * ‖L‖ < ‖p‖) : +lemma ext_add₁_pow2 {L i S₁ S₂ p : V} (pp : Pow2 p) (h : (i + 1) * ‖L‖ < ‖p‖) : (S₁ + S₂ * p){L}[i] = S₁{L}[i] := by rcases zero_le S₂ with (rfl | pos₂) · simp @@ -133,12 +133,12 @@ lemma ext_add₁_pow2 {L i S₁ S₂ p : M} (pp : Pow2 p) (h : (i + 1) * ‖L‖ _ = S₁ / bexp (S₁ + S₂ * p) (i * ‖L‖) % L # 1 := by simp [mul_assoc] _ = S₁{L}[i] := ext_eq_of_ge le_self_add -lemma ext_add₁_bexp {L i j S₁ S₂ : M} (hi : i ≤ ‖I‖) (hij : j < i) : +lemma ext_add₁_bexp {L i j S₁ S₂ : V} (hi : i ≤ ‖I‖) (hij : j < i) : (S₁ + S₂ * bexp (I # L) (i * ‖L‖)){L}[j] = S₁{L}[j] := ext_add₁_pow2 (bexp_pow2 $ mul_len_lt_len_hash hi) (by rw [len_bexp (mul_len_lt_len_hash hi), lt_succ_iff_le]; exact mul_le_mul_right (succ_le_iff_lt.mpr hij)) -lemma ext_add₂_bexp {I i j S₁ S₂ : M} (hij : i + j ≤ ‖I‖) (hS₁ : ‖S₁‖ ≤ i * ‖L‖) : +lemma ext_add₂_bexp {I i j S₁ S₂ : V} (hij : i + j ≤ ‖I‖) (hS₁ : ‖S₁‖ ≤ i * ‖L‖) : (S₁ + S₂ * bexp (I # L) (i * ‖L‖)){L}[i + j] = S₂{L}[j] := by have hie : Exponential (i * ‖L‖) (bexp (I # L) (i * ‖L‖)) := exp_bexp_of_lt (mul_len_lt_len_hash $ le_trans le_self_add hij) calc (S₁ + S₂ * bexp (I # L) (i * ‖L‖)){L}[i + j] @@ -149,18 +149,18 @@ lemma ext_add₂_bexp {I i j S₁ S₂ : M} (hij : i + j ≤ ‖I‖) (hS₁ : congr 2; rw [div_add_mul_self, div_eq_zero_of_lt] <;> simp [hie.lt_iff_len_le.mpr hS₁, hie.range_pos] _ = S₂{L}[j] := ext_eq_hash_of_le (le_trans le_add_self hij) -def append (I L S i X : M) : M := S % bexp (I # L) (i * ‖L‖) + X * bexp (I # L) (i * ‖L‖) +def append (I L S i X : V) : V := S % bexp (I # L) (i * ‖L‖) + X * bexp (I # L) (i * ‖L‖) -lemma append_nil (I L S i : M) : append I L S i 0 = S % bexp (I # L) (i * ‖L‖) := by simp [append] +lemma append_nil (I L S i : V) : append I L S i 0 = S % bexp (I # L) (i * ‖L‖) := by simp [append] -lemma len_append (I L S : M) {i X} (hi : i ≤ ‖I‖) (hX : 0 < X) : ‖append I L S i X‖ = ‖X‖ + i * ‖L‖ := calc +lemma len_append (I L S : V) {i X} (hi : i ≤ ‖I‖) (hX : 0 < X) : ‖append I L S i X‖ = ‖X‖ + i * ‖L‖ := calc ‖append I L S i X‖ = ‖X * bexp (I # L) (i * ‖L‖) + S % bexp (I # L) (i * ‖L‖)‖ := by simp [append, add_comm] _ = ‖X‖ + log (bexp (I # L) (i * ‖L‖)) := length_mul_pow2_add_of_lt hX (bexp_pow2 $ mul_len_lt_len_hash hi) (mod_lt _ $ bexp_pos $ mul_len_lt_len_hash hi) _ = ‖X‖ + i * ‖L‖ := by simp [log_bexp (mul_len_lt_len_hash hi)] -lemma append_lt_hash (I L S : M) {i X} (hi : i < ‖I‖) (hX : ‖X‖ ≤ ‖L‖) : append I L S i X < I # L := by +lemma append_lt_hash (I L S : V) {i X} (hi : i < ‖I‖) (hX : ‖X‖ ≤ ‖L‖) : append I L S i X < I # L := by rcases zero_le X with (rfl | pos) · simp [append_nil] exact lt_of_lt_of_le (mod_lt _ (bexp_pos $ mul_len_lt_len_hash $ le_of_lt hi)) (by simp) @@ -169,7 +169,7 @@ lemma append_lt_hash (I L S : M) {i X} (hi : i < ‖I‖) (hX : ‖X‖ ≤ ‖L ‖X‖ + i * ‖L‖ ≤ (i + 1) * ‖L‖ := by simp [add_mul, add_comm (i * ‖L‖), hX] _ ≤ ‖I‖ * ‖L‖ := mul_le_mul_right (succ_le_iff_lt.mpr hi) -lemma append_lt_sq_hash (I L S : M) {i X} (hi : i ≤ ‖I‖) (hX : ‖X‖ ≤ ‖L‖) (Ipos : 0 < I) : append I L S i X < (I # L)^2 := by +lemma append_lt_sq_hash (I L S : V) {i X} (hi : i ≤ ‖I‖) (hX : ‖X‖ ≤ ‖L‖) (Ipos : 0 < I) : append I L S i X < (I # L)^2 := by rcases hi with (rfl | hi) · calc append I L S ‖I‖ X = S % I # L + X * I # L := by simp [append, bexp_eq_hash] @@ -178,14 +178,14 @@ lemma append_lt_sq_hash (I L S : M) {i X} (hi : i ≤ ‖I‖) (hX : ‖X‖ ≤ _ ≤ (I # L) ^ 2 := by simp [sq, hash_comm L 1]; exact hash_monotone (pos_iff_one_le.mp Ipos) (by rfl) · exact lt_of_lt_of_le (append_lt_hash I L S hi hX) (by simp) -lemma ext_append_last (I L S : M) {i X} (hi : i ≤ ‖I‖) (hX : ‖X‖ ≤ ‖L‖) : (append I L S i X){L}[i] = X := calc +lemma ext_append_last (I L S : V) {i X} (hi : i ≤ ‖I‖) (hX : ‖X‖ ≤ ‖L‖) : (append I L S i X){L}[i] = X := calc (append I L S i X){L}[i] = (S % bexp (I # L) (i * ‖L‖) + X * bexp (I # L) (i * ‖L‖)){L}[i + 0] := by simp [append] _ = X{L}[0] := ext_add₂_bexp (by simpa using hi) ((exp_bexp_of_lt (mul_len_lt_len_hash hi)).lt_iff_len_le.mp (mod_lt _ $ bexp_pos $ mul_len_lt_len_hash hi)) _ = X := ext_zero_eq_self_of_le hX -lemma ext_append_lt (I L S : M) {i j X} (hi : i ≤ ‖I‖) (hij : j < i) : +lemma ext_append_lt (I L S : V) {i j X} (hi : i ≤ ‖I‖) (hij : j < i) : (append I L S i X){L}[j] = S{L}[j] := calc (append I L S i X){L}[j] = (S % bexp (I # L) (i * ‖L‖) + X * bexp (I # L) (i * ‖L‖)){L}[j] := rfl _ = (S % bexp (I # L) (i * ‖L‖)){L}[j] := ext_add₁_bexp hi hij @@ -194,24 +194,24 @@ lemma ext_append_lt (I L S : M) {i j X} (hi : i ≤ ‖I‖) (hij : j < i) : section -variable {L A : M} +variable {L A : V} -def IsSegment (L A start intv S : M) : Prop := ∀ i < intv, S{L}[i + 1] = S{L}[i] + fbit A (start + i) +def IsSegment (L A start intv S : V) : Prop := ∀ i < intv, S{L}[i + 1] = S{L}[i] + fbit A (start + i) -def Segment (U L A start intv nₛ nₑ : M) : Prop := ∃ S < U, IsSegment L A start intv S ∧ S{L}[0] = nₛ ∧ S{L}[intv] = nₑ +def Segment (U L A start intv nₛ nₑ : V) : Prop := ∃ S < U, IsSegment L A start intv S ∧ S{L}[0] = nₛ ∧ S{L}[intv] = nₑ -def IsSeries (U I L A iter T : M) : Prop := ∀ l < iter, Segment U L A (‖I‖ * l) ‖I‖ (T{L}[l]) (T{L}[l + 1]) +def IsSeries (U I L A iter T : V) : Prop := ∀ l < iter, Segment U L A (‖I‖ * l) ‖I‖ (T{L}[l]) (T{L}[l + 1]) -def Series (U I L A iter n : M) : Prop := ∃ T < U, IsSeries U I L A iter T ∧ T{L}[0] = 0 ∧ T{L}[iter] = n +def Series (U I L A iter n : V) : Prop := ∃ T < U, IsSeries U I L A iter T ∧ T{L}[0] = 0 ∧ T{L}[iter] = n -def SeriesSegment (U I L A k n : M) : Prop := ∃ nₖ ≤ n, Series U I L A (k / ‖I‖) nₖ ∧ Segment U L A (‖I‖ * (k / ‖I‖)) (k % ‖I‖) nₖ n +def SeriesSegment (U I L A k n : V) : Prop := ∃ nₖ ≤ n, Series U I L A (k / ‖I‖) nₖ ∧ Segment U L A (‖I‖ * (k / ‖I‖)) (k % ‖I‖) nₖ n -lemma SeriesSegment.series {U I L A k n : M} (H : SeriesSegment U I L A k n) : +lemma SeriesSegment.series {U I L A k n : V} (H : SeriesSegment U I L A k n) : ∃ T S, IsSeries U I L A (k / ‖I‖) T ∧ IsSegment L A (‖I‖ * (k / ‖I‖)) (k % ‖I‖) S ∧ T{L}[0] = 0 ∧ T{L}[k / ‖I‖] = S{L}[0] ∧ S{L}[k % ‖I‖] = n := by rcases H with ⟨_, _, ⟨T, _, hT, hTₛ, hTₑ⟩, ⟨S, _, hS, rfl, rfl⟩⟩ exact ⟨T, S, hT, hS, hTₛ, hTₑ, rfl⟩ -lemma IsSegment.le_add {L A start intv S : M} (H : IsSegment L A start intv S) : ∀ i ≤ intv, S{L}[i] ≤ S{L}[0] + i := by +lemma IsSegment.le_add {L A start intv S : V} (H : IsSegment L A start intv S) : ∀ i ≤ intv, S{L}[i] ≤ S{L}[0] + i := by intro i induction i using induction_sigma0 · definability @@ -224,16 +224,16 @@ lemma IsSegment.le_add {L A start intv S : M} (H : IsSegment L A start intv S) : _ ≤ S{L}[i] + 1 := by simp _ ≤ S{L}[0] + (i + 1) := by simp [←add_assoc, IH] --- lemma Segment.refl (U L A start n : M) (hU : n < U) (hn : ‖n‖ ≤ ‖L‖) : Segment U L A start 0 n n := +-- lemma Segment.refl (U L A start n : V) (hU : n < U) (hn : ‖n‖ ≤ ‖L‖) : Segment U L A start 0 n n := -- ⟨n, hU, by intro; simp, ext_zero_eq_self_of_le hn, ext_zero_eq_self_of_le hn⟩ -lemma Segment.refl (U L A start n : M) (hL : L # 1 ≤ U) (hn : ‖n‖ ≤ ‖L‖) : Segment U L A start 0 n n := +lemma Segment.refl (U L A start n : V) (hL : L # 1 ≤ U) (hn : ‖n‖ ≤ ‖L‖) : Segment U L A start 0 n n := ⟨n, lt_of_lt_of_le (lt_hash_one_iff.mpr hn) hL, by intro; simp, ext_zero_eq_self_of_le hn, ext_zero_eq_self_of_le hn⟩ -lemma Segment.le_add {U L A start intv nₛ nₑ : M} (H : Segment U L A start intv nₛ nₑ) : nₑ ≤ nₛ + intv := by +lemma Segment.le_add {U L A start intv nₛ nₑ : V} (H : Segment U L A start intv nₛ nₑ) : nₑ ≤ nₛ + intv := by rcases H with ⟨S, _, hS, rfl, rfl⟩; exact hS.le_add intv (by rfl) -lemma Segment.uniq {U L A start intv nₛ nₑ₁ nₑ₂ : M} +lemma Segment.uniq {U L A start intv nₛ nₑ₁ nₑ₂ : V} (H₁ : Segment U L A start intv nₛ nₑ₁) (H₂ : Segment U L A start intv nₛ nₑ₂) : nₑ₁ = nₑ₂ := by rcases H₁ with ⟨S₁, _, HS₁, Hₛ, rfl⟩ rcases H₂ with ⟨S₂, _, HS₂, rfl, rfl⟩ @@ -247,7 +247,7 @@ lemma Segment.uniq {U L A start intv nₛ nₑ₁ nₑ₂ : M} have h₂ : S₂{L}[i + 1] = S₂{L}[i] + fbit A (start + i) := HS₂ i (lt_iff_succ_le.mpr hi) simp [IH (le_trans (by simp) hi), h₁, h₂] -lemma IsSeries.le_add {U I L A iter T : M} (H : IsSeries U I L A iter T) : ∀ l ≤ iter, T{L}[l] ≤ T{L}[0] + ‖I‖ * l := by +lemma IsSeries.le_add {U I L A iter T : V} (H : IsSeries U I L A iter T) : ∀ l ≤ iter, T{L}[l] ≤ T{L}[0] + ‖I‖ * l := by intro l induction l using induction_sigma0 · definability @@ -259,10 +259,10 @@ lemma IsSeries.le_add {U I L A iter T : M} (H : IsSeries U I L A iter T) : ∀ l T{L}[l + 1] ≤ T{L}[l] + ‖I‖ := (H l (succ_le_iff_lt.mp h)).le_add _ ≤ T{L}[0] + ‖I‖ * (l + 1) := by simpa [mul_add, ←add_assoc] using IH -lemma Series.le_add {U I L A iter n : M} (H : Series U I L A iter n) : n ≤ ‖I‖ * iter := by +lemma Series.le_add {U I L A iter n : V} (H : Series U I L A iter n) : n ≤ ‖I‖ * iter := by rcases H with ⟨T, _, hT, hzero, rfl⟩; simpa [hzero] using hT.le_add iter (by rfl) -lemma Series.uniq {U I L A iter n₁ n₂ : M} (H₁ : Series U I L A iter n₁) (H₂ : Series U I L A iter n₂) : n₁ = n₂ := by +lemma Series.uniq {U I L A iter n₁ n₂ : V} (H₁ : Series U I L A iter n₁) (H₂ : Series U I L A iter n₂) : n₁ = n₂ := by rcases H₁ with ⟨T₁, _, HT₁, Hₛ₁, rfl⟩ rcases H₂ with ⟨T₂, _, HT₂, Hₛ₂, rfl⟩ suffices ∀ i ≤ iter, T₁{L}[i] = T₂{L}[i] from this iter (by rfl) @@ -276,7 +276,7 @@ lemma Series.uniq {U I L A iter n₁ n₂ : M} (H₁ : Series U I L A iter n₁) have h₂ : Segment U L A (‖I‖ * i) ‖I‖ (T₁{L}[i]) (T₂{L}[i + 1]) := by simpa [IH] using HT₂ i (lt_iff_succ_le.mpr hi) exact h₁.uniq h₂ -lemma SeriesSegment.le {U I L A k n : M} (H : SeriesSegment U I L A k n) : +lemma SeriesSegment.le {U I L A k n : V} (H : SeriesSegment U I L A k n) : n ≤ k := by rcases H with ⟨nₖ, _, hT, hS⟩ calc @@ -284,22 +284,22 @@ lemma SeriesSegment.le {U I L A k n : M} (H : SeriesSegment U I L A k n) : _ ≤ ‖I‖ * (k / ‖I‖) + k % ‖I‖ := by simpa [mul_comm] using hT.le_add _ = k := div_add_mod k ‖I‖ -lemma SeriesSegment.initial {U I L A : M} (Upos : 0 < U) : SeriesSegment U I L A 0 0 := +lemma SeriesSegment.initial {U I L A : V} (Upos : 0 < U) : SeriesSegment U I L A 0 0 := ⟨0, by rfl, ⟨0, Upos, by simp [IsSeries]⟩, ⟨0, Upos, by simp [IsSegment]⟩⟩ -lemma SeriesSegment.zero {U I L k : M} (Upos : 0 < U) : SeriesSegment U I L 0 k 0 := +lemma SeriesSegment.zero {U I L k : V} (Upos : 0 < U) : SeriesSegment U I L 0 k 0 := ⟨0, by rfl, ⟨0, Upos, fun _ _ ↦ ⟨0, Upos, fun _ _ ↦ by simp, by simp⟩, by simp⟩, ⟨0, Upos, fun _ _ ↦ by simp, by simp⟩⟩ -lemma SeriesSegment.uniq {U I L A k n₁ n₂ : M} (H₁ : SeriesSegment U I L A k n₁) (H₂ : SeriesSegment U I L A k n₂) : +lemma SeriesSegment.uniq {U I L A k n₁ n₂ : V} (H₁ : SeriesSegment U I L A k n₁) (H₂ : SeriesSegment U I L A k n₂) : n₁ = n₂ := by rcases H₁ with ⟨nₘ₁, _, hT₁, hS₁⟩ rcases H₂ with ⟨nₘ₂, _, hT₂, hS₂⟩ rcases show nₘ₁ = nₘ₂ from hT₁.uniq hT₂ exact hS₁.uniq hS₂ -variable {U I L A : M} (hU : (I # L)^2 ≤ U) (hIL : ‖‖I‖^2‖ ≤ ‖L‖) (Ipos : 0 < I) +variable {U I L A : V} (hU : (I # L)^2 ≤ U) (hIL : ‖‖I‖^2‖ ≤ ‖L‖) (Ipos : 0 < I) -lemma Segment.succ {start intv nₛ nₑ : M} (H : Segment U L A start intv nₛ nₑ) (hintv : intv < ‖I‖) (hnₛ : ‖nₛ + ‖I‖‖ ≤ ‖L‖) : +lemma Segment.succ {start intv nₛ nₑ : V} (H : Segment U L A start intv nₛ nₑ) (hintv : intv < ‖I‖) (hnₛ : ‖nₛ + ‖I‖‖ ≤ ‖L‖) : Segment U L A start (intv + 1) nₛ (nₑ + fbit A (start + intv)) := by rcases H with ⟨S, _, H, rfl, rfl⟩ let S' := append I L S (intv + 1) (S{L}[intv] + fbit A (start + intv)) @@ -325,7 +325,7 @@ lemma Segment.succ {start intv nₛ nₑ : M} (H : Segment U L A start intv nₛ ext_append_lt I L S (succ_le_iff_lt.mpr hintv) (by simp), ext_append_last I L S (succ_le_iff_lt.mpr hintv) le_len_L ⟩ -lemma Series.succ {iter n n' : M} (HT : Series U I L A iter n) (HS : Segment U L A (‖I‖ * iter) ‖I‖ n n') (hiter : iter < ‖I‖) : +lemma Series.succ {iter n n' : V} (HT : Series U I L A iter n) (HS : Segment U L A (‖I‖ * iter) ‖I‖ n n') (hiter : iter < ‖I‖) : Series U I L A (iter + 1) n' := by have Hn : n ≤ ‖I‖ * iter := HT.le_add rcases HT with ⟨T, _, HT, Tₛ, rfl⟩ @@ -351,7 +351,7 @@ lemma Series.succ {iter n n' : M} (HT : Series U I L A iter n) (HS : Segment U L Eq.trans (ext_append_lt I L T (succ_le_iff_lt.mpr hiter) (by simp)) Tₛ, hTlast ⟩ -lemma div_mod_succ (a b : M) : ((a + 1) / b = a / b + 1 ∧ (a + 1) % b = 0 ∧ a % b + 1 = b) ∨ ((a + 1) / b = a / b ∧ (a + 1) % b = a % b + 1) := by +lemma div_mod_succ (a b : V) : ((a + 1) / b = a / b + 1 ∧ (a + 1) % b = 0 ∧ a % b + 1 = b) ∨ ((a + 1) / b = a / b ∧ (a + 1) % b = a % b + 1) := by rcases zero_le b with (rfl | pos) · simp have : a % b + 1 ≤ b := lt_iff_succ_le.mp <| mod_lt a (by simp [pos]) @@ -369,7 +369,7 @@ lemma div_mod_succ (a b : M) : ((a + 1) / b = a / b + 1 ∧ (a + 1) % b = 0 ∧ · rw [←this, mul_comm b, mod_mul_add _ _ pos] simp [ltb] -lemma SeriesSegment.succ {k n : M} (hk : k < ‖I‖^2) (H : SeriesSegment U I L A k n) : +lemma SeriesSegment.succ {k n : V} (hk : k < ‖I‖^2) (H : SeriesSegment U I L A k n) : SeriesSegment U I L A (k + 1) (n + fbit A k) := by have hhk : (k + 1)/‖I‖ ≤ ‖I‖ := by simpa using div_monotone (succ_le_iff_lt.mpr hk) ‖I‖ have hnk : n ≤ k := H.le @@ -413,24 +413,24 @@ section 4. $\| A \| < \|I\|^2$ -/ -def polyI (A : M) : M := bexp (2 * A) (√‖A‖) +def polyI (A : V) : V := bexp (2 * A) (√‖A‖) -def polyL (A : M) : M := ‖polyI A‖ ^ 2 +def polyL (A : V) : V := ‖polyI A‖ ^ 2 -def polyU (A : M) : M := (2 * A + 1) ^ 128 +def polyU (A : V) : V := (2 * A + 1) ^ 128 -lemma len_polyI {A : M} (pos : 0 < A) : ‖polyI A‖ = √‖A‖ + 1 := +lemma len_polyI {A : V} (pos : 0 < A) : ‖polyI A‖ = √‖A‖ + 1 := len_bexp (show √‖A‖ < ‖2 * A‖ from by simp [length_two_mul_of_pos pos, lt_succ_iff_le]) -lemma polyI_le {A : M} (pos : 0 < A) : ‖A‖ < ‖polyI A‖ ^ 2 := by simp [len_polyI pos] +lemma polyI_le {A : V} (pos : 0 < A) : ‖A‖ < ‖polyI A‖ ^ 2 := by simp [len_polyI pos] -lemma two_add_two_eq_four : 2 + 2 = (4 : M) := by simp [←three_add_one_eq_four, ←two_add_one_eq_three, ←one_add_one_eq_two, add_assoc] +lemma two_add_two_eq_four : 2 + 2 = (4 : V) := by simp [←three_add_one_eq_four, ←two_add_one_eq_three, ←one_add_one_eq_two, add_assoc] -lemma four_mul_eq_two_mul_two_mul (a : M) : 4 * a = 2 * (2 * a) := by simp [←two_add_two_eq_four, add_mul, two_mul] +lemma four_mul_eq_two_mul_two_mul (a : V) : 4 * a = 2 * (2 * a) := by simp [←two_add_two_eq_four, add_mul, two_mul] -@[simp] lemma two_mul_sqrt_le_self (a : M) : 2 * √a ≤ a + 1 := le_trans (two_mul_le_sq_add_one (√a)) (by simp) +@[simp] lemma two_mul_sqrt_le_self (a : V) : 2 * √a ≤ a + 1 := le_trans (two_mul_le_sq_add_one (√a)) (by simp) -lemma four_mul_hash_self (a : M) : (4 * a) # (4 * a) ≤ (a # a) ^ 16 := calc +lemma four_mul_hash_self (a : V) : (4 * a) # (4 * a) ≤ (a # a) ^ 16 := calc (4 * a) # (4 * a) ≤ ((4 * a) # (2 * a)) ^ 2 := by simp [four_mul_eq_two_mul_two_mul, hash_two_mul_le_sq_hash] _ ≤ ((4 * a) # a) ^ 4 := by simp [pow_four_eq_sq_sq, hash_two_mul_le_sq_hash] _ ≤ ((a # (2 * a)) ^ 2) ^ 4 := by rw [hash_comm (4 * a) a] @@ -438,14 +438,14 @@ lemma four_mul_hash_self (a : M) : (4 * a) # (4 * a) ≤ (a # a) ^ 16 := calc _ ≤ ((a # a) ^ 4) ^ 4 := by simp [pow_four_eq_sq_sq, hash_two_mul_le_sq_hash] _ ≤ (a # a) ^ 16 := by simp [←pow_mul] -@[simp] lemma pos_sq_iff {a : M} : 0 < √a ↔ 0 < a := +@[simp] lemma pos_sq_iff {a : V} : 0 < √a ↔ 0 < a := ⟨fun h ↦ lt_of_lt_of_le h (by simp), by intro h; by_contra A; simp at A; simp [show a = 0 from by simpa [A] using sqrt_lt_sq a] at h⟩ -@[simp] lemma pow_four_le_pow_four {a b : M} : a ^ 4 ≤ b ^ 4 ↔ a ≤ b := by simp [pow_four_eq_sq_sq] +@[simp] lemma pow_four_le_pow_four {a b : V} : a ^ 4 ≤ b ^ 4 ↔ a ≤ b := by simp [pow_four_eq_sq_sq] -lemma bexp_four_mul {a a' x : M} (hx : 4 * x < ‖a‖) (hx' : x < ‖a'‖) : +lemma bexp_four_mul {a a' x : V} (hx : 4 * x < ‖a‖) (hx' : x < ‖a'‖) : bexp a (4 * x) = (bexp a' x) ^ 4 := by rw [four_mul_eq_two_mul_two_mul, bexp_two_mul (a' := a), bexp_two_mul (a := a), pow_four_eq_sq_sq] · exact lt_of_le_of_lt (by simp [four_mul_eq_two_mul_two_mul]) hx @@ -453,7 +453,7 @@ lemma bexp_four_mul {a a' x : M} (hx : 4 * x < ‖a‖) (hx' : x < ‖a'‖) : · simpa [four_mul_eq_two_mul_two_mul] using hx · exact lt_of_le_of_lt (by simp [four_mul_eq_two_mul_two_mul]) hx -lemma polyI_hash_self_polybounded {A : M} (pos : 0 < A) : (polyI A) # (polyI A) ≤ (2 * A + 1) ^ 4 := calc +lemma polyI_hash_self_polybounded {A : V} (pos : 0 < A) : (polyI A) # (polyI A) ≤ (2 * A + 1) ^ 4 := calc (polyI A) # (polyI A) = bexp ((polyI A) # (polyI A)) ((√‖A‖ + 1) ^ 2) := Eq.symm <| by simpa [sq, len_polyI pos] using bexp_eq_hash (polyI A) (polyI A) _ ≤ bexp ((2 * A) # (2 * A)) ((2 * √‖A‖) ^ 2) := (bexp_monotone_le @@ -470,18 +470,18 @@ lemma polyI_hash_self_polybounded {A : M} (pos : 0 < A) : (polyI A) # (polyI A) _ = (A # 1) ^ 4 := by congr 1; simpa using bexp_eq_hash A 1 _ ≤ (2 * A + 1) ^ 4 := by simp -lemma polyI_hash_polyL_polybounded {A : M} (pos : 0 < A) : (polyI A) # (polyL A) ≤ (2 * A + 1) ^ 64 := calc +lemma polyI_hash_polyL_polybounded {A : V} (pos : 0 < A) : (polyI A) # (polyL A) ≤ (2 * A + 1) ^ 64 := calc (polyI A) # (polyL A) ≤ (polyI A) # (3 * polyI A) := hash_monotone (by rfl) (by simp [polyL, sq_len_le_three_mul]) _ ≤ (4 * polyI A) # (4 * polyI A) := hash_monotone (le_mul_of_pos_left $ by simp) (mul_le_mul_right $ by simp [←three_add_one_eq_four]) _ ≤ ((polyI A) # (polyI A)) ^ (4 * 4) := by simpa using four_mul_hash_self _ _ ≤ ((2 * A + 1) ^ 4) ^ (4 * 4) := by simp only [pow_mul, pow_four_le_pow_four, polyI_hash_self_polybounded pos] _ = (2 * A + 1) ^ 64 := by simp [←pow_mul] -lemma sq_polyI_hash_polyL_polybounded {A : M} (pos : 0 < A) : ((polyI A) # (polyL A)) ^ 2 ≤ polyU A := calc +lemma sq_polyI_hash_polyL_polybounded {A : V} (pos : 0 < A) : ((polyI A) # (polyL A)) ^ 2 ≤ polyU A := calc ((polyI A) # (polyL A)) ^ 2 ≤ ((2 * A + 1) ^ 64) ^ 2 := by simp [polyI_hash_polyL_polybounded pos] _ = polyU A := by simp [polyU, ←pow_mul] -def NuonAux (A k n : M) : Prop := SeriesSegment (polyU A) (polyI A) (polyL A) A k n +def NuonAux (A k n : V) : Prop := SeriesSegment (polyU A) (polyI A) (polyL A) A k n def isSegmentDef : 𝚺₀.Semisentence 5 := .mkSigma “L A start intv S | @@ -492,7 +492,7 @@ def isSegmentDef : 𝚺₀.Semisentence 5 := .mkSigma S_L_i_succ = S_L_i + fb” (by simp) -lemma isSegmentDef_defined : Arith.Defined (M := M) (λ v ↦ IsSegment (v 0) (v 1) (v 2) (v 3) (v 4)) isSegmentDef := by +lemma isSegmentDef_defined : 𝚺₀.Defined (V := V) (λ v ↦ IsSegment (v 0) (v 1) (v 2) (v 3) (v 4)) isSegmentDef := by intro v; simp [IsSegment, isSegmentDef, ext_defined.df.iff, fbit_defined.df.iff, lt_succ_iff_le, numeral_eq_natCast] apply forall₂_congr; intro x _ constructor @@ -502,7 +502,7 @@ lemma isSegmentDef_defined : Arith.Defined (M := M) (λ v ↦ IsSegment (v 0) (v def segmentDef : 𝚺₀.Semisentence 7 := .mkSigma “U L A start intv nₛ nₑ | ∃ S < U, !isSegmentDef L A start intv S ∧ !extDef nₛ L S 0 ∧ !extDef nₑ L S intv” (by simp) -lemma segmentDef_defined : Arith.Defined (M := M) (λ v ↦ Segment (v 0) (v 1) (v 2) (v 3) (v 4) (v 5) (v 6)) segmentDef := by +lemma segmentDef_defined : 𝚺₀.Defined (V := V) (λ v ↦ Segment (v 0) (v 1) (v 2) (v 3) (v 4) (v 5) (v 6)) segmentDef := by intro v; simp [Segment, segmentDef, ext_defined.df.iff, isSegmentDef_defined.df.iff, @Eq.comm _ (v 5), @Eq.comm _ (v 6)] rfl @@ -514,15 +514,15 @@ def isSeriesDef : 𝚺₀.Semisentence 6 := .mkSigma ∃ y <⁺ T, !extDef y L T (l + 1) ∧ !segmentDef U L A (lI * l) lI x y” (by simp) -lemma bex_eq_le_iff {p : M → Prop} {b : M} : +lemma bex_eq_le_iff {p : V → Prop} {b : V} : (∃ a ≤ z, a = b ∧ p a) ↔ (b ≤ z ∧ p b) := ⟨by rintro ⟨a, hp, rfl, hr⟩; exact ⟨hp, hr⟩, by rintro ⟨hp, hr⟩; exact ⟨b, hp, rfl, hr⟩⟩ -lemma bex_eq_lt_iff {p : M → Prop} {b : M} : +lemma bex_eq_lt_iff {p : V → Prop} {b : V} : (∃ a < z, a = b ∧ p a) ↔ (b < z ∧ p b) := ⟨by rintro ⟨a, hp, rfl, hr⟩; exact ⟨hp, hr⟩, by rintro ⟨hp, hr⟩; exact ⟨b, hp, rfl, hr⟩⟩ -lemma isSerieDef_defined : Arith.Defined (M := M) (λ v ↦ IsSeries (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) isSeriesDef := by +lemma isSerieDef_defined : 𝚺₀.Defined (V := V) (λ v ↦ IsSeries (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) isSeriesDef := by intro v; simp [IsSeries, isSeriesDef, length_defined.df.iff, ext_defined.df.iff, segmentDef_defined.df.iff, lt_succ_iff_le] apply forall₂_congr; intro x _ rw [bex_eq_le_iff, bex_eq_le_iff, bex_eq_le_iff] @@ -532,7 +532,7 @@ lemma isSerieDef_defined : Arith.Defined (M := M) (λ v ↦ IsSeries (v 0) (v 1) def seriesDef : 𝚺₀.Semisentence 6 := .mkSigma “U I L A iter n | ∃ T < U, !isSeriesDef U I L A iter T ∧ !extDef 0 L T 0 ∧ !extDef n L T iter” (by simp) -lemma seriesDef_defined : Arith.Defined (M := M) (λ v ↦ Series (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) seriesDef := by +lemma seriesDef_defined : 𝚺₀.Defined (V := V) (λ v ↦ Series (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) seriesDef := by intro v; simp [Series, seriesDef, isSerieDef_defined.df.iff, ext_defined.df.iff] apply exists_congr; intro T apply and_congr_right; intros @@ -548,7 +548,7 @@ def seriesSegmentDef : 𝚺₀.Semisentence 6 := .mkSigma !seriesDef U I L A d nₖ ∧ !segmentDef U L A (l * d) r nₖ n” (by simp) -lemma seriesSegmentDef_defined : Arith.Defined (M := M) (λ v ↦ SeriesSegment (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) seriesSegmentDef := by +lemma seriesSegmentDef_defined : 𝚺₀.Defined (V := V) (λ v ↦ SeriesSegment (v 0) (v 1) (v 2) (v 3) (v 4) (v 5)) seriesSegmentDef := by intro v; simp [SeriesSegment, seriesSegmentDef, length_defined.df.iff, div_defined.df.iff, rem_defined.df.iff, seriesDef_defined.df.iff, segmentDef_defined.df.iff, lt_succ_iff_le] apply exists_congr; intro nₖ @@ -563,31 +563,31 @@ def nuonAuxDef : 𝚺₀.Semisentence 3 := .mkSigma ∃ lg <⁺ g, !lengthDef lg g ∧ !seriesSegmentDef ((2 * A + 1) ^' 128) g (lg ²) A k n” (by simp) -lemma nuonAux_defined : 𝚺₀-Relation₃ (NuonAux : M → M → M → Prop) via nuonAuxDef := by +lemma nuonAux_defined : 𝚺₀-Relation₃ (NuonAux : V → V → V → Prop) via nuonAuxDef := by intro v; simp [NuonAux, polyU, polyI, polyL, nuonAuxDef, length_defined.df.iff, sqrt_defined.df.iff, bexp_defined.df.iff, seriesSegmentDef_defined.df.iff, lt_succ_iff_le, numeral_eq_natCast] rw [bex_eq_le_iff, bex_eq_le_iff, bex_eq_le_iff, bex_eq_le_iff]; simp -instance nuonAux_definable : DefinableRel₃ ℒₒᵣ 𝚺₀ (NuonAux : M → M → M → Prop) := Defined.to_definable _ nuonAux_defined +instance nuonAux_definable : 𝚺₀-Relation₃ (NuonAux : V → V → V → Prop) := nuonAux_defined.to_definable -instance : Bounded₃ ℒₒᵣ (ext : M → M → M → M) := ⟨#1, λ _ ↦ by simp⟩ +instance : Bounded₃ (ext : V → V → V → V) := ⟨#1, λ _ ↦ by simp⟩ -@[simp] lemma NuonAux.initial (A : M) : NuonAux A 0 0 := SeriesSegment.initial (by simp [polyU]) +@[simp] lemma NuonAux.initial (A : V) : NuonAux A 0 0 := SeriesSegment.initial (by simp [polyU]) -@[simp] lemma NuonAux.initial_iff (A n : M) : NuonAux A 0 n ↔ n = 0 := ⟨fun h ↦ h.uniq (NuonAux.initial A), by rintro rfl; simp⟩ +@[simp] lemma NuonAux.initial_iff (A n : V) : NuonAux A 0 n ↔ n = 0 := ⟨fun h ↦ h.uniq (NuonAux.initial A), by rintro rfl; simp⟩ -@[simp] lemma NuonAux.zero (k : M) : NuonAux 0 k 0 := SeriesSegment.zero (by simp [polyU]) +@[simp] lemma NuonAux.zero (k : V) : NuonAux 0 k 0 := SeriesSegment.zero (by simp [polyU]) -lemma NuonAux.le {A k n : M} (H : NuonAux A k n) : n ≤ k := SeriesSegment.le H +lemma NuonAux.le {A k n : V} (H : NuonAux A k n) : n ≤ k := SeriesSegment.le H -lemma NuonAux.uniq {A k n₁ n₂ : M} (H₁ : NuonAux A k n₁) (H₂ : NuonAux A k n₂) : n₁ = n₂ := SeriesSegment.uniq H₁ H₂ +lemma NuonAux.uniq {A k n₁ n₂ : V} (H₁ : NuonAux A k n₁) (H₂ : NuonAux A k n₂) : n₁ = n₂ := SeriesSegment.uniq H₁ H₂ -lemma NuonAux.succ {A k : M} (H : NuonAux A k n) (hk : k ≤ ‖A‖) : NuonAux A (k + 1) (n + fbit A k) := by +lemma NuonAux.succ {A k : V} (H : NuonAux A k n) (hk : k ≤ ‖A‖) : NuonAux A (k + 1) (n + fbit A k) := by rcases zero_le A with (rfl | pos) · rcases show n = 0 from H.uniq (NuonAux.zero k); simp exact SeriesSegment.succ (sq_polyI_hash_polyL_polybounded pos) (by simp [polyL]) (lt_of_le_of_lt hk $ polyI_le pos) H -lemma NuonAux.exists {k : M} (hk : k ≤ ‖A‖) : ∃ n, NuonAux A k n := by +lemma NuonAux.exists {k : V} (hk : k ≤ ‖A‖) : ∃ n, NuonAux A k n := by suffices ∃ n ≤ k, NuonAux A k n by rcases this with ⟨n, _, h⟩; exact ⟨n, h⟩ revert hk @@ -600,12 +600,12 @@ lemma NuonAux.exists {k : M} (hk : k ≤ ‖A‖) : ∃ n, NuonAux A k n := by rcases IH (le_trans (by simp) hk) with ⟨n, hn, Hn⟩ exact ⟨n + fbit A k, add_le_add hn (by simp), Hn.succ (le_trans (by simp) hk)⟩ -lemma NuonAux.succ_elim {A k : M} (hk : k ≤ ‖A‖) (H : NuonAux A (k + 1) n) : ∃ n', n = n' + fbit A k ∧ NuonAux A k n' := by +lemma NuonAux.succ_elim {A k : V} (hk : k ≤ ‖A‖) (H : NuonAux A (k + 1) n) : ∃ n', n = n' + fbit A k ∧ NuonAux A k n' := by rcases NuonAux.exists hk with ⟨n', H'⟩ rcases H.uniq (H'.succ hk) exact ⟨n', rfl, H'⟩ -lemma NuonAux.succ_iff {A k : M} (hk : k ≤ ‖A‖) : NuonAux A (k + 1) (n + fbit A k) ↔ NuonAux A k n := by +lemma NuonAux.succ_iff {A k : V} (hk : k ≤ ‖A‖) : NuonAux A (k + 1) (n + fbit A k) ↔ NuonAux A k n := by constructor · intro H rcases NuonAux.exists hk with ⟨n', H'⟩ @@ -613,7 +613,7 @@ lemma NuonAux.succ_iff {A k : M} (hk : k ≤ ‖A‖) : NuonAux A (k + 1) (n + f exact H' · exact (NuonAux.succ · hk) -lemma NuonAux.two_mul {k n : M} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux (2 * A) (k + 1) n := by +lemma NuonAux.two_mul {k n : V} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux (2 * A) (k + 1) n := by revert n hk suffices ∀ n ≤ k, k ≤ ‖A‖ → NuonAux A k n → NuonAux (2 * A) (k + 1) n by intro n hk H @@ -628,7 +628,7 @@ lemma NuonAux.two_mul {k n : M} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux have IH : NuonAux (2 * A) (k + 1) n' := IH n' H'.le (le_trans (by simp) hk) H' simpa using IH.succ (le_trans hk (length_monotone $ by simp)) -lemma NuonAux.two_mul_add_one {k n : M} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux (2 * A + 1) (k + 1) (n + 1) := by +lemma NuonAux.two_mul_add_one {k n : V} (hk : k ≤ ‖A‖) : NuonAux A k n → NuonAux (2 * A + 1) (k + 1) (n + 1) := by revert n hk suffices ∀ n ≤ k, k ≤ ‖A‖ → NuonAux A k n → NuonAux (2 * A + 1) (k + 1) (n + 1) by intro n hk H @@ -647,44 +647,44 @@ end end Nuon -def Nuon (A n : M) : Prop := Nuon.NuonAux A ‖A‖ n +def Nuon (A n : V) : Prop := Nuon.NuonAux A ‖A‖ n -lemma Nuon.exists_unique (A : M) : ∃! n, Nuon A n := by +lemma Nuon.exists_unique (A : V) : ∃! n, Nuon A n := by rcases show ∃ n, Nuon A n from NuonAux.exists (by simp) with ⟨n, hn⟩ exact ExistsUnique.intro n hn (fun n' hn' ↦ hn'.uniq hn) -def nuon (a : M) : M := Classical.choose! (Nuon.exists_unique a) +def nuon (a : V) : V := Classical.choose! (Nuon.exists_unique a) -@[simp] lemma nuon_nuon (a : M) : Nuon a (nuon a) := Classical.choose!_spec (Nuon.exists_unique a) +@[simp] lemma nuon_nuon (a : V) : Nuon a (nuon a) := Classical.choose!_spec (Nuon.exists_unique a) -lemma Nuon.nuon_eq {a b : M} (h : Nuon a b) : nuon a = b := (nuon_nuon a).uniq h +lemma Nuon.nuon_eq {a b : V} (h : Nuon a b) : nuon a = b := (nuon_nuon a).uniq h -lemma Nuon.nuon_eq_iff {a b : M} : b = nuon a ↔ Nuon a b := Classical.choose!_eq_iff (Nuon.exists_unique a) +lemma Nuon.nuon_eq_iff {a b : V} : b = nuon a ↔ Nuon a b := Classical.choose!_eq_iff (Nuon.exists_unique a) -lemma nuon_bit0 (a : M) : nuon (2 * a) = nuon a := by +lemma nuon_bit0 (a : V) : nuon (2 * a) = nuon a := by rcases zero_le a with (rfl | pos) · simp · have : Nuon (2 * a) (nuon a) := by simpa [Nuon, length_two_mul_of_pos pos] using (nuon_nuon a).two_mul (by simp) exact this.nuon_eq -lemma nuon_bit1 (a : M) : nuon (2 * a + 1) = nuon a + 1 := by +lemma nuon_bit1 (a : V) : nuon (2 * a + 1) = nuon a + 1 := by have : Nuon (2 * a + 1) (nuon a + 1) := by simpa [Nuon, length_two_mul_add_one] using (nuon_nuon a).two_mul_add_one (by simp) exact this.nuon_eq -@[simp] lemma nuon_zero : nuon (0 : M) = 0 := Nuon.nuon_eq (by simp [Nuon]) +@[simp] lemma nuon_zero : nuon (0 : V) = 0 := Nuon.nuon_eq (by simp [Nuon]) def _root_.LO.FirstOrder.Arith.nuonDef : 𝚺₀.Semisentence 2 := .mkSigma “n A | ∃ l <⁺ A, !lengthDef l A ∧ !Nuon.nuonAuxDef A l n” (by simp) -lemma nuon_defined : 𝚺₀-Function₁ (nuon : M → M) via nuonDef := by +lemma nuon_defined : 𝚺₀-Function₁ (nuon : V → V) via nuonDef := by intro v; simp [Nuon.nuon_eq_iff, Nuon, nuonDef, length_defined.df.iff, Nuon.nuonAux_defined.df.iff, lt_succ_iff_le] rw [Nuon.bex_eq_le_iff]; simp @[simp] lemma eval_nuon_iff (v) : - Semiformula.Evalbm M v nuonDef.val ↔ v 0 = nuon (v 1) :=nuon_defined.df.iff v + Semiformula.Evalbm V v nuonDef.val ↔ v 0 = nuon (v 1) :=nuon_defined.df.iff v -instance nuon_definable : 𝚺₀-Function₁ (nuon : M → M) := Defined.to_definable _ nuon_defined +instance nuon_definable : 𝚺₀-Function₁ (nuon : V → V) := nuon_defined.to_definable end From e898862f6c1c2e7746793a9b69d348c929305e21 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Fri, 16 Aug 2024 05:31:41 +0900 Subject: [PATCH 08/10] wip --- Arithmetization/ISigmaOne/HFS/Coding.lean | 2 +- Arithmetization/ISigmaOne/HFS/Vec.lean | 62 ++++++++++--------- .../ISigmaOne/Metamath/Formula/Basic.lean | 36 +++++------ .../ISigmaOne/Metamath/Formula/Functions.lean | 12 ++-- .../ISigmaOne/Metamath/Formula/Iteration.lean | 10 +-- .../ISigmaOne/Metamath/Language.lean | 18 +++--- .../ISigmaOne/Metamath/Proof/Derivation.lean | 26 ++++---- .../ISigmaOne/Metamath/Term/Basic.lean | 62 +++++++++---------- .../ISigmaOne/Metamath/Term/Functions.lean | 26 ++++---- .../ISigmaOne/Metamath/Theory/Theory.lean | 10 +-- 10 files changed, 133 insertions(+), 131 deletions(-) diff --git a/Arithmetization/ISigmaOne/HFS/Coding.lean b/Arithmetization/ISigmaOne/HFS/Coding.lean index 6fbcdb3..eecf13c 100644 --- a/Arithmetization/ISigmaOne/HFS/Coding.lean +++ b/Arithmetization/ISigmaOne/HFS/Coding.lean @@ -12,7 +12,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] def finsetArithmetizeAux : List V → V | [] => ∅ diff --git a/Arithmetization/ISigmaOne/HFS/Vec.lean b/Arithmetization/ISigmaOne/HFS/Vec.lean index 47ee0ab..354e0d3 100644 --- a/Arithmetization/ISigmaOne/HFS/Vec.lean +++ b/Arithmetization/ISigmaOne/HFS/Vec.lean @@ -169,6 +169,8 @@ lemma graph_defined : 𝚺₁-Predicate (Graph : V → Prop) via graphDef := instance graph_definable : 𝚺₁-Predicate (Graph : V → Prop) := graph_defined.to_definable +instance graph_definable' : 𝚺-[0 + 1]-Predicate (Graph : V → Prop) := graph_definable + end /-- TODO: move-/ @@ -201,7 +203,7 @@ lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by suffices ∀ i' ≤ i, ∀ v' ≤ v, ∃ x, Graph ⟪v', i', x⟫ from this i (by simp) v (by simp) intro i' hi' induction i' using induction_sigma1 - · { } + · definability case zero => intro v' _ exact ⟨fstIdx v', graph_case.mpr <| Or.inl ⟨v', rfl⟩⟩ @@ -209,7 +211,7 @@ lemma graph_exists (v i : V) : ∃ x, Graph ⟪v, i, x⟫ := by intro v' hv' rcases ih (le_trans le_self_add hi') (sndIdx v') (le_trans (by simp) hv') with ⟨x, hx⟩ exact ⟨x, graph_case.mpr <| Or.inr ⟨v', i', x, rfl, hx⟩⟩ -/--/ + lemma graph_unique {v i x₁ x₂ : V} : Graph ⟪v, i, x₁⟫ → Graph ⟪v, i, x₂⟫ → x₁ = x₂ := by induction i using induction_pi1 generalizing v x₁ x₂ · definability @@ -262,7 +264,7 @@ lemma cons_cases (x : V) : x = 0 ∨ ∃ y v, x = y ∷ v := by lemma cons_induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P) (nil : P 0) (cons : ∀ x v, P v → P (x ∷ v)) : ∀ v, P v := - order_induction_hh ℒₒᵣ Γ 1 hP (by + order_induction_hh Γ 1 hP (by intro v ih rcases nil_or_cons v with (rfl | ⟨x, v, rfl⟩) · exact nil @@ -292,9 +294,9 @@ lemma nth_defined : 𝚺₁-Function₂ (nth : V → V → V) via nthDef := by @[simp] lemma eval_nthDef (v) : Semiformula.Evalbm V v nthDef.val ↔ v 0 = nth (v 1) (v 2) := nth_defined.df.iff v -instance nth_definable : 𝚺₁-Function₂ (nth : V → V → V) := Defined.to_definable _ nth_defined +instance nth_definable : 𝚺₁-Function₂ (nth : V → V → V) := nth_defined.to_definable -instance nth_definable' (Γ) : (Γ, m + 1)-Function₂ (nth : V → V → V) := .of_sigmaOne nth_definable _ _ +instance nth_definable' (Γ m) : Γ-[m + 1]-Function₂ (nth : V → V → V) := .of_sigmaOne nth_definable _ _ end @@ -380,8 +382,8 @@ variable (V) structure Construction {arity : ℕ} (β : Blueprint arity) where nil (param : Fin arity → V) : V cons (param : Fin arity → V) (x xs ih) : V - nil_defined : DefinedFunction nil β.nil - cons_defined : DefinedFunction (fun v ↦ cons (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) β.cons + nil_defined : 𝚺₁.DefinedFunction nil β.nil + cons_defined : 𝚺₁.DefinedFunction (fun v ↦ cons (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) β.cons variable {V} @@ -427,14 +429,16 @@ def Graph : V → Prop := c.construction.Fixpoint param section -lemma graph_defined : Arith.Defined (fun v ↦ c.Graph (v ·.succ) (v 0)) β.graphDef := +lemma graph_defined : 𝚺₁.Defined (fun v ↦ c.Graph (v ·.succ) (v 0)) β.graphDef := c.construction.fixpoint_defined -instance graph_definable : Arith.Definable ℒₒᵣ 𝚺₁ (fun v ↦ c.Graph (v ·.succ) (v 0)) := Defined.to_definable _ c.graph_defined +instance graph_definable : 𝚺₁.Boldface (fun v ↦ c.Graph (v ·.succ) (v 0)) := c.graph_defined.to_definable instance graph_definable' (param) : 𝚺₁-Predicate (c.Graph param) := by simpa using HierarchySymbol.Boldface.retractiont (n := 1) c.graph_definable (#0 :> fun i ↦ &(param i)) +instance graph_definable'' (param) : 𝚺-[0 + 1]-Predicate (c.Graph param) := c.graph_definable' param + end variable {param} @@ -507,7 +511,7 @@ lemma result_eq_of_graph {xs y : V} (h : c.Graph param ⟪xs, y⟫) : c.result p section -lemma result_defined : Arith.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0)) β.resultDef := by +lemma result_defined : 𝚺₁.DefinedFunction (fun v ↦ c.result (v ·.succ) (v 0)) β.resultDef := by intro v; simp [Blueprint.resultDef, c.graph_defined.df.iff] constructor · intro h; rw [h]; exact c.result_graph _ _ @@ -516,11 +520,11 @@ lemma result_defined : Arith.DefinedFunction (fun v ↦ c.result (v ·.succ) (v @[simp] lemma eval_resultDef (v) : Semiformula.Evalbm V v β.resultDef.val ↔ v 0 = c.result (v ·.succ.succ) (v 1) := c.result_defined.df.iff v -instance result_definable : Arith.DefinableFunction ℒₒᵣ 𝚺₁ (fun v ↦ c.result (v ·.succ) (v 0)) := - Defined.to_definable _ c.result_defined +instance result_definable : 𝚺₁.BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0)) := + c.result_defined.to_definable instance result_definable' (Γ m) : - Arith.DefinableFunction ℒₒᵣ (Γ, m + 1) (fun v ↦ c.result (v ·.succ) (v 0)) := .of_sigmaOne c.result_definable _ _ + Γ-[m + 1].BoldfaceFunction (fun v ↦ c.result (v ·.succ) (v 0)) := .of_sigmaOne c.result_definable _ _ end @@ -567,9 +571,9 @@ lemma len_defined : 𝚺₁-Function₁ (len : V → V) via lenDef := constructi @[simp] lemma eval_lenDef (v) : Semiformula.Evalbm V v lenDef.val ↔ v 0 = len (v 1) := len_defined.df.iff v -instance len_definable : 𝚺₁-Function₁ (len : V → V) := Defined.to_definable _ len_defined +instance len_definable : 𝚺₁-Function₁ (len : V → V) := len_defined.to_definable -instance len_definable' (Γ) : (Γ, m + 1)-Function₁ (len : V → V) := .of_sigmaOne len_definable _ _ +instance len_definable' (Γ m) : Γ-[m + 1]-Function₁ (len : V → V) := .of_sigmaOne len_definable _ _ end @@ -688,9 +692,9 @@ lemma takeLast_defined : 𝚺₁-Function₂ (takeLast : V → V → V) via take @[simp] lemma eval_takeLastDef (v) : Semiformula.Evalbm V v takeLastDef.val ↔ v 0 = takeLast (v 1) (v 2) := takeLast_defined.df.iff v -instance takeLast_definable : 𝚺₁-Function₂ (takeLast : V → V → V) := Defined.to_definable _ takeLast_defined +instance takeLast_definable : 𝚺₁-Function₂ (takeLast : V → V → V) := takeLast_defined.to_definable -instance takeLast_definable' (Γ) : (Γ, m + 1)-Function₂ (takeLast : V → V → V) := .of_sigmaOne takeLast_definable _ _ +instance takeLast_definable' (Γ m) : Γ-[m + 1]-Function₂ (takeLast : V → V → V) := .of_sigmaOne takeLast_definable _ _ end @@ -775,9 +779,9 @@ lemma concat_defined : 𝚺₁-Function₂ (concat : V → V → V) via concatDe @[simp] lemma eval_concatDef (v) : Semiformula.Evalbm V v concatDef.val ↔ v 0 = concat (v 1) (v 2) := concat_defined.df.iff v -instance concat_definable : 𝚺₁-Function₂ (concat : V → V → V) := Defined.to_definable _ concat_defined +instance concat_definable : 𝚺₁-Function₂ (concat : V → V → V) := concat_defined.to_definable -instance concat_definable' (Γ) : (Γ, m + 1)-Function₂ (concat : V → V → V) := .of_sigmaOne concat_definable _ _ +instance concat_definable' (Γ m) : Γ-[m + 1]-Function₂ (concat : V → V → V) := .of_sigmaOne concat_definable _ _ end @@ -850,9 +854,9 @@ lemma memVec_defined : 𝚫₁-Relation (MemVec : V → V → Prop) via memVecDe @[simp] lemma eval_memVecDef (v) : Semiformula.Evalbm V v memVecDef.val ↔ v 0 ∈ᵥ v 1 := memVec_defined.df.iff v -instance memVec_definable : 𝚫₁-Relation (MemVec : V → V → Prop) := Defined.to_definable _ memVec_defined +instance memVec_definable : 𝚫₁-Relation (MemVec : V → V → Prop) := memVec_defined.to_definable -instance memVec_definable' (Γ) : (Γ, m + 1)-Relation (MemVec : V → V → Prop) := .of_deltaOne memVec_definable _ _ +instance memVec_definable' (Γ m) : Γ-[m + 1]-Relation (MemVec : V → V → Prop) := .of_deltaOne memVec_definable _ _ end @@ -881,9 +885,9 @@ def _root_.LO.FirstOrder.Arith.subsetVecDef : 𝚫₁.Semisentence 2 := .mkDelta (.mkPi “v w | ∀ x <⁺ v, !memVecDef.sigma x v → !memVecDef.pi x w” (by simp)) lemma subsetVec_defined : 𝚫₁-Relation (SubsetVec : V → V → Prop) via subsetVecDef := - ⟨by intro v; simp [subsetVecDef, HSemiformula.val_sigma, memVec_defined.proper.iff'], + ⟨by intro v; simp [subsetVecDef, HierarchySymbol.Semiformula.val_sigma, memVec_defined.proper.iff'], by intro v - simp [subsetVecDef, HSemiformula.val_sigma, memVec_defined.proper.iff'] + simp [subsetVecDef, HierarchySymbol.Semiformula.val_sigma, memVec_defined.proper.iff'] constructor · intro h x _; exact h x · intro h x hx; exact h x (le_of_memVec hx) hx⟩ @@ -891,9 +895,9 @@ lemma subsetVec_defined : 𝚫₁-Relation (SubsetVec : V → V → Prop) via su @[simp] lemma eval_subsetVecDef (v) : Semiformula.Evalbm V v subsetVecDef.val ↔ v 0 ⊆ᵥ v 1 := subsetVec_defined.df.iff v -instance subsetVec_definable : 𝚫₁-Relation (SubsetVec : V → V → Prop) := Defined.to_definable _ subsetVec_defined +instance subsetVec_definable : 𝚫₁-Relation (SubsetVec : V → V → Prop) := subsetVec_defined.to_definable -instance subsetVec_definable' (Γ) : (Γ, m + 1)-Relation (SubsetVec : V → V → Prop) := .of_deltaOne subsetVec_definable _ _ +instance subsetVec_definable' (Γ m) : Γ-[m + 1]-Relation (SubsetVec : V → V → Prop) := .of_deltaOne subsetVec_definable _ _ end @@ -934,9 +938,9 @@ lemma repeatVec_defined : 𝚺₁-Function₂ (repeatVec : V → V → V) via re @[simp] lemma eval_repeatVec (v) : Semiformula.Evalbm V v repeatVecDef.val ↔ v 0 = repeatVec (v 1) (v 2) := repeatVec_defined.df.iff v -instance repeatVec_definable : 𝚺₁-Function₂ (repeatVec : V → V → V) := Defined.to_definable _ repeatVec_defined +instance repeatVec_definable : 𝚺₁-Function₂ (repeatVec : V → V → V) := repeatVec_defined.to_definable -@[simp] instance repeatVec_definable' (Γ) : (Γ, m + 1)-Function₂ (repeatVec : V → V → V) := +@[simp] instance repeatVec_definable' (Γ) : Γ-[m + 1]-Function₂ (repeatVec : V → V → V) := .of_sigmaOne repeatVec_definable _ _ end @@ -1004,9 +1008,9 @@ lemma vecToSet_defined : 𝚺₁-Function₁ (vecToSet : V → V) via vecToSetDe @[simp] lemma eval_vecToSetDef (v) : Semiformula.Evalbm V v vecToSetDef.val ↔ v 0 = vecToSet (v 1) := vecToSet_defined.df.iff v -instance vecToSet_definable : 𝚺₁-Function₁ (vecToSet : V → V) := Defined.to_definable _ vecToSet_defined +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) := .of_sigmaOne vecToSet_definable _ _ end diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index 099d1a7..5716ff2 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -294,22 +294,22 @@ def construction : Fixpoint.Construction V (blueprint pL) where Φ := fun _ ↦ Phi L defined := ⟨ by intro v - -- simp [blueprint, HSemiformula.val_sigma, (termSeq_defined L).proper.iff'] - simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, HSemiformula.val_sigma, - HSemiformula.sigma_mkDelta, HSemiformula.val_mkSigma, LogicalConnective.HomClass.map_or, + -- 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, 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, HSemiformula.pi_mkDelta, HSemiformula.val_mkPi, + LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, (semitermVec_defined L).proper.iff'], by intro v - -- simpa [blueprint, Language.Defined.eval_rel_iff (L := L), eval_termSeq L, HSemiformula.val_sigma, formulaAux] using phi_iff L _ _ + -- simpa [blueprint, Language.Defined.eval_rel_iff (L := L), eval_termSeq L, HierarchySymbol.Semiformula.val_sigma, formulaAux] using phi_iff L _ _ simpa only [Fin.isValue, Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, - HSemiformula.val_sigma, formulaAux, HSemiformula.val_mkSigma, - LogicalConnective.HomClass.map_or, HSemiformula.val_mkDelta, Semiformula.eval_bexLT, + 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, @@ -375,7 +375,7 @@ lemma uformula_defined : 𝚫₁-Predicate L.UFormula via pL.uformulaDef := instance uformulaDef_definable : 𝚫₁-Predicate L.UFormula := Defined.to_definable _ (uformula_defined L) -@[simp, definability] instance uformulaDef_definable' (Γ) : (Γ, m + 1)-Predicate L.UFormula := +@[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 @@ -390,12 +390,12 @@ def _root_.LO.FirstOrder.Arith.LDef.isSemiformulaDef (pL : LDef) : 𝚫₁.Semis (.mkPi “n p | !pL.uformulaDef.pi p ∧ !fstIdxDef n p” (by simp)) lemma semiformula_defined : 𝚫₁-Relation L.Semiformula via pL.isSemiformulaDef where - left := by intro v; simp [LDef.isSemiformulaDef, HSemiformula.val_sigma, (uformula_defined L).proper.iff'] - right := by intro v; simp [LDef.isSemiformulaDef, HSemiformula.val_sigma, eval_uformulaDef L, Language.Semiformula, eq_comm] + 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 semiformula_definable : 𝚫₁-Relation L.Semiformula := Defined.to_definable _ (semiformula_defined L) -@[simp, definability] instance semiformula_defined' (Γ) : (Γ, m + 1)-Relation L.Semiformula := +@[simp, definability] instance semiformula_defined' (Γ) : Γ-[m + 1]-Relation L.Semiformula := .of_deltaOne (semiformula_definable L) _ _ variable {L} @@ -750,7 +750,7 @@ def construction : Fixpoint.Construction V (β.blueprint) where defined := ⟨by intro v /- - simp? [HSemiformula.val_sigma, Blueprint.blueprint, + simp? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_uformulaDef L, (uformula_defined L).proper.iff', c.rel_defined.iff, c.rel_defined.graph_delta.proper.iff', c.nrel_defined.iff, c.nrel_defined.graph_delta.proper.iff', @@ -764,7 +764,7 @@ 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, - HSemiformula.val_sigma, HSemiformula.sigma_mkDelta, HSemiformula.val_mkSigma, + 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', @@ -778,16 +778,16 @@ def construction : Fixpoint.Construction V (β.blueprint) where 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, HSemiformula.pi_mkDelta, HSemiformula.val_mkPi, + LogicalConnective.Prop.or_eq, HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, (uformula_defined L).proper.iff', c.rel_defined.graph_delta.proper.iff', - HSemiformula.graphDelta_val, c.nrel_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 [HSemiformula.val_sigma, Blueprint.blueprint, + simpa [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_uformulaDef L, c.rel_defined.iff, c.nrel_defined.iff, @@ -801,7 +801,7 @@ 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, - HSemiformula.val_sigma, HSemiformula.val_mkDelta, HSemiformula.val_mkSigma, + 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', @@ -1174,7 +1174,7 @@ section lemma result_defined : 𝚺₁-Function₂ c.result via β.result := by intro v - simp [Blueprint.result, HSemiformula.val_sigma, eval_uformulaDef L, (uformula_defined L).proper.iff', c.eval_graphDef] + simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, eval_uformulaDef L, (uformula_defined L).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 _ diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index d5ce81d..0606c03 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -76,7 +76,7 @@ lemma neg_defined : 𝚺₁-Function₁ L.neg via pL.negDef := fun v ↦ by instance neg_definable : 𝚺₁-Function₁ L.neg := Defined.to_definable _ (neg_defined L) -@[simp, definability] instance neg_definable' (Γ) : (Γ, m + 1)-Function₁ L.neg := +@[simp, definability] instance neg_definable' (Γ) : Γ-[m + 1]-Function₁ L.neg := .of_sigmaOne (neg_definable L) _ _ end @@ -183,7 +183,7 @@ lemma imp_defined : 𝚺₁-Function₃ L.imp via pL.impDef := fun v ↦ by instance imp_definable : 𝚺₁-Function₃ L.imp := Defined.to_definable _ (imp_defined L) -instance imp_definable' (Γ) : (Γ, m + 1)-Function₃ L.imp := +instance imp_definable' (Γ) : Γ-[m + 1]-Function₃ L.imp := .of_sigmaOne (imp_definable L) _ _ end @@ -255,7 +255,7 @@ lemma shift_defined : 𝚺₁-Function₁ L.shift via pL.shiftDef := fun v ↦ b instance shift_definable : 𝚺₁-Function₁ L.shift := Defined.to_definable _ (shift_defined L) -@[simp, definability] instance shift_definable' (Γ) : (Γ, m + 1)-Function₁ L.shift := +@[simp, definability] instance shift_definable' (Γ) : Γ-[m + 1]-Function₁ L.shift := .of_sigmaOne (shift_definable L) _ _ end @@ -341,7 +341,7 @@ lemma qVec_defined : 𝚺₁-Function₃ L.qVec via pL.qVecDef := by instance qVec_definable : 𝚺₁-Function₃ L.qVec := Defined.to_definable _ (qVec_defined L) -@[simp, definability] instance qVec_definable' (Γ m) : (Γ, m + 1)-Function₃ L.qVec := .of_sigmaOne (qVec_definable L) _ _ +@[simp, definability] instance qVec_definable' (Γ m) : Γ-[m + 1]-Function₃ L.qVec := .of_sigmaOne (qVec_definable L) _ _ end @@ -409,7 +409,7 @@ lemma substs_defined : 𝚺₁-Function₃ L.substs via pL.substsDef := fun v instance substs_definable : 𝚺₁-Function₃ L.substs := Defined.to_definable _ (substs_defined L) -@[simp, definability] instance substs_definable' (Γ) : (Γ, m + 1)-Function₃ L.substs := +@[simp, definability] instance substs_definable' (Γ) : Γ-[m + 1]-Function₃ L.substs := .of_sigmaOne (substs_definable L) _ _ end @@ -756,7 +756,7 @@ def _root_.LO.FirstOrder.Arith.LDef.isFVFreeDef (pL : LDef) : 𝚺₁.Semisenten .mkSigma “n p | !pL.isSemiformulaDef.sigma n p ∧ !pL.shiftDef p p” (by simp) lemma isFVFree_defined : 𝚺₁-Relation L.IsFVFree via pL.isFVFreeDef := by - intro v; simp [LDef.isFVFreeDef, HSemiformula.val_sigma, (semiformula_defined L).df.iff, (shift_defined L).df.iff] + intro v; simp [LDef.isFVFreeDef, HierarchySymbol.Semiformula.val_sigma, (semiformula_defined L).df.iff, (shift_defined L).df.iff] simp [Language.IsFVFree, eq_comm] end diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean index 9e0e4b1..3d5bb66 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean @@ -51,7 +51,7 @@ lemma qqConj_defined : 𝚺₁-Function₂ (qqConj : V → V → V) via qqConjDe instance qqConj_definable : 𝚺₁-Function₂ (qqConj : V → V → V) := Defined.to_definable _ qqConj_defined -instance qqConj_definable' (Γ) : (Γ, m + 1)-Function₂ (qqConj : V → V → V) := .of_sigmaOne qqConj_definable _ _ +instance qqConj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqConj : V → V → V) := .of_sigmaOne qqConj_definable _ _ end @@ -123,7 +123,7 @@ lemma qqDisj_defined : 𝚺₁-Function₂ (qqDisj : V → V → V) via qqDisjDe instance qqDisj_definable : 𝚺₁-Function₂ (qqDisj : V → V → V) := Defined.to_definable _ qqDisj_defined -instance qqDisj_definable' (Γ) : (Γ, m + 1)-Function₂ (qqDisj : V → V → V) := .of_sigmaOne qqDisj_definable _ _ +instance qqDisj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqDisj : V → V → V) := .of_sigmaOne qqDisj_definable _ _ end @@ -186,13 +186,13 @@ lemma substItr_defined : 𝚺₁-Function₄ (substItr : V → V → V → V → instance substItr_definable : 𝚺₁-Function₄ (substItr : V → V → V → V → V) := Defined.to_definable _ substItr_defined -@[simp, definability] instance substItr_definable' (Γ m) : (Γ, m + 1)-Function₄ (substItr : V → V → V → V → V) := +@[simp, definability] instance substItr_definable' (Γ m) : Γ-[m + 1]-Function₄ (substItr : V → V → V → V → V) := .of_sigmaOne substItr_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) := +instance substItr_definable₁' (n w p : V) (Γ m) : Γ-[m + 1]-Function₁ (substItr n w p) := .of_sigmaOne (substItr_definable₁ n w p) _ _ end @@ -302,7 +302,7 @@ lemma qqVerums_defined : 𝚺₁-Function₂ (qqVerums : V → V → V) via qqVe instance qqVerums_definable : 𝚺₁-Function₂ (qqVerums : V → V → V) := Defined.to_definable _ qqVerums_defined -@[simp] instance qqVerums_definable' (Γ) : (Γ, m + 1)-Function₂ (qqVerums : V → V → V) := +@[simp] instance qqVerums_definable' (Γ) : Γ-[m + 1]-Function₂ (qqVerums : V → V → V) := .of_sigmaOne qqVerums_definable _ _ end diff --git a/Arithmetization/ISigmaOne/Metamath/Language.lean b/Arithmetization/ISigmaOne/Metamath/Language.lean index 38a4030..bc6bce3 100644 --- a/Arithmetization/ISigmaOne/Metamath/Language.lean +++ b/Arithmetization/ISigmaOne/Metamath/Language.lean @@ -8,13 +8,13 @@ open FirstOrder FirstOrder.Arith section -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable (V) structure _root_.LO.FirstOrder.Arith.LDef where - func : HSemisentence ℒₒᵣ 2 𝚺₀ - rel : HSemisentence ℒₒᵣ 2 𝚺₀ + func : 𝚺₀.Semisentence 2 + rel : 𝚺₀.Semisentence 2 protected structure Language where Func (arity : V) : V → Prop @@ -36,15 +36,15 @@ variable {L : Arith.Language V} {pL : LDef} [L.Defined pL] @[simp] lemma Defined.eval_rel_iff (v) : Semiformula.Evalbm V v pL.rel.val ↔ L.Rel (v 0) (v 1) := Defined.rel.df.iff v -instance Defined.func_definable : 𝚺₀-Relation L.Func := Defined.to_definable _ Defined.func +instance Defined.func_definable : 𝚺₀-Relation L.Func := Defined.func.to_definable -instance Defined.rel_definable : 𝚺₀-Relation L.Rel := Defined.to_definable _ Defined.rel +instance Defined.rel_definable : 𝚺₀-Relation L.Rel := Defined.rel.to_definable -@[simp, definability] instance Defined.func_definable' (Γ) : Γ-Relation L.Func := - HierarchySymbol.Boldface.of_zero Defined.func_definable _ +@[simp, definability] instance Defined.func_definable' (ℌ) : ℌ-Relation L.Func := + HierarchySymbol.Boldface.of_zero Defined.func_definable -@[simp, definability] instance Defined.rel_definable' (Γ) : Γ-Relation L.Rel := - HierarchySymbol.Boldface.of_zero Defined.rel_definable _ +@[simp, definability] instance Defined.rel_definable' (ℌ) : ℌ-Relation L.Rel := + HierarchySymbol.Boldface.of_zero Defined.rel_definable end Language diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean index 3339f36..332ad7e 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean @@ -29,8 +29,8 @@ def _root_.LO.FirstOrder.Arith.LDef.formulaSetDef (pL : LDef) : 𝚫₁.Semisent variable (L) lemma formulaSet_defined : 𝚫₁-Predicate L.FormulaSet via pL.formulaSetDef := - ⟨by intro v; simp [LDef.formulaSetDef, HSemiformula.val_sigma, (semiformula_defined L).df.iff, (semiformula_defined L).proper.iff'], - by intro v; simp [LDef.formulaSetDef, HSemiformula.val_sigma, (semiformula_defined L).df.iff]; rfl⟩ + ⟨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⟩ @[simp] instance formulaSet_definable : 𝚫₁-Predicate L.FormulaSet := Defined.to_definable _ (formulaSet_defined L) @@ -477,7 +477,7 @@ def construction : Fixpoint.Construction V (blueprint pT) where ⟨by intro v /- - simp? [blueprint, HSemiformula.val_sigma, + 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, @@ -487,8 +487,8 @@ def construction : Fixpoint.Construction V (blueprint pT) where T.mem_defined.df.iff, T.mem_defined.proper.iff'] -/ - simp only [Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, Fin.isValue, HSemiformula.val_sigma, - HSemiformula.sigma_mkDelta, HSemiformula.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, @@ -503,14 +503,14 @@ def construction : Fixpoint.Construction V (blueprint pT) where 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, HSemiformula.pi_mkDelta, HSemiformula.val_mkPi, + 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'], by intro v /- - simp? [phi_iff, blueprint, HSemiformula.val_sigma, + 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, @@ -520,7 +520,7 @@ def construction : Fixpoint.Construction V (blueprint pT) where T.mem_defined.df.iff] -/ simp only [Fin.isValue, phi_iff, Nat.succ_eq_add_one, Nat.reduceAdd, blueprint, - HSemiformula.val_sigma, HSemiformula.val_mkDelta, HSemiformula.val_mkSigma, + 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, @@ -593,7 +593,7 @@ lemma derivation_defined : 𝚫₁-Predicate T.Derivation via pT.derivationDef : instance derivation_definable : 𝚫₁-Predicate T.Derivation := Defined.to_definable _ (derivation_defined T) -@[simp] instance derivatin_definable' (Γ) : (Γ, m + 1)-Predicate T.Derivation := +@[simp] instance derivatin_definable' (Γ) : Γ-[m + 1]-Predicate T.Derivation := .of_deltaOne (derivation_definable T) _ _ def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationOfDef {pL : LDef} (pT : pL.TDef) : 𝚫₁.Semisentence 2 := .mkDelta @@ -601,19 +601,19 @@ def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationOfDef {pL : LDef} (pT : pL.TD (.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, HSemiformula.val_sigma, (derivation_defined T).proper.iff'], - by intro v; simp [LDef.TDef.derivationOfDef, HSemiformula.val_sigma, (derivation_defined T).df.iff, eq_comm (b := fstIdx (v 0))]; rfl⟩ + ⟨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⟩ instance derivationOf_definable : 𝚫₁-Relation T.DerivationOf := Defined.to_definable _ (derivationOf_defined T) -@[simp] instance derivatinOf_definable' (Γ) : (Γ, m + 1)-Relation T.DerivationOf := +@[simp] instance derivatinOf_definable' (Γ) : Γ-[m + 1]-Relation T.DerivationOf := .of_deltaOne (derivationOf_definable T) _ _ 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) lemma Language.Theory.derivable_defined : 𝚺₁-Predicate T.Derivable via pT.derivableDef := by - intro v; simp [LDef.TDef.derivableDef, HSemiformula.val_sigma, (derivationOf_defined T).df.iff, Language.Theory.Derivable] + intro v; simp [LDef.TDef.derivableDef, HierarchySymbol.Semiformula.val_sigma, (derivationOf_defined T).df.iff, Language.Theory.Derivable] instance Language.Theory.derivable_definable : 𝚺₁-Predicate T.Derivable := Defined.to_definable _ (derivable_defined T) diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean index 7d87189..2c6218a 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Basic.lean @@ -7,7 +7,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -156,15 +156,15 @@ abbrev Language.Term : V → Prop := L.Semiterm 0 def _root_.LO.FirstOrder.Arith.LDef.isSemitermDef (pL : LDef) : 𝚫₁.Semisentence 2 := (blueprint pL).fixpointDefΔ₁.rew (Rew.substs ![#1, #0]) lemma isSemiterm_defined : 𝚫₁-Relation L.Semiterm via pL.isSemitermDef := - ⟨HSemiformula.ProperOn.rew (construction L).fixpoint_definedΔ₁.proper _, + ⟨HierarchySymbol.Semiformula.ProperOn.rew (construction L).fixpoint_definedΔ₁.proper _, by intro v; simp [LDef.isSemitermDef, (construction L).eval_fixpointDefΔ₁]; rfl⟩ @[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 isSemitermDef_definable : 𝚫₁-Relation (L.Semiterm) := Defined.to_definable _ (isSemiterm_defined L) +instance isSemitermDef_definable : 𝚫₁-Relation (L.Semiterm) := (isSemiterm_defined L).to_definable -@[simp, definability] instance isSemitermDef_definable' (Γ) : (Γ, m + 1)-Relation (L.Semiterm) := +@[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]) @@ -209,15 +209,15 @@ def _root_.LO.FirstOrder.Arith.LDef.semitermVecDef (pL : LDef) : 𝚫₁.Semisen variable (L) lemma semitermVec_defined : 𝚫₁-Relation₃ L.SemitermVec via pL.semitermVecDef := - ⟨by intro v; simp [LDef.semitermVecDef, HSemiformula.val_sigma, eval_isSemitermDef L, (isSemiterm_defined L).proper.iff'], - by intro v; simp [LDef.semitermVecDef, HSemiformula.val_sigma, eval_isSemitermDef L, Language.SemitermVec]⟩ + ⟨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]⟩ @[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 -instance semitermVecDef_definable : 𝚫₁-Relation₃ (L.SemitermVec) := Defined.to_definable _ (semitermVec_defined L) +instance semitermVecDef_definable : 𝚫₁-Relation₃ (L.SemitermVec) := (semitermVec_defined L).to_definable -@[simp, definability] instance semitermVecDef_definable' (Γ) : (Γ, m + 1)-Relation₃ (L.SemitermVec) := +@[simp, definability] instance semitermVecDef_definable' (Γ) : Γ-[m + 1]-Relation₃ (L.SemitermVec) := .of_deltaOne (semitermVecDef_definable L) _ _ end @@ -316,9 +316,9 @@ structure Construction (L : Arith.Language V) {k : ℕ} (φ : Blueprint pL k) wh 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_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 variable {V} @@ -376,15 +376,15 @@ def construction : Fixpoint.Construction V β.blueprint where defined := ⟨by intro v /- - simp? [HSemiformula.val_sigma, Blueprint.blueprint, + simp? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_isSemitermDef L, (isSemiterm_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'] -/ - simp only [Nat.succ_eq_add_one, Blueprint.blueprint, Nat.reduceAdd, HSemiformula.val_sigma, - BinderNotation.finSuccItr_one, Nat.add_zero, HSemiformula.sigma_mkDelta, - HSemiformula.val_mkSigma, Semiformula.eval_bexLTSucc', Semiterm.val_bvar, + simp only [Nat.succ_eq_add_one, Blueprint.blueprint, Nat.reduceAdd, HierarchySymbol.Semiformula.val_sigma, + BinderNotation.finSuccItr_one, Nat.add_zero, HierarchySymbol.Semiformula.sigma_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, @@ -396,21 +396,21 @@ def construction : Fixpoint.Construction V β.blueprint where 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, - HSemiformula.pi_mkDelta, HSemiformula.val_mkPi, (isSemiterm_defined L).proper.iff', - c.bvar_defined.graph_delta.proper.iff', HSemiformula.graphDelta_val, + HierarchySymbol.Semiformula.pi_mkDelta, HierarchySymbol.Semiformula.val_mkPi, (isSemiterm_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, LogicalConnective.Prop.arrow_eq, forall_eq, c.func_defined.graph_delta.proper.iff'] , by intro v /- - simpa? [HSemiformula.val_sigma, Blueprint.blueprint, eval_isSemitermDef L, + simpa? [HierarchySymbol.Semiformula.val_sigma, Blueprint.blueprint, eval_isSemitermDef L, c.bvar_defined.iff, c.fvar_defined.iff, c.func_defined.iff] using c.phi_iff _ _ _ _ -/ simpa only [Nat.succ_eq_add_one, BinderNotation.finSuccItr_one, Fin.succ_zero_eq_one, - Blueprint.blueprint, Nat.reduceAdd, HSemiformula.val_sigma, Nat.add_zero, - HSemiformula.val_mkDelta, HSemiformula.val_mkSigma, Semiformula.eval_bexLTSucc', + 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, LogicalConnective.HomClass.map_and, Semiformula.eval_substs, Matrix.comp_vecCons', Matrix.cons_val_two, Matrix.vecTail, Function.comp_apply, Matrix.cons_val_succ, @@ -458,7 +458,7 @@ lemma Graph.case_iff {t y : V} : variable (c) -lemma graph_defined : Arith.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.succ) (v 0) (v 1) (v 2)) β.graph := by intro v simp [Blueprint.graph, c.construction.fixpoint_defined.iff] constructor @@ -468,10 +468,10 @@ lemma graph_defined : Arith.Defined (fun v ↦ c.Graph (v ·.succ.succ.succ) (v @[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 -instance termSubst_definable : Arith.Definable ℒₒᵣ 𝚺₁ (fun v ↦ c.Graph (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) := - Defined.to_definable _ (graph_defined c) +instance termSubst_definable : 𝚺₁.Boldface (fun v ↦ c.Graph (v ·.succ.succ.succ) (v 0) (v 1) (v 2)) := + (graph_defined c).to_definable -@[simp, definability] instance termSubst_definable₂ (param n) : 𝚺₁-Relation (c.Graph param n) := by +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)) lemma graph_dom_isSemiterm {t y} : @@ -523,8 +523,7 @@ lemma graph_exists {t : V} : L.Semiterm n t → ∃ y, c.Graph param n t y := by intro x; exact ⟨c.fvar param n 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 apply HierarchySymbol.Boldface.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) ih + have : ∃ w, len w = k ∧ ∀ i < k, c.Graph param n 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⟩ @@ -587,8 +586,7 @@ 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)⟩ - rcases sigmaOne_skolem_vec - (by apply HierarchySymbol.Boldface.comp₂ (by definability) (by definability) (c.termSubst_definable₂ param n)) this + rcases sigmaOne_skolem_vec (by definability) this with ⟨w', hw'k, hw'⟩ refine ExistsUnique.intro w' ⟨hw'k.symm, hw'⟩ ?_ intro w'' ⟨hkw'', hw''⟩ @@ -642,9 +640,9 @@ variable (c) section -lemma result_defined : Arith.DefinedFunction (fun v ↦ c.result (v ·.succ.succ) (v 0) (v 1)) β.result := by +lemma result_defined : 𝚺₁.DefinedFunction (fun v ↦ c.result (v ·.succ.succ) (v 0) (v 1)) β.result := by intro v - simp [Blueprint.result, HSemiformula.val_sigma, (isSemiterm_defined L).proper.iff', + simp [Blueprint.result, HierarchySymbol.Semiformula.val_sigma, (isSemiterm_defined L).proper.iff', eval_isSemitermDef L, c.eval_graphDef, result, Classical.choose!_eq_iff] rfl @@ -657,9 +655,9 @@ private lemma resultVec_graph {w' k n w} : (¬L.SemitermVec k n w → w' = 0) ) := Classical.choose!_eq_iff (c.graph_existsUnique_vec_total param k n w) -lemma resultVec_defined : Arith.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.succ) (v 0) (v 1) (v 2)) β.resultVec := by intro v - simpa [Blueprint.resultVec, HSemiformula.val_sigma, (semitermVec_defined L).proper.iff', + simpa [Blueprint.resultVec, HierarchySymbol.Semiformula.val_sigma, (semitermVec_defined L).proper.iff', eval_semitermVecDef L, c.eval_graphDef] using c.resultVec_graph lemma eval_resultVec (v : Fin (arity + 4) → V) : diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean index cd12e89..d49f9ee 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean @@ -6,7 +6,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] section @@ -63,34 +63,34 @@ def _root_.LO.FirstOrder.Arith.LDef.termSubstVecDef (pL : LDef) : 𝚺₁.Semise variable (L) -lemma termSubst_defined : Arith.DefinedFunction (fun v ↦ L.termSubst (v 0) (v 1) (v 2) (v 3)) pL.termSubstDef := by +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] @[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 termSubst_definable : Arith.DefinableFunction ℒₒᵣ 𝚺₁ (fun v : Fin 4 → V ↦ L.termSubst (v 0) (v 1) (v 2) (v 3)) := - Defined.to_definable _ (termSubst_defined L) +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 termSubst_definable₂ (n m : V) : 𝚺₁-Function₂ (L.termSubst n m) := by - simpa using DefinableFunction.retractiont (n := 2) (termSubst_definable L) ![&n, &m, #0, #1] + simpa using HierarchySymbol.BoldfaceFunction.retractiont (n := 2) (termSubst_definable L) ![&n, &m, #0, #1] -@[simp, definability] instance termSubst_definable₂' (Γ k) (n m : V) : (Γ, k + 1)-Function₂ (L.termSubst n m) := +@[simp, definability] instance termSubst_definable₂' (Γ k) (n m : V) : Γ-[k + 1]-Function₂ (L.termSubst n m) := .of_sigmaOne (termSubst_definable₂ L n m) _ _ -lemma termSubstVec_defined : Arith.DefinedFunction (fun v ↦ L.termSubstVec (v 0) (v 1) (v 2) (v 3) (v 4)) pL.termSubstVecDef := by +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 : Arith.DefinableFunction ℒₒᵣ 𝚺₁ (fun v : Fin 5 → V ↦ L.termSubstVec (v 0) (v 1) (v 2) (v 3) (v 4)) := - Defined.to_definable _ (termSubstVec_defined L) +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 DefinableFunction.retractiont (n := 2) (termSubstVec_definable L) ![&k, &n, &m, #0, #1] + 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) := +@[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) _ _ end @@ -230,7 +230,7 @@ lemma termShift_defined : 𝚺₁-Function₂ L.termShift via pL.termShiftDef := 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 := - Defined.to_definable _ (termShift_defined L) + (termShift_defined L).to_definable @[definability, simp] instance termShift_definable' (Γ i) : (Γ, i + 1)-Function₂ L.termShift := .of_sigmaOne (termShift_definable L) _ _ @@ -633,7 +633,7 @@ lemma numeral_defined : 𝚺₁-Function₁ (numeral : V → V) via numeralDef : @[simp] instance numeral_definable : 𝚺₁-Function₁ (numeral : V → V) := Defined.to_definable _ numeral_defined -@[simp] instance numeral_definable' (Γ m) : (Γ, m + 1)-Function₁ (numeral : V → V) := .of_sigmaOne numeral_definable _ _ +@[simp] instance numeral_definable' (Γ m) : Γ-[m + 1]-Function₁ (numeral : V → V) := .of_sigmaOne numeral_definable _ _ end diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean b/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean index 6fdadd2..383c18c 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/Theory.lean @@ -169,11 +169,11 @@ def _root_.LO.FirstOrder.Arith.LDef.TDef.union {pL : LDef} (t u : pL.TDef) : pL. instance union_Defined_union (T U : L.Theory) {t u : pL.TDef} [T.Defined t] [U.Defined u] : (T.union U).Defined (t.union u) where defined := ⟨by simp [Arith.LDef.TDef.union] - apply HSemiformula.ProperOn.or + apply HierarchySymbol.Semiformula.ProperOn.or (Language.Theory.Defined.defined (T := T)).proper (Language.Theory.Defined.defined (T := U)).proper, by - intro v; simp [Arith.LDef.TDef.union, HSemiformula.or, Language.Theory.union, - HSemiformula.val_sigma, + intro v; simp [Arith.LDef.TDef.union, HierarchySymbol.Semiformula.or, Language.Theory.union, + HierarchySymbol.Semiformula.val_sigma, (Language.Theory.Defined.defined (T := T)).df.iff, (Language.Theory.Defined.defined (T := U)).df.iff]⟩ @@ -208,7 +208,7 @@ instance : (eqReplaceC (V := V)).Defined eqReplaceCDef where defined := by intro v simp [eqReplaceC, eqReplaceCDef, - HSemiformula.val_sigma, + 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] @@ -272,7 +272,7 @@ instance : (thEQ V).Defined thEQDef where defined := ⟨by { intro v simp [thEQDef, - HSemiformula.val_sigma, + HierarchySymbol.Semiformula.val_sigma, (imp_defined (Language.codeIn ℒₒᵣ V)).df.iff, (substs_defined (Language.codeIn ℒₒᵣ V)).df.iff, (semiformula_defined (Language.codeIn ℒₒᵣ V)).df.iff] From 6bc27ede71b29708fc8a0009220f658d552abc10 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sat, 17 Aug 2024 01:17:05 +0900 Subject: [PATCH 09/10] refactor --- Arithmetization/Definability/Boldface.lean | 2 +- Arithmetization/ISigmaOne/Ind.lean | 25 ++++--- .../ISigmaOne/Metamath/Coding.lean | 14 ++-- .../Metamath/DerivabilityConditions/D1.lean | 4 +- .../Metamath/DerivabilityConditions/D3.lean | 15 ++-- .../ISigmaOne/Metamath/Formula/Basic.lean | 68 +++++++++---------- .../ISigmaOne/Metamath/Formula/Functions.lean | 48 ++++++------- .../ISigmaOne/Metamath/Formula/Iteration.lean | 12 ++-- .../ISigmaOne/Metamath/Formula/Typed.lean | 2 +- .../ISigmaOne/Metamath/Proof/Derivation.lean | 18 ++--- .../ISigmaOne/Metamath/Proof/Thy.lean | 4 +- .../ISigmaOne/Metamath/Proof/Typed.lean | 2 +- .../ISigmaOne/Metamath/Term/Functions.lean | 18 ++--- .../ISigmaOne/Metamath/Term/Typed.lean | 4 +- .../ISigmaOne/Metamath/Theory/R.lean | 19 +++--- .../Metamath/Theory/SigmaOneDefinable.lean | 4 +- 16 files changed, 130 insertions(+), 129 deletions(-) diff --git a/Arithmetization/Definability/Boldface.lean b/Arithmetization/Definability/Boldface.lean index 9fbf8a1..cfccc89 100644 --- a/Arithmetization/Definability/Boldface.lean +++ b/Arithmetization/Definability/Boldface.lean @@ -708,7 +708,7 @@ lemma retraction {n k} {f : (Fin k → V) → V} (hf : ℌ.BoldfaceFunction f) ( have := Boldface.retraction (n := n + 1) hf (0 :> fun i ↦ (e i).succ); simp at this exact this.of_iff (by intro x; simp) -lemma retractiont {f : (Fin k → V) → V} (hf : ℌ.BoldfaceFunction f) (t : Fin k → Semiterm ℒₒᵣ V n) : +lemma retractiont {n k} {f : (Fin k → V) → V} (hf : ℌ.BoldfaceFunction f) (t : Fin k → Semiterm ℒₒᵣ V n) : ℌ.BoldfaceFunction fun v ↦ f (fun i ↦ Semiterm.valm V v id (t i)) := by have := Boldface.retractiont (n := n + 1) hf (#0 :> fun i ↦ Rew.bShift (t i)); simp at this exact this.of_iff (by intro x; simp [Semiterm.val_bShift']) diff --git a/Arithmetization/ISigmaOne/Ind.lean b/Arithmetization/ISigmaOne/Ind.lean index 7b9b052..e63555f 100644 --- a/Arithmetization/ISigmaOne/Ind.lean +++ b/Arithmetization/ISigmaOne/Ind.lean @@ -12,12 +12,11 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] - +variable {V : Type*} [ORingStruc V] variable (m : ℕ) [Fact (1 ≤ m)] [V ⊧ₘ* 𝐈𝐍𝐃𝚺 m] -lemma induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ : (𝚷, m)-Predicate Q) +lemma induction_sigma_or_pi {P Q : V → Prop} (hP : 𝚺-[m]-Predicate P) (hQ : 𝚷-[m]-Predicate Q) (zero : P 0 ∨ Q 0) (succ : ∀ x, P x ∨ Q x → P (x + 1) ∨ Q (x + 1)) : ∀ x, P x ∨ Q x := by haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) intro a @@ -40,7 +39,7 @@ lemma induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ have := this a (by rfl) simpa [hp, hq] using this -lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P) (hQ : (𝚷, m)-Predicate Q) +lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : 𝚺-[m]-Predicate P) (hQ : 𝚷-[m]-Predicate Q) (ind : ∀ x, (∀ y < x, P y ∨ Q y) → P x ∨ Q x) : ∀ x, P x ∨ Q x := by haveI : V ⊧ₘ* 𝐈𝚺₁ := mod_ISigma_of_le (show 1 ≤ m from Fact.out) intro a @@ -56,17 +55,17 @@ lemma order_induction_sigma_or_pi {P Q : V → Prop} (hP : (𝚺, m)-Predicate P · clear hp hq ind apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · simp_all only [zero_add, DefinableFunction.const] + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂ + · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] + · simp_all only [zero_add, HierarchySymbol.BoldfaceFunction.const] · simp_all only [Fin.isValue] apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.or - · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · simp_all only [zero_add, DefinableFunction.const] - · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · simp_all only [zero_add, DefinableFunction.const] + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂ + · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] + · simp_all only [zero_add, HierarchySymbol.BoldfaceFunction.const] + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂ + · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] + · simp_all only [zero_add, HierarchySymbol.BoldfaceFunction.const] case ind z ih => have : P z ∨ Q z := ind z (fun y hy ↦ by diff --git a/Arithmetization/ISigmaOne/Metamath/Coding.lean b/Arithmetization/ISigmaOne/Metamath/Coding.lean index eae612f..c3d23e1 100644 --- a/Arithmetization/ISigmaOne/Metamath/Coding.lean +++ b/Arithmetization/ISigmaOne/Metamath/Coding.lean @@ -21,7 +21,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] lemma nat_cast_empty : ((∅ : ℕ) : V) = ∅ := rfl @@ -79,7 +79,7 @@ namespace LO.FirstOrder.Semiterm open LO.Arith FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -136,7 +136,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Semiterm FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -229,7 +229,7 @@ namespace LO.FirstOrder.Semiformula open LO.Arith FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -347,7 +347,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith FirstOrder.Semiformula -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -411,7 +411,7 @@ namespace LO.FirstOrder.Derivation2 open LO.Arith FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → DecidableEq (L.Func k)] [(k : ℕ) → DecidableEq (L.Rel k)] [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -432,7 +432,7 @@ namespace LO.FirstOrder open LO.Arith FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] diff --git a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean index 7061f71..0f6d17c 100644 --- a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean +++ b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D1.lean @@ -8,7 +8,7 @@ namespace LO.FirstOrder open LO.Arith FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] @@ -90,7 +90,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith FirstOrder.Semiformula -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Language} [(k : ℕ) → Encodable (L.Func k)] [(k : ℕ) → Encodable (L.Rel k)] [DefinableLanguage L] diff --git a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean index 1f6b92d..069ecf3 100644 --- a/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean +++ b/Arithmetization/ISigmaOne/Metamath/DerivabilityConditions/D3.lean @@ -15,7 +15,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] namespace Formalized @@ -71,16 +71,17 @@ noncomputable def termEqComplete {n : ℕ} (e : Fin n → V) : | Semiterm.func Language.Zero.zero v => by simpa using eqRefl T _ | Semiterm.func Language.One.one v => by simpa using eqRefl T _ | Semiterm.func Language.Add.add v => by - simp [Rew.func, Semiterm.val_func] + simp only [Rew.func, Semiterm.codeIn'_add, Fin.isValue, subst_add, Semiterm.val_func, + 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 0).valbm V e + (v 1).valbm V e) := addComplete T _ _ + 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 _ _ 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 0).valbm V e * (v 1).valbm V e) := mulComplete T _ _ + 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 _ _ exact eqTrans T _ _ _ ⨀ ih ⨀ this lemma termEq_complete! {n : ℕ} (e : Fin n → V) (t : Semiterm ℒₒᵣ Empty n) : @@ -168,8 +169,8 @@ open LO.Arith LO.Arith.Formalized variable (T : Theory ℒₒᵣ) [Arith.DefinableSigma₁Theory T] class ISigma₁EQaddR₀ where - EQ : ∀ (V : Type) [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁], EQTheory (Theory.codeIn V T) - R0 : ∀ (V : Type) [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁], R₀Theory (Theory.codeIn V T) + EQ : ∀ (V : Type) [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁], EQTheory (Theory.codeIn V T) + R0 : ∀ (V : Type) [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁], R₀Theory (Theory.codeIn V T) end LO.FirstOrder.Theory @@ -177,7 +178,7 @@ namespace LO.Arith.Formalized open LO.FirstOrder -variable {V : Type} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable (T : Theory ℒₒᵣ) [Arith.DefinableSigma₁Theory T] diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean index 5716ff2..590673d 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Basic.lean @@ -6,7 +6,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -373,7 +373,7 @@ lemma uformula_defined : 𝚫₁-Predicate L.UFormula via pL.uformulaDef := @[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 := Defined.to_definable _ (uformula_defined L) +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) _ _ @@ -393,7 +393,7 @@ lemma semiformula_defined : 𝚫₁-Relation L.Semiformula via pL.isSemiformulaD 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 semiformula_definable : 𝚫₁-Relation L.Semiformula := Defined.to_definable _ (semiformula_defined L) +instance semiformula_definable : 𝚫₁-Relation L.Semiformula := (semiformula_defined L).to_definable @[simp, definability] instance semiformula_defined' (Γ) : Γ-[m + 1]-Relation L.Semiformula := .of_deltaOne (semiformula_definable L) _ _ @@ -530,7 +530,7 @@ lemma Language.UFormula.induction (Γ) {P : V → Prop} (hP : Γ-[1]-Predicate P · 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) +lemma Language.Semiformula.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]) @@ -658,14 +658,14 @@ structure Construction (L : Arith.Language V) (φ : Blueprint pL) where 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 + 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 @@ -887,7 +887,7 @@ lemma graph_defined : 𝚺₁-Relation₃ c.Graph via β.graph := by @[simp] lemma eval_graphDef (v) : Semiformula.Evalbm V v β.graph.val ↔ c.Graph (v 0) (v 1) (v 2) := (graph_defined β c).df.iff v -@[definability, simp] instance graph_definable : 𝚺₁-Relation₃ c.Graph := Defined.to_definable _ c.graph_defined +instance graph_definable : 𝚺-[0 + 1]-Relation₃ c.Graph := (c.graph_defined).to_definable variable {β} @@ -1203,18 +1203,18 @@ lemma uformula_result_induction {P : V → V → V → Prop} (hP : 𝚺₁-Relat 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 := - DefinableFunction.comp₂_infer (f := Max.max) - (DefinableFunction.var _) - (DefinableFunction.comp₂_infer - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _))) - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.comp₁_infer (DefinableFunction.var _)))) + 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 HierarchySymbol.Boldface.imp - (HierarchySymbol.Boldface.comp₁_infer (DefinableFunction.var _)) - (HierarchySymbol.Boldface.comp₃_infer - (DefinableFunction.var _) - (DefinableFunction.var _) - (DefinableFunction.comp₂_infer (DefinableFunction.var _) (DefinableFunction.var _))) + (HierarchySymbol.Boldface.comp₁ (HierarchySymbol.BoldfaceFunction.var _)) + (HierarchySymbol.Boldface.comp₃ + (HierarchySymbol.BoldfaceFunction.var _) + (HierarchySymbol.BoldfaceFunction.var _) + (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⟩ | @@ -1255,23 +1255,23 @@ lemma semiformula_result_induction {P : V → V → V → V → Prop} (hP : 𝚺 intro param p hp apply c.uformula_result_induction (P := fun param p y ↦ ∀ n ≤ p, n = fstIdx p → P param n p y) ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ ?_ hp - · apply HierarchySymbol.Boldface.ball_le (DefinableFunction.var _) + · 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, Fin.succ_zero_eq_one] apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.imp · simp_all only [SigmaPiDelta.alt_sigma, Fin.isValue] - apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂_infer - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] + apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₂ + · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] · simp_all only [zero_add, Fin.isValue] - apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer - simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₄_infer - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] + apply LO.FirstOrder.Arith.HierarchySymbol.BoldfaceFunction.comp₁ + simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] + · apply LO.FirstOrder.Arith.HierarchySymbol.Boldface.comp₄ + · simp_all only [zero_add, Fin.isValue, HierarchySymbol.BoldfaceFunction.var] · simp_all only [zero_add, Fin.isValue] - apply LO.FirstOrder.Arith.DefinableFunction.comp₁_infer - simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] - · simp_all only [zero_add, Fin.isValue, DefinableFunction.var] + apply LO.FirstOrder.Arith.HierarchySymbol.BoldfaceFunction.comp₁ + 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 diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean index 0606c03..abac961 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Functions.lean @@ -74,9 +74,9 @@ lemma neg_defined : 𝚺₁-Function₁ L.neg via pL.negDef := fun v ↦ by Semiformula.Evalbm (L := ℒₒᵣ) V v pL.negDef ↔ v 0 = L.neg (v 1) := (neg_defined L).df.iff v instance neg_definable : 𝚺₁-Function₁ L.neg := - Defined.to_definable _ (neg_defined L) + (neg_defined L).to_definable -@[simp, definability] instance neg_definable' (Γ) : Γ-[m + 1]-Function₁ L.neg := +instance neg_definable' (Γ) : Γ-[m + 1]-Function₁ L.neg := .of_sigmaOne (neg_definable L) _ _ end @@ -181,7 +181,7 @@ lemma imp_defined : 𝚺₁-Function₃ L.imp via pL.impDef := fun v ↦ by Semiformula.Evalbm V v pL.impDef.val ↔ v 0 = L.imp (v 1) (v 2) (v 3) := (imp_defined L).df.iff v instance imp_definable : 𝚺₁-Function₃ L.imp := - Defined.to_definable _ (imp_defined L) + (imp_defined L).to_definable instance imp_definable' (Γ) : Γ-[m + 1]-Function₃ L.imp := .of_sigmaOne (imp_definable L) _ _ @@ -253,9 +253,9 @@ lemma shift_defined : 𝚺₁-Function₁ L.shift via pL.shiftDef := fun v ↦ b Semiformula.Evalbm (L := ℒₒᵣ) V v pL.shiftDef ↔ v 0 = L.shift (v 1) := (shift_defined L).df.iff v instance shift_definable : 𝚺₁-Function₁ L.shift := - Defined.to_definable _ (shift_defined L) + (shift_defined L).to_definable -@[simp, definability] instance shift_definable' (Γ) : Γ-[m + 1]-Function₁ L.shift := +instance shift_definable' (Γ) : Γ-[m + 1]-Function₁ L.shift := .of_sigmaOne (shift_definable L) _ _ end @@ -339,7 +339,7 @@ def _root_.LO.FirstOrder.Arith.LDef.qVecDef (pL : LDef) : 𝚺₁.Semisentence 4 lemma qVec_defined : 𝚺₁-Function₃ L.qVec via pL.qVecDef := by intro v; simp [LDef.qVecDef, eval_termBShiftVecDef L]; rfl -instance qVec_definable : 𝚺₁-Function₃ L.qVec := Defined.to_definable _ (qVec_defined L) +instance qVec_definable : 𝚺₁-Function₃ L.qVec := (qVec_defined L).to_definable @[simp, definability] instance qVec_definable' (Γ m) : Γ-[m + 1]-Function₃ L.qVec := .of_sigmaOne (qVec_definable L) _ _ @@ -407,9 +407,9 @@ lemma substs_defined : 𝚺₁-Function₃ L.substs via pL.substsDef := fun v Semiformula.Evalbm (L := ℒₒᵣ) V v pL.substsDef ↔ v 0 = L.substs (v 1) (v 2) (v 3) := (substs_defined L).df.iff v instance substs_definable : 𝚺₁-Function₃ L.substs := - Defined.to_definable _ (substs_defined L) + (substs_defined L).to_definable -@[simp, definability] instance substs_definable' (Γ) : Γ-[m + 1]-Function₃ L.substs := +instance substs_definable' (Γ) : Γ-[m + 1]-Function₃ L.substs := .of_sigmaOne (substs_definable L) _ _ end @@ -459,11 +459,11 @@ lemma uformula_subst_induction {P : V → V → V → V → Prop} (hP : 𝚺₁- 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₄_infer - (DefinableFunction.comp₁_infer (DefinableFunction.var _)) - (DefinableFunction.comp₁_infer (DefinableFunction.var _)) - (DefinableFunction.var _) - (DefinableFunction.var _) + · 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) @@ -498,12 +498,12 @@ lemma semiformula_subst_induction {P : V → V → V → V → V → Prop} (hP : 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₅_infer - (DefinableFunction.var _) - (DefinableFunction.comp₁_infer (DefinableFunction.var _)) - (DefinableFunction.comp₁_infer (DefinableFunction.var _)) - (DefinableFunction.var _) - (DefinableFunction.var _) + · 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) @@ -611,9 +611,9 @@ lemma substs_substs {p} (hp : L.Semiformula l p) : apply HierarchySymbol.Boldface.all apply HierarchySymbol.Boldface.imp (by definability) apply HierarchySymbol.Boldface.imp (by definability) - apply HierarchySymbol.Boldface.comp₂_infer (by simp; definability) - apply DefinableFunction.comp₃_infer (by definability) ?_ (by definability) - apply DefinableFunction₅.comp (termSubstVec_definable _) <;> definability + apply HierarchySymbol.Boldface.comp₂ (by simp; definability) + apply HierarchySymbol.BoldfaceFunction.comp₃ (by definability) ?_ (by definability) + apply HierarchySymbol.BoldfaceFunction₅.comp (termSubstVec_definable _) <;> 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]) @@ -708,7 +708,7 @@ 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 -@[simp] instance substs₁_definable : 𝚺₁-Function₂ L.substs₁ := Defined.to_definable _ (substs₁_defined L) +@[simp] instance substs₁_definable : 𝚺₁-Function₂ L.substs₁ := (substs₁_defined L).to_definable end @@ -735,7 +735,7 @@ 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] -@[simp] instance free_definable : 𝚺₁-Function₁ L.free := Defined.to_definable _ (free_defined L) +instance free_definable : 𝚺-[0 + 1]-Function₁ L.free := (free_defined L).to_definable end diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean index 3d5bb66..7dce5b1 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Iteration.lean @@ -7,7 +7,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -49,7 +49,7 @@ lemma qqConj_defined : 𝚺₁-Function₂ (qqConj : V → V → V) via qqConjDe @[simp] lemma eval_qqConj (v) : Semiformula.Evalbm V v qqConjDef.val ↔ v 0 = qqConj (v 1) (v 2) := qqConj_defined.df.iff v -instance qqConj_definable : 𝚺₁-Function₂ (qqConj : V → V → V) := Defined.to_definable _ qqConj_defined +instance qqConj_definable : 𝚺₁-Function₂ (qqConj : V → V → V) := qqConj_defined.to_definable instance qqConj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqConj : V → V → V) := .of_sigmaOne qqConj_definable _ _ @@ -121,7 +121,7 @@ lemma qqDisj_defined : 𝚺₁-Function₂ (qqDisj : V → V → V) via qqDisjDe @[simp] lemma eval_qqDisj (v) : Semiformula.Evalbm V v qqDisjDef.val ↔ v 0 = qqDisj (v 1) (v 2) := qqDisj_defined.df.iff v -instance qqDisj_definable : 𝚺₁-Function₂ (qqDisj : V → V → V) := Defined.to_definable _ qqDisj_defined +instance qqDisj_definable : 𝚺₁-Function₂ (qqDisj : V → V → V) := qqDisj_defined.to_definable instance qqDisj_definable' (Γ) : Γ-[m + 1]-Function₂ (qqDisj : V → V → V) := .of_sigmaOne qqDisj_definable _ _ @@ -184,7 +184,7 @@ lemma substItr_defined : 𝚺₁-Function₄ (substItr : V → V → V → V → @[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) := Defined.to_definable _ substItr_defined +instance substItr_definable : 𝚺₁-Function₄ (substItr : V → V → V → V → V) := substItr_defined.to_definable @[simp, definability] instance substItr_definable' (Γ m) : Γ-[m + 1]-Function₄ (substItr : V → V → V → V → V) := .of_sigmaOne substItr_definable _ _ @@ -300,9 +300,9 @@ lemma qqVerums_defined : 𝚺₁-Function₂ (qqVerums : V → V → V) via qqVe @[simp] lemma qqVerums_repeatVec (v) : Semiformula.Evalbm V v qqVerumsDef.val ↔ v 0 = qqVerums (v 1) (v 2) := qqVerums_defined.df.iff v -instance qqVerums_definable : 𝚺₁-Function₂ (qqVerums : V → V → V) := Defined.to_definable _ qqVerums_defined +instance qqVerums_definable : 𝚺₁-Function₂ (qqVerums : V → V → V) := qqVerums_defined.to_definable -@[simp] instance qqVerums_definable' (Γ) : Γ-[m + 1]-Function₂ (qqVerums : V → V → V) := +instance qqVerums_definable' (Γ) : Γ-[m + 1]-Function₂ (qqVerums : V → V → V) := .of_sigmaOne qqVerums_definable _ _ end diff --git a/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean index 3e45f97..24cdb5f 100644 --- a/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Formula/Typed.lean @@ -13,7 +13,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean index 332ad7e..7711a2b 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Derivation.lean @@ -6,7 +6,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -32,7 +32,7 @@ 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⟩ -@[simp] instance formulaSet_definable : 𝚫₁-Predicate L.FormulaSet := Defined.to_definable _ (formulaSet_defined L) +@[simp] instance formulaSet_definable : 𝚫₁-Predicate L.FormulaSet := (formulaSet_defined L).to_definable end @@ -123,7 +123,7 @@ variable (L) lemma setShift_defined : 𝚺₁-Function₁ L.setShift via pL.setShiftDef := by intro v; simp [LDef.setShiftDef, setShift_graph, (shift_defined L).df.iff] -@[simp, definability] instance setShift_definable : 𝚺₁-Function₁ L.setShift := Defined.to_definable _ (setShift_defined L) +@[simp, definability] instance setShift_definable : 𝚺₁-Function₁ L.setShift := (setShift_defined L).to_definable end @@ -591,7 +591,7 @@ def _root_.LO.FirstOrder.Arith.LDef.TDef.derivationDef {pL : LDef} (pT : pL.TDef lemma derivation_defined : 𝚫₁-Predicate T.Derivation via pT.derivationDef := (construction T).fixpoint_definedΔ₁ -instance derivation_definable : 𝚫₁-Predicate T.Derivation := Defined.to_definable _ (derivation_defined T) +instance derivation_definable : 𝚫₁-Predicate T.Derivation := (derivation_defined T).to_definable @[simp] instance derivatin_definable' (Γ) : Γ-[m + 1]-Predicate T.Derivation := .of_deltaOne (derivation_definable T) _ _ @@ -604,7 +604,7 @@ lemma derivationOf_defined : 𝚫₁-Relation T.DerivationOf via pT.derivationOf ⟨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⟩ -instance derivationOf_definable : 𝚫₁-Relation T.DerivationOf := Defined.to_definable _ (derivationOf_defined T) +instance derivationOf_definable : 𝚫₁-Relation T.DerivationOf := (derivationOf_defined T).to_definable @[simp] instance derivatinOf_definable' (Γ) : Γ-[m + 1]-Relation T.DerivationOf := .of_deltaOne (derivationOf_definable T) _ _ @@ -615,10 +615,10 @@ def _root_.LO.FirstOrder.Arith.LDef.TDef.derivableDef {pL : LDef} (pT : pL.TDef) lemma Language.Theory.derivable_defined : 𝚺₁-Predicate T.Derivable via pT.derivableDef := by intro v; simp [LDef.TDef.derivableDef, HierarchySymbol.Semiformula.val_sigma, (derivationOf_defined T).df.iff, Language.Theory.Derivable] -instance Language.Theory.derivable_definable : 𝚺₁-Predicate T.Derivable := Defined.to_definable _ (derivable_defined T) +instance Language.Theory.derivable_definable : 𝚺₁-Predicate T.Derivable := (derivable_defined T).to_definable /-- instance for definability tactic-/ -@[simp] instance Language.Theory.derivable_definable' : (𝚺, 0 + 1)-Predicate T.Derivable := derivable_definable T +instance Language.Theory.derivable_definable' : 𝚺-[0 + 1]-Predicate T.Derivable := derivable_definable T def _root_.LO.FirstOrder.Arith.LDef.TDef.prv {pL : LDef} (pT : pL.TDef) : 𝚺₁.Semisentence 1 := .mkSigma “p | ∃ s, !insertDef s p 0 ∧ !pT.derivableDef s” (by simp) @@ -626,10 +626,10 @@ def _root_.LO.FirstOrder.Arith.LDef.TDef.prv {pL : LDef} (pT : pL.TDef) : 𝚺 protected lemma Language.Theory.provable_defined : 𝚺₁-Predicate T.Provable via pT.prv := by intro v; simp [LDef.TDef.prv, (derivable_defined T).df.iff, Language.Theory.Provable, singleton_eq_insert, emptyset_def] -instance Language.Theory.provable_definable : 𝚺₁-Predicate T.Provable := Defined.to_definable _ T.provable_defined +instance Language.Theory.provable_definable : 𝚺₁-Predicate T.Provable := T.provable_defined.to_definable /-- instance for definability tactic-/ -@[simp] instance Language.Theory.provable_definable' : (𝚺, 0 + 1)-Predicate T.Provable := T.provable_definable +instance Language.Theory.provable_definable' : 𝚺-[0 + 1]-Predicate T.Provable := T.provable_definable end diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Thy.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Thy.lean index e1b6a35..8064cdf 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Thy.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Thy.lean @@ -7,7 +7,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -16,7 +16,7 @@ section theory variable (L) structure _root_.LO.FirstOrder.Arith.LDef.TDef (pL : LDef) where - ch : HSemisentence ℒₒᵣ 1 𝚫₁ + ch : 𝚫₁.Semisentence 1 protected structure Language.Theory (L : Arith.Language V) {pL : LDef} [Arith.Language.Defined L pL] where set : Set V diff --git a/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean index 88081b3..6ba0082 100644 --- a/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Proof/Typed.lean @@ -14,7 +14,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean index d49f9ee..3502a95 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Functions.lean @@ -232,7 +232,7 @@ lemma termShift_defined : 𝚺₁-Function₂ L.termShift via pL.termShiftDef := instance termShift_definable : 𝚺₁-Function₂ L.termShift := (termShift_defined L).to_definable -@[definability, simp] instance termShift_definable' (Γ i) : (Γ, i + 1)-Function₂ L.termShift := .of_sigmaOne (termShift_definable L) _ _ +@[definability, simp] instance termShift_definable' (Γ i) : Γ-[i + 1]-Function₂ L.termShift := .of_sigmaOne (termShift_definable L) _ _ lemma termShiftVec_defined : 𝚺₁-Function₃ L.termShiftVec via pL.termShiftVecDef := by intro v; simpa [LDef.termShiftVecDef, Language.termShiftVec] using (construction L).resultVec_defined v @@ -241,9 +241,9 @@ lemma termShiftVec_defined : 𝚺₁-Function₃ L.termShiftVec via pL.termShift Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termShiftVecDef ↔ v 0 = L.termShiftVec (v 1) (v 2) (v 3) := (termShiftVec_defined L).df.iff v instance termShiftVec_definable : 𝚺₁-Function₃ L.termShiftVec := - Defined.to_definable _ (termShiftVec_defined L) + (termShiftVec_defined L).to_definable -@[simp, definability] instance termShiftVec_definable' (Γ i) : (Γ, i + 1)-Function₃ L.termShiftVec := +@[simp, definability] instance termShiftVec_definable' (Γ i) : Γ-[i + 1]-Function₃ L.termShiftVec := .of_sigmaOne (termShiftVec_definable L) _ _ end @@ -347,9 +347,9 @@ lemma termBShift_defined : 𝚺₁-Function₂ L.termBShift via pL.termBShiftDef 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 := - Defined.to_definable _ (termBShift_defined L) + (termBShift_defined L).to_definable -@[definability, simp] instance termBShift_definable' (Γ i) : (Γ, i + 1)-Function₂ L.termBShift := .of_sigmaOne (termBShift_definable L) _ _ +@[definability, simp] instance termBShift_definable' (Γ i) : Γ-[i + 1]-Function₂ L.termBShift := .of_sigmaOne (termBShift_definable L) _ _ lemma termBShiftVec_defined : 𝚺₁-Function₃ L.termBShiftVec via pL.termBShiftVecDef := by intro v; simpa using (construction L).resultVec_defined v @@ -358,9 +358,9 @@ lemma termBShiftVec_defined : 𝚺₁-Function₃ L.termBShiftVec via pL.termBSh Semiformula.Evalbm (L := ℒₒᵣ) V v pL.termBShiftVecDef ↔ v 0 = L.termBShiftVec (v 1) (v 2) (v 3) := (termBShiftVec_defined L).df.iff v instance termBShiftVec_definable : 𝚺₁-Function₃ L.termBShiftVec := - Defined.to_definable _ (termBShiftVec_defined L) + (termBShiftVec_defined L).to_definable -@[simp, definability] instance termBShiftVec_definable' (Γ i) : (Γ, i + 1)-Function₃ L.termBShiftVec := +@[simp, definability] instance termBShiftVec_definable' (Γ i) : Γ-[i + 1]-Function₃ L.termBShiftVec := .of_sigmaOne (termBShiftVec_definable L) _ _ end @@ -584,7 +584,7 @@ lemma numeralAux_defined : 𝚺₁-Function₁ (numeralAux : V → V) via numera @[simp] lemma eval_numeralAuxDef (v) : Semiformula.Evalbm V v numeralAuxDef.val ↔ v 0 = numeralAux (v 1) := numeralAux_defined.df.iff v -@[definability, simp] instance seqExp_definable : 𝚺₁-Function₁ (numeralAux : V → V) := Defined.to_definable _ numeralAux_defined +instance seqExp_definable : 𝚺-[0 + 1]-Function₁ (numeralAux : V → V) := numeralAux_defined.to_definable end @@ -631,7 +631,7 @@ 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) := Defined.to_definable _ numeral_defined +@[simp] 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 _ _ diff --git a/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean b/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean index 2f856d8..8b17be4 100644 --- a/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean +++ b/Arithmetization/ISigmaOne/Metamath/Term/Typed.lean @@ -12,7 +12,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -278,7 +278,7 @@ instance (n : V) : Add (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨add⟩ instance (n : V) : Mul (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨mul⟩ -instance (n : V) : Coe V (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨typedNumeral n⟩ +instance coeNumeral (n : V) : Coe V (⌜ℒₒᵣ⌝.TSemiterm n) := ⟨typedNumeral n⟩ variable {n : V} diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean index 3d1a6e4..027254a 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/R.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/R.lean @@ -14,7 +14,7 @@ namespace LO.Arith open FirstOrder FirstOrder.Arith -variable {V : Type*} [Zero V] [One V] [Add V] [Mul V] [LT V] [V ⊧ₘ* 𝐈𝚺₁] +variable {V : Type*} [ORingStruc V] [V ⊧ₘ* 𝐈𝚺₁] variable {L : Arith.Language V} {pL : LDef} [Arith.Language.Defined L pL] @@ -22,7 +22,7 @@ namespace Formalized variable (V) -abbrev LOR.Theory := @Language.Theory V _ _ _ _ _ ⌜ℒₒᵣ⌝ (Language.lDef ℒₒᵣ) _ +abbrev LOR.Theory := @Language.Theory V _ ⌜ℒₒᵣ⌝ (Language.lDef ℒₒᵣ) _ variable {V} @@ -37,12 +37,13 @@ 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 +variable (v : V) + class R₀Theory (T : LOR.TTheory (V := V)) extends EQTheory T where - add (n m : V) : T ⊢ (↑n + ↑m) =' ↑(n + m) - mul (n m : V) : T ⊢ (↑n * ↑m) =' ↑(n * m) + add (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) + mul (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 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 - /- section @@ -227,13 +228,13 @@ section R₀ variable [R₀Theory T] -def addComplete (n m : V) : T ⊢ (↑n + ↑m) =' ↑(n + m) := R₀Theory.add n m +def addComplete (n m : V) : T ⊢ (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) := R₀Theory.add n m -lemma add_complete! (n m : V) : T ⊢! (↑n + ↑m) =' ↑(n + m) := ⟨addComplete T n m⟩ +lemma add_complete! (n m : V) : T ⊢! (n + m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n + m) := ⟨addComplete T n m⟩ -def mulComplete (n m : V) : T ⊢ (↑n * ↑m) =' ↑(n * m) := R₀Theory.mul n m +def mulComplete (n m : V) : T ⊢ (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) := R₀Theory.mul n m -lemma mul_complete! (n m : V) : T ⊢! (↑n * ↑m) =' ↑(n * m) := ⟨mulComplete T n m⟩ +lemma mul_complete! (n m : V) : T ⊢! (n * m : ⌜ℒₒᵣ⌝[V].TSemiterm 0) =' ↑(n * m) := ⟨mulComplete T n m⟩ def neComplete {n m : V} (h : n ≠ m) : T ⊢ ↑n ≠' ↑m := R₀Theory.ne h diff --git a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean index 952d4f0..30d1277 100644 --- a/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean +++ b/Arithmetization/ISigmaOne/Metamath/Theory/SigmaOneDefinable.lean @@ -33,14 +33,14 @@ variable {T V} lemma Language.SyntacticTheory.codeIn_iff : x ∈ T.codeIn V ↔ V ⊧/![x] T.tDef.ch.val := iff_of_eq rfl lemma mem_coded_theory {σ} (h : σ ∈ T) : ⌜σ⌝ ∈ T.codeIn V := Language.SyntacticTheory.codeIn_iff.mpr <| by - have := consequence_iff_add_eq.mp (sound! <| SyntacticTheory.Δ₁HierarchySymbol.Boldface.mem_iff.mp h) V inferInstance + 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 instance tDef_defined : (T.codeIn V).Defined T.tDef where defined := ⟨by intro v rw [show v = ![v 0] from Matrix.constant_eq_singleton'] - have := consequence_iff_add_eq.mp (sound! <| FirstOrder.SyntacticTheory.Δ₁HierarchySymbol.Boldface.isΔ₁ (T := T)) V inferInstance + have := consequence_iff_add_eq.mp (sound! <| FirstOrder.SyntacticTheory.Δ₁Definable.isΔ₁ (T := T)) V inferInstance simp [models_iff] at this ⊢ simp [SyntacticTheory.tDef, this], by intro v; simp [SyntacticTheory.codeIn, ←Matrix.constant_eq_singleton']⟩ From 59c1c1bcc391053c4ccbd843415dee2f50db97f8 Mon Sep 17 00:00:00 2001 From: palalansouki Date: Sat, 17 Aug 2024 01:22:29 +0900 Subject: [PATCH 10/10] fix --- Arithmetization.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Arithmetization.lean b/Arithmetization.lean index 1f54f85..1f836eb 100644 --- a/Arithmetization.lean +++ b/Arithmetization.lean @@ -3,7 +3,10 @@ import Arithmetization.Vorspiel.Graph import Arithmetization.Vorspiel.Lemmata import Arithmetization.Definability.Init -import Arithmetization.Definability.Definability +import Arithmetization.Definability.Hierarchy +import Arithmetization.Definability.Boldface +import Arithmetization.Definability.BoundedBoldface +import Arithmetization.Definability.Absoluteness import Arithmetization.Basic.PeanoMinus import Arithmetization.Basic.Ind @@ -41,7 +44,6 @@ import Arithmetization.ISigmaOne.Metamath.Coding import Arithmetization.ISigmaOne.Metamath.Theory.R import Arithmetization.ISigmaOne.Metamath.Theory.SigmaOneDefinable -import Arithmetization.ISigmaOne.Metamath.Theory.Theory import Arithmetization.ISigmaOne.Metamath.DerivabilityConditions.D1 import Arithmetization.ISigmaOne.Metamath.DerivabilityConditions.D3