From 0fb6da6a76573eea2eedef4c2b4088245e586a80 Mon Sep 17 00:00:00 2001 From: twwar Date: Tue, 22 Jul 2025 10:56:30 +0200 Subject: [PATCH 01/19] reduction system attribute --- .../Computability/CombinatoryLogic/Basic.lean | 62 ++++++------ .../CombinatoryLogic/Confluence.lean | 28 +++--- .../Computability/CombinatoryLogic/Defs.lean | 64 +++++------- .../CombinatoryLogic/Recursion.lean | 98 +++++++++---------- Cslib/Semantics/ReductionSystem/Basic.lean | 20 ++-- 5 files changed, 132 insertions(+), 140 deletions(-) diff --git a/Cslib/Computability/CombinatoryLogic/Basic.lean b/Cslib/Computability/CombinatoryLogic/Basic.lean index 2926527e..10e830b2 100644 --- a/Cslib/Computability/CombinatoryLogic/Basic.lean +++ b/Cslib/Computability/CombinatoryLogic/Basic.lean @@ -61,7 +61,7 @@ def Polynomial.eval {n : Nat} (Γ : SKI.Polynomial n) (l : List SKI) (hl : List. def Polynomial.varFreeToSKI (Γ : SKI.Polynomial 0) : SKI := Γ.eval [] (by trivial) /-- Inductively define a polynomial `Γ'` so that (up to the fact that we haven't -defined reduction on polynomials) `Γ' ⬝ t ⇒* Γ[xₙ ← t]`. -/ +defined reduction on polynomials) `Γ' ⬝ t ↠ Γ[xₙ ← t]`. -/ def Polynomial.elimVar {n : Nat} : SKI.Polynomial (n+1) → SKI.Polynomial n /- The K-combinator leaves plain terms unchanged by substitution `K ⬝ x ⬝ t ⇒ x` -/ | SKI.Polynomial.term x => K ⬝' x @@ -83,7 +83,7 @@ for the inner variables. -/ theorem Polynomial.elimVar_correct {n : Nat} (Γ : SKI.Polynomial (n+1)) {ys : List SKI} (hys : ys.length = n) (z : SKI) : - Γ.elimVar.eval ys hys ⬝ z ⇒* Γ.eval (ys ++ [z]) + Γ.elimVar.eval ys hys ⬝ z ↠ Γ.eval (ys ++ [z]) (by rw [List.length_append, hys, List.length_singleton]) := by match n, Γ with @@ -125,13 +125,13 @@ def Polynomial.toSKI {n : Nat} (Γ : SKI.Polynomial n) : SKI := /-- Correctness for the toSKI (bracket abstraction) algorithm. -/ theorem Polynomial.toSKI_correct {n : Nat} (Γ : SKI.Polynomial n) (xs : List SKI) - (hxs : xs.length = n) : Γ.toSKI.applyList xs ⇒* Γ.eval xs hxs := by + (hxs : xs.length = n) : Γ.toSKI.applyList xs ↠ Γ.eval xs hxs := by match n with | 0 => unfold toSKI varFreeToSKI applyList rw [List.length_eq_zero_iff] at hxs simp_rw [hxs, List.foldl_nil] - apply MRed.refl + rfl | n+1 => -- show that xs = ys + [z] have : xs ≠ [] := List.ne_nil_of_length_eq_add_one hxs @@ -165,7 +165,7 @@ choose a descriptive name. def RPoly : SKI.Polynomial 2 := &1 ⬝' &0 /-- A SKI term representing R -/ def R : SKI := RPoly.toSKI -theorem R_def (x y : SKI) : R ⬝ x ⬝ y ⇒* y ⬝ x := +theorem R_def (x y : SKI) : R ⬝ x ⬝ y ↠ y ⬝ x := RPoly.toSKI_correct [x, y] (by simp) @@ -173,7 +173,7 @@ theorem R_def (x y : SKI) : R ⬝ x ⬝ y ⇒* y ⬝ x := def BPoly : SKI.Polynomial 3 := &0 ⬝' (&1 ⬝' &2) /-- A SKI term representing B -/ def B : SKI := BPoly.toSKI -theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ⇒* f ⬝ (g ⬝ x) := +theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ↠ f ⬝ (g ⬝ x) := BPoly.toSKI_correct [f, g, x] (by simp) @@ -181,7 +181,7 @@ theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ⇒* f ⬝ (g ⬝ x) := def CPoly : SKI.Polynomial 3 := &0 ⬝' &2 ⬝' &1 /-- A SKI term representing C -/ def C : SKI := CPoly.toSKI -theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ⇒* f ⬝ y ⬝ x := +theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ↠ f ⬝ y ⬝ x := CPoly.toSKI_correct [f, x, y] (by simp) @@ -189,7 +189,7 @@ theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ⇒* f ⬝ y ⬝ x := def RotRPoly : SKI.Polynomial 3 := &2 ⬝' &0 ⬝' &1 /-- A SKI term representing RotR -/ def RotR : SKI := RotRPoly.toSKI -theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ⇒* z ⬝ x ⬝ y := +theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ↠ z ⬝ x ⬝ y := RotRPoly.toSKI_correct [x, y, z] (by simp) @@ -197,7 +197,7 @@ theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ⇒* z ⬝ x ⬝ y := def RotLPoly : SKI.Polynomial 3 := &1 ⬝' &2 ⬝' &0 /-- A SKI term representing RotL -/ def RotL : SKI := RotLPoly.toSKI -theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ⇒* y ⬝ z ⬝ x := +theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ↠ y ⬝ z ⬝ x := RotLPoly.toSKI_correct [x, y, z] (by simp) @@ -205,7 +205,7 @@ theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ⇒* y ⬝ z ⬝ x := def δPoly : SKI.Polynomial 1 := &0 ⬝' &0 /-- A SKI term representing δ -/ def δ : SKI := δPoly.toSKI -theorem δ_def (x : SKI) : δ ⬝ x ⇒* x ⬝ x := +theorem δ_def (x : SKI) : δ ⬝ x ↠ x ⬝ x := δPoly.toSKI_correct [x] (by simp) @@ -213,7 +213,7 @@ theorem δ_def (x : SKI) : δ ⬝ x ⇒* x ⬝ x := def HPoly : SKI.Polynomial 2 := &0 ⬝' (&1 ⬝' &1) /-- A SKI term representing H -/ def H : SKI := HPoly.toSKI -theorem H_def (f x : SKI) : H ⬝ f ⬝ x ⇒* f ⬝ (x ⬝ x) := +theorem H_def (f x : SKI) : H ⬝ f ⬝ x ↠ f ⬝ (x ⬝ x) := HPoly.toSKI_correct [f, x] (by simp) @@ -221,7 +221,7 @@ theorem H_def (f x : SKI) : H ⬝ f ⬝ x ⇒* f ⬝ (x ⬝ x) := def YPoly : SKI.Polynomial 1 := H ⬝' &0 ⬝' (H ⬝' &0) /-- A SKI term representing Y -/ def Y : SKI := YPoly.toSKI -theorem Y_def (f : SKI) : Y ⬝ f ⇒* H ⬝ f ⬝ (H ⬝ f) := +theorem Y_def (f : SKI) : Y ⬝ f ↠ H ⬝ f ⬝ (H ⬝ f) := YPoly.toSKI_correct [f] (by simp) @@ -239,29 +239,29 @@ rather than up to a common reduct. An alternative is to use Turing's fixed-point (defined below). -/ def fixedPoint (f : SKI) : SKI := H ⬝ f ⬝ (H ⬝ f) -theorem fixedPoint_correct (f : SKI) : f.fixedPoint ⇒* f ⬝ f.fixedPoint := H_def f (H ⬝ f) +theorem fixedPoint_correct (f : SKI) : f.fixedPoint ↠ f ⬝ f.fixedPoint := H_def f (H ⬝ f) /-- Auxiliary definition for Turing's fixed-point combinator: ΘAux := λ x y. y (x x y) -/ def ΘAuxPoly : SKI.Polynomial 2 := &1 ⬝' (&0 ⬝' &0 ⬝' &1) /-- A term representing ΘAux -/ def ΘAux : SKI := ΘAuxPoly.toSKI -theorem ΘAux_def (x y : SKI) : ΘAux ⬝ x ⬝ y ⇒* y ⬝ (x ⬝ x ⬝ y) := +theorem ΘAux_def (x y : SKI) : ΘAux ⬝ x ⬝ y ↠ y ⬝ (x ⬝ x ⬝ y) := ΘAuxPoly.toSKI_correct [x, y] (by simp) /--Turing's fixed-point combinator: Θ := (λ x y. y (x x y)) (λ x y. y (x x y)) -/ def Θ : SKI := ΘAux ⬝ ΘAux /-- A SKI term representing Θ -/ -theorem Θ_correct (f : SKI) : Θ ⬝ f ⇒* f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f +theorem Θ_correct (f : SKI) : Θ ⬝ f ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f /-! ### Church Booleans -/ /-- A term a represents the boolean value u if it is βη-equivalent to a standard Church boolean. -/ def IsBool (u : Bool) (a : SKI) : Prop := - ∀ x y : SKI, a ⬝ x ⬝ y ⇒* (if u then x else y) + ∀ x y : SKI, a ⬝ x ⬝ y ↠ (if u then x else y) -theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ⇒* a') (ha' : IsBool u a') : +theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠ a') (ha' : IsBool u a') : IsBool u a := by intro x y trans a' ⬝ x ⬝ y @@ -278,13 +278,13 @@ theorem TT_correct : IsBool true TT := fun x y ↦ MRed.K x y def FF : SKI := K ⬝ I theorem FF_correct : IsBool false FF := fun x y ↦ calc - FF ⬝ x ⬝ y ⇒ I ⬝ y := by apply red_head; exact red_K I x - _ ⇒ y := red_I y + FF ⬝ x ⬝ y ⭢ I ⬝ y := by apply red_head; exact red_K I x + _ ⭢ y := red_I y /-- Conditional: Cond x y b := if b then x else y -/ protected def Cond : SKI := RotR theorem cond_correct (a x y : SKI) (u : Bool) (h : IsBool u a) : - SKI.Cond ⬝ x ⬝ y ⬝ a ⇒* if u then x else y := by + SKI.Cond ⬝ x ⬝ y ⬝ a ↠ if u then x else y := by trans a ⬝ x ⬝ y · exact rotR_def x y a · exact h x y @@ -302,7 +302,7 @@ theorem neg_correct (a : SKI) (ua : Bool) (h : IsBool ua a) : IsBool (¬ ua) (SK def AndPoly : SKI.Polynomial 2 := SKI.Cond ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' FF ⬝' &0 /-- A SKI term representing And -/ protected def And : SKI := AndPoly.toSKI -theorem and_def (a b : SKI) : SKI.And ⬝ a ⬝ b ⇒* SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := +theorem and_def (a b : SKI) : SKI.And ⬝ a ⬝ b ↠ SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := AndPoly.toSKI_correct [a, b] (by simp) theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -321,7 +321,7 @@ theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool u def OrPoly : SKI.Polynomial 2 := SKI.Cond ⬝' TT ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' &0 /-- A SKI term representing Or -/ protected def Or : SKI := OrPoly.toSKI -theorem or_def (a b : SKI) : SKI.Or ⬝ a ⬝ b ⇒* SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := +theorem or_def (a b : SKI) : SKI.Or ⬝ a ⬝ b ↠ SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := OrPoly.toSKI_correct [a, b] (by simp) theorem or_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -350,22 +350,22 @@ def Fst : SKI := R ⬝ TT /-- Second projection -/ def Snd : SKI := R ⬝ FF -theorem fst_correct (a b : SKI) : Fst ⬝ (MkPair ⬝ a ⬝ b) ⇒* a := by calc - _ ⇒* SKI.Cond ⬝ a ⬝ b ⬝ TT := R_def _ _ - _ ⇒* a := cond_correct TT a b true TT_correct +theorem fst_correct (a b : SKI) : Fst ⬝ (MkPair ⬝ a ⬝ b) ↠ a := by calc + _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ TT := R_def _ _ + _ ↠ a := cond_correct TT a b true TT_correct -theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ⇒* b := by calc - _ ⇒* SKI.Cond ⬝ a ⬝ b ⬝ FF := R_def _ _ - _ ⇒* b := cond_correct FF a b false FF_correct +theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ↠ b := by calc + _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ FF := R_def _ _ + _ ↠ b := cond_correct FF a b false FF_correct /-- Unpaired f ⟨x, y⟩ := f x y, cf `Nat.unparied`. -/ def UnpairedPoly : SKI.Polynomial 2 := &0 ⬝' (Fst ⬝' &1) ⬝' (Snd ⬝' &1) /-- A term representing Unpaired -/ protected def Unpaired : SKI := UnpairedPoly.toSKI -theorem unpaired_def (f p : SKI) : SKI.Unpaired ⬝ f ⬝ p ⇒* f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := +theorem unpaired_def (f p : SKI) : SKI.Unpaired ⬝ f ⬝ p ↠ f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := UnpairedPoly.toSKI_correct [f, p] (by simp) -theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y) ⇒* f ⬝ x ⬝ y := by +theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y) ↠ f ⬝ x ⬝ y := by trans f ⬝ (Fst ⬝ (MkPair ⬝ x ⬝ y)) ⬝ (Snd ⬝ (MkPair ⬝ x ⬝ y)) . exact unpaired_def f _ . apply parallel_mRed @@ -377,5 +377,5 @@ theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x def PairPoly : SKI.Polynomial 3 := MkPair ⬝' (&0 ⬝' &2) ⬝' (&1 ⬝' &2) /-- A SKI term representing Pair -/ protected def Pair : SKI := PairPoly.toSKI -theorem pair_def (f g x : SKI) : SKI.Pair ⬝ f ⬝ g ⬝ x ⇒* MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := +theorem pair_def (f g x : SKI) : SKI.Pair ⬝ f ⬝ g ⬝ x ↠ MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := PairPoly.toSKI_correct [f, g, x] (by simp) diff --git a/Cslib/Computability/CombinatoryLogic/Confluence.lean b/Cslib/Computability/CombinatoryLogic/Confluence.lean index 5c0fb508..4c426396 100644 --- a/Cslib/Computability/CombinatoryLogic/Confluence.lean +++ b/Cslib/Computability/CombinatoryLogic/Confluence.lean @@ -9,11 +9,11 @@ import Cslib.Utils.Relation /-! # SKI reduction is confluent -This file proves the **Church-Rosser** theorem for the SKI calculus, that is, if `a ⇒* b` and -`a ⇒* c`, `b ⇒* d` and `c ⇒* d` for some term `d`. More strongly (though equivalently), we show +This file proves the **Church-Rosser** theorem for the SKI calculus, that is, if `a ↠ b` and +`a ↠ c`, `b ↠ d` and `c ↠ d` for some term `d`. More strongly (though equivalently), we show that the relation of having a common reduct is transitive — in the above situation, `a` and `b`, and `a` and `c` have common reducts, so the result implies the same of `b` and `c`. Note that -`CommonReduct` is symmetric (trivially) and reflexive (since `⇒*` is), so we in fact show that +`CommonReduct` is symmetric (trivially) and reflexive (since `↠` is), so we in fact show that `CommonReduct` is an equivalence. Our proof @@ -23,7 +23,7 @@ Chapter 4 of Peter Selinger's notes: ## Main definitions -- `ParallelReduction` : a relation `⇒ₚ` on terms such that `⇒ ⊆ ⇒ₚ ⊆ ⇒*`, allowing simultaneous +- `ParallelReduction` : a relation `⇒ₚ` on terms such that `⇒ ⊆ ⇒ₚ ⊆ ↠`, allowing simultaneous reduction on the head and tail of a term. ## Main results @@ -31,13 +31,13 @@ reduction on the head and tail of a term. - `parallelReduction_diamond` : parallel reduction satisfies the diamond property, that is, it is confluent in a single step. - `commonReduct_equivalence` : by a general result, the diamond property for `⇒ₚ` implies the same -for its reflexive-transitive closure. This closure is exactly `⇒*`, which implies the +for its reflexive-transitive closure. This closure is exactly `↠`, which implies the **Church-Rosser** theorem as sketched above. -/ namespace SKI -open Red MRed +open Red MRed ReductionSystem /-- A reduction step allowing simultaneous reduction of disjoint redexes -/ inductive ParallelReduction : SKI → SKI → Prop @@ -55,8 +55,8 @@ inductive ParallelReduction : SKI → SKI → Prop /-- Notation for parallel reduction -/ scoped infix:90 " ⇒ₚ " => ParallelReduction -/-- The inclusion `⇒ₚ ⊆ ⇒*` -/ -theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ⇒* a' := by +/-- The inclusion `⇒ₚ ⊆ ↠` -/ +theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ↠ a' := by cases h case refl => exact Relation.ReflTransGen.refl case par a a' b b' ha hb => @@ -68,7 +68,7 @@ theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ⇒* a' := case red_S a b c => exact Relation.ReflTransGen.single (red_S a b c) /-- The inclusion `⇒ ⊆ ⇒ₚ` -/ -theorem parallelReduction_of_red {a a' : SKI} (h : a ⇒ a') : a ⇒ₚ a' := by +theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢ a') : a ⇒ₚ a' := by cases h case red_S => apply ParallelReduction.red_S case red_K => apply ParallelReduction.red_K @@ -86,12 +86,12 @@ theorem parallelReduction_of_red {a a' : SKI} (h : a ⇒ a') : a ⇒ₚ a' := by `parallelReduction_of_red` imply that `⇒` and `⇒ₚ` have the same reflexive-transitive closure. -/ theorem reflTransGen_parallelReduction_mRed : - Relation.ReflTransGen ParallelReduction = MRed := by + Relation.ReflTransGen ParallelReduction = RedSKI.MRed := by ext a b constructor · apply Relation.reflTransGen_minimal - · exact MRed.reflexive - · exact MRed.transitive + · exact λ _ => by rfl + · exact instTransitiveMRed RedSKI · exact @mRed_of_parallelReduction · apply Relation.reflTransGen_minimal · exact Relation.reflexive_reflTransGen @@ -101,7 +101,7 @@ theorem reflTransGen_parallelReduction_mRed : /-! Irreducibility for the (partially applied) primitive combinators. -TODO: possibly these should be proven more generally (in another file) for `⇒*`. +TODO: possibly these should be proven more generally (in another file) for `↠`. -/ lemma I_irreducible (a : SKI) (h : I ⇒ₚ a) : a = I := by @@ -233,7 +233,7 @@ theorem commonReduct_equivalence : Equivalence CommonReduct := by exact join_parallelReduction_equivalence /-- The **Church-Rosser** theorem in the form it is usually stated. -/ -theorem MRed.diamond (a b c : SKI) (hab : a ⇒* b) (hac : a ⇒* c) : CommonReduct b c := by +theorem MRed.diamond (a b c : SKI) (hab : a ↠ b) (hac : a ↠ c) : CommonReduct b c := by apply commonReduct_equivalence.trans (y := a) · exact commonReduct_equivalence.symm (commonReduct_of_single hab) · exact commonReduct_of_single hac diff --git a/Cslib/Computability/CombinatoryLogic/Defs.lean b/Cslib/Computability/CombinatoryLogic/Defs.lean index c0b993db..ca4c8b09 100644 --- a/Cslib/Computability/CombinatoryLogic/Defs.lean +++ b/Cslib/Computability/CombinatoryLogic/Defs.lean @@ -5,6 +5,7 @@ Authors: Thomas Waring -/ import Mathlib.Logic.Relation import Cslib.Utils.Relation +import Cslib.Semantics.ReductionSystem.Basic /-! # SKI Combinatory Logic @@ -23,7 +24,7 @@ using the SKI basis. - `⬝` : application between SKI terms, - `⇒` : single-step reduction, -- `⇒*` : multi-step reduction, +- `↠` : multi-step reduction, ## References @@ -60,6 +61,7 @@ lemma applyList_concat (f : SKI) (ys : List SKI) (z : SKI) : /-! ### Reduction relations between SKI terms -/ /-- Single-step reduction of SKI terms -/ +@[reduction_sys RedSKI] inductive Red : SKI → SKI → Prop where /-- The operational semantics of the `S`, -/ | red_S (x y z : SKI) : Red (S ⬝ x ⬝ y ⬝ z) (x ⬝ z ⬝ (y ⬝ z)) @@ -72,60 +74,47 @@ inductive Red : SKI → SKI → Prop where /-- and tail of an SKI term. -/ | red_tail (x y y' : SKI) (_ : Red y y') : Red (x ⬝ y) (x ⬝ y') -/-- Notation for single-step reduction -/ -scoped infix:90 " ⇒ " => Red -/-- Multi-step reduction of SKI terms -/ -def MRed : SKI → SKI → Prop := Relation.ReflTransGen Red +open Red ReductionSystem -/-- Notation for multi-step reduction (by analogy with the Kleene star) -/ -scoped infix:90 " ⇒* " => MRed +theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) := MRed.single RedSKI <| red_S .. +theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. +theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. -open Red - -@[refl] -theorem MRed.refl (a : SKI) : a ⇒* a := Relation.ReflTransGen.refl - -theorem MRed.single {a b : SKI} (h : a ⇒ b) : a ⇒* b := Relation.ReflTransGen.single h - -theorem MRed.S (x y z : SKI) : MRed (S ⬝ x ⬝ y ⬝ z) (x ⬝ z ⬝ (y ⬝ z)) := MRed.single <| red_S .. -theorem MRed.K (x y : SKI) : MRed (K ⬝ x ⬝ y) x := MRed.single <| red_K .. -theorem MRed.I (x : SKI) : MRed (I ⬝ x) x := MRed.single <| red_I .. - -theorem MRed.head {a a' : SKI} (b : SKI) (h : a ⇒* a') : (a ⬝ b) ⇒* (a' ⬝ b) := by +theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠ a') : (a ⬝ b) ↠ (a' ⬝ b) := by induction h with | refl => apply MRed.refl | @tail a' a'' _ ha'' ih => apply Relation.ReflTransGen.tail (b := a' ⬝ b) ih exact Red.red_head a' a'' b ha'' -theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ⇒* b') : (a ⬝ b) ⇒* (a ⬝ b') := by +theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠ b') : (a ⬝ b) ↠ (a ⬝ b') := by induction h with | refl => apply MRed.refl | @tail b' b'' _ hb'' ih => apply Relation.ReflTransGen.tail (b := a ⬝ b') ih exact Red.red_tail a b' b'' hb'' -instance MRed.instTrans : IsTrans SKI MRed := Relation.instIsTransReflTransGen -theorem MRed.transitive : Transitive MRed := transitive_of_trans MRed +-- instance MRed.instTrans : IsTrans SKI MRed := Relation.instIsTransReflTransGen +-- theorem MRed.transitive : Transitive MRed := transitive_of_trans MRed -instance MRed.instIsRefl : IsRefl SKI MRed := Relation.instIsReflReflTransGen -theorem MRed.reflexive : Reflexive MRed := IsRefl.reflexive +-- instance MRed.instIsRefl : IsRefl SKI MRed := Relation.instIsReflReflTransGen +-- theorem MRed.reflexive : Reflexive MRed := IsRefl.reflexive -instance MRedTrans : Trans Red MRed MRed := - ⟨fun hab => Relation.ReflTransGen.trans (MRed.single hab)⟩ +-- instance MRedTrans : Trans Red MRed MRed := +-- ⟨fun hab => Relation.ReflTransGen.trans (MRed.single hab)⟩ -instance MRedRedTrans : Trans MRed Red MRed := - ⟨fun hab hbc => Relation.ReflTransGen.trans hab (MRed.single hbc)⟩ +-- instance MRedRedTrans : Trans MRed Red MRed := +-- ⟨fun hab hbc => Relation.ReflTransGen.trans hab (MRed.single hbc)⟩ -instance RedMRedTrans : Trans Red Red MRed := - ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ +-- instance RedMRedTrans : Trans Red Red MRed := +-- ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ -lemma parallel_mRed {a a' b b' : SKI} (ha : a ⇒* a') (hb : b ⇒* b') : - (a ⬝ b) ⇒* (a' ⬝ b') := +lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠ a') (hb : b ↠ b') : + (a ⬝ b) ↠ (a' ⬝ b') := Trans.simple (MRed.head b ha) (MRed.tail a' hb) -lemma parallel_red {a a' b b' : SKI} (ha : a ⇒ a') (hb : b ⇒ b') : (a ⬝ b) ⇒* (a' ⬝ b') := by +lemma parallel_red {a a' b b' : SKI} (ha : a ⭢ a') (hb : b ⭢ b') : (a ⬝ b) ↠ (a' ⬝ b') := by trans a' ⬝ b all_goals apply MRed.single · exact Red.red_head a a' b ha @@ -133,11 +122,10 @@ lemma parallel_red {a a' b b' : SKI} (ha : a ⇒ a') (hb : b ⇒ b') : (a ⬝ b) /-- Express that two terms have a reduce to a common term. -/ -def CommonReduct : SKI → SKI → Prop := Relation.Join MRed +def CommonReduct : SKI → SKI → Prop := Relation.Join RedSKI.MRed -lemma commonReduct_of_single {a b : SKI} (h : a ⇒* b) : CommonReduct a b := by - refine Relation.join_of_single MRed.reflexive h +lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨b, h, by rfl⟩ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join -theorem reflexive_commonReduct : Reflexive CommonReduct := - Relation.reflexive_join MRed.reflexive +theorem reflexive_commonReduct : Reflexive CommonReduct := λ x => by + refine ⟨x,?_,?_⟩ <;> rfl diff --git a/Cslib/Computability/CombinatoryLogic/Recursion.lean b/Cslib/Computability/CombinatoryLogic/Recursion.lean index 9ef05ba6..18448030 100644 --- a/Cslib/Computability/CombinatoryLogic/Recursion.lean +++ b/Cslib/Computability/CombinatoryLogic/Recursion.lean @@ -15,16 +15,16 @@ formalisation of `Mathlib.Computability.Partrec`. Since composition (`B`-combina what remains are the following definitions and proofs of their correctness. - Church numerals : a predicate `IsChurch n a` expressing that the term `a` is βη-equivalent to -the standard church numeral `n` — that is, `a ⬝ f ⬝ x ⇒* f ⬝ (f ⬝ ... ⬝ (f ⬝ x)))`. +the standard church numeral `n` — that is, `a ⬝ f ⬝ x ↠ f ⬝ (f ⬝ ... ⬝ (f ⬝ x)))`. - SKI numerals : `Zero` and `Succ`, corresponding to `Partrec.zero` and `Partrec.succ`, and correctness proofs `zero_correct` and `succ_correct`. - Predecessor : a term `Pred` so that (`pred_correct`) `IsChurch n a → IsChurch n.pred (Pred ⬝ a)`. - Primitive recursion : a term `Rec` so that (`rec_correct_succ`) `IsChurch (n+1) a` implies -`Rec ⬝ x ⬝ g ⬝ a ⇒* g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` implies -`Rec ⬝ x ⬝ g ⬝ a ⇒* x`. +`Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` implies +`Rec ⬝ x ⬝ g ⬝ a ↠ x`. - Unbounded root finding (μ-recursion) : given a term `f` representing a function `fℕ: Nat → Nat`, -which takes on the value 0 a term `RFind` such that (`rFind_correct`) `RFind ⬝ f ⇒* a` such that +which takes on the value 0 a term `RFind` such that (`rFind_correct`) `RFind ⬝ f ↠ a` such that `IsChurch n a` for `n` the smallest root of `fℕ`. ## References @@ -49,7 +49,7 @@ sense of `Mathlib.Data.Part` (as used in `Mathlib.Computability.Partrec`). namespace SKI -open Red MRed +open Red MRed ReductionSystem /-- Function form of the church numerals. -/ def Church (n : Nat) (f x : SKI) : SKI := @@ -58,22 +58,22 @@ match n with | n+1 => f ⬝ (Church n f x) /-- `church` commutes with reduction. -/ -lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ⇒* f') (hx : x ⇒* x') : - Church n f x ⇒* Church n f' x' := by +lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠ f') (hx : x ↠ x') : + Church n f x ↠ Church n f' x' := by induction n with | zero => exact hx | succ n ih => exact parallel_mRed hf ih /-- The term `a` is βη-equivalent to a standard church numeral. -/ -def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x : SKI, a ⬝ f ⬝ x ⇒* Church n f x +def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x : SKI, a ⬝ f ⬝ x ↠ Church n f x /-- To show `IsChurch n a` it suffices to show the same for a reduct of `a`. -/ -theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ⇒* a') : IsChurch n a' → IsChurch n a := by +theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠ a') : IsChurch n a' → IsChurch n a := by simp_rw [IsChurch] intro ha' f x calc - _ ⇒* a' ⬝ f ⬝ x := by apply MRed.head; apply MRed.head; exact h - _ ⇒* Church n f x := by apply ha' + _ ↠ a' ⬝ f ⬝ x := by apply MRed.head; apply MRed.head; exact h + _ ↠ Church n f x := by apply ha' /-! ### Church numeral basics -/ @@ -84,24 +84,24 @@ theorem zero_correct : IsChurch 0 SKI.Zero := by unfold IsChurch SKI.Zero Church intro f x calc - _ ⇒ I ⬝ x := by apply red_head; apply red_K - _ ⇒ x := by apply red_I + _ ⭢ I ⬝ x := by apply red_head; apply red_K + _ ⭢ x := by apply red_I /-- Church one := λ f x. f x -/ protected def One : SKI := I theorem one_correct : IsChurch 1 SKI.One := by intro f x apply head - exact single (red_I f) + exact MRed.single RedSKI (red_I f) /-- Church succ := λ a f x. f (a f x) ~ λ a f. B f (a f) ~ λ a. S B a ~ S B -/ protected def Succ : SKI := S ⬝ B theorem succ_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch (n+1) (SKI.Succ ⬝ a) := by intro f x calc - _ ⇒ B ⬝ f ⬝ (a ⬝ f) ⬝ x := by apply red_head; apply red_S - _ ⇒* f ⬝ (a ⬝ f ⬝ x) := by apply B_def - _ ⇒* f ⬝ (Church n f x) := by apply MRed.tail; exact h f x + _ ⭢ B ⬝ f ⬝ (a ⬝ f) ⬝ x := by apply red_head; apply red_S + _ ↠ f ⬝ (a ⬝ f ⬝ x) := by apply B_def + _ ↠ f ⬝ (Church n f x) := by apply MRed.tail; exact h f x /-- To define the predecessor, iterate the function `PredAux` ⟨i, j⟩ ↦ ⟨j, j+1⟩ on ⟨0,0⟩, then take @@ -110,14 +110,14 @@ the first component. def PredAuxPoly : SKI.Polynomial 1 := MkPair ⬝' (Snd ⬝' &0) ⬝' (SKI.Succ ⬝' (Snd ⬝' &0)) /-- A term representing PredAux-/ def PredAux : SKI := PredAuxPoly.toSKI -theorem predAux_def (p : SKI) : PredAux ⬝ p ⇒* MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := +theorem predAux_def (p : SKI) : PredAux ⬝ p ↠ MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := PredAuxPoly.toSKI_correct [p] (by simp) /-- Useful auxiliary definition expressing that `p` represents ns ∈ Nat × Nat. -/ def IsChurchPair (ns : Nat × Nat) (x : SKI) : Prop := IsChurch ns.1 (Fst ⬝ x) ∧ IsChurch ns.2 (Snd ⬝ x) -theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ⇒* a') : +theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠ a') : IsChurchPair ns a' → IsChurchPair ns a := by simp_rw [IsChurchPair] intro ⟨ha₁,ha₂⟩ @@ -154,7 +154,7 @@ theorem predAux_correct' (n : Nat) : def PredPoly : SKI.Polynomial 1 := Fst ⬝' (&0 ⬝' PredAux ⬝' (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) /-- A term representing Pred -/ def Pred : SKI := PredPoly.toSKI -theorem pred_def (a : SKI) : Pred ⬝ a ⇒* Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := +theorem pred_def (a : SKI) : Pred ⬝ a ↠ Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := PredPoly.toSKI_correct [a] (by simp) theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (Pred ⬝ a) := by @@ -172,7 +172,7 @@ theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (P def IsZeroPoly : SKI.Polynomial 1 := &0 ⬝' (K ⬝ FF) ⬝' TT /-- A term representing IsZero -/ def IsZero : SKI := IsZeroPoly.toSKI -theorem isZero_def (a : SKI) : IsZero ⬝ a ⇒* a ⬝ (K ⬝ FF) ⬝ TT := +theorem isZero_def (a : SKI) : IsZero ⬝ a ↠ a ⬝ (K ⬝ FF) ⬝ TT := IsZeroPoly.toSKI_correct [a] (by simp) theorem isZero_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsBool (n = 0) (IsZero ⬝ a) := by @@ -189,8 +189,8 @@ theorem isZero_correct (n : Nat) (a : SKI) (h : IsChurch n a) : rw [hk] at h apply isBool_trans (ha' := FF_correct) calc - _ ⇒* (K ⬝ FF) ⬝ Church k (K ⬝ FF) TT := h _ _ - _ ⇒ FF := red_K _ _ + _ ↠ (K ⬝ FF) ⬝ Church k (K ⬝ FF) TT := h _ _ + _ ⭢ FF := red_K _ _ /-- @@ -202,33 +202,33 @@ def RecAuxPoly : SKI.Polynomial 4 := /-- A term representing RecAux -/ def RecAux : SKI := RecAuxPoly.toSKI theorem recAux_def (R₀ x g a : SKI) : - RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := + RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := RecAuxPoly.toSKI_correct [R₀, x, g, a] (by simp) /-- We define Rec so that -`Rec ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a)` +`Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a)` -/ def Rec : SKI := fixedPoint RecAux theorem rec_def (x g a : SKI) : - Rec ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc - _ ⇒* RecAux ⬝ Rec ⬝ x ⬝ g ⬝ a := by + Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc + _ ↠ RecAux ⬝ Rec ⬝ x ⬝ g ⬝ a := by apply MRed.head; apply MRed.head; apply MRed.head apply fixedPoint_correct - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := recAux_def Rec x g a + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := recAux_def Rec x g a -theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ⇒* x := by +theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ↠ x := by calc - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ - _ ⇒* if (Nat.beq 0 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ + _ ↠ if (Nat.beq 0 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by apply cond_correct exact isZero_correct 0 a ha theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n+1) a) : - Rec ⬝ x ⬝ g ⬝ a ⇒* g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by + Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by calc - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ - _ ⇒* if (Nat.beq (n+1) 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ + _ ↠ if (Nat.beq (n+1) 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by apply cond_correct exact isZero_correct (n+1) a ha @@ -245,19 +245,19 @@ def RFindAboveAuxPoly : SKI.Polynomial 3 := /-- A term representing RFindAboveAux -/ def RFindAboveAux : SKI := RFindAboveAuxPoly.toSKI lemma rfindAboveAux_def (R₀ f a : SKI) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := RFindAboveAuxPoly.toSKI_correct [R₀, a, f] (by trivial) theorem rfindAboveAux_base (R₀ f a : SKI) (hfa : IsChurch 0 (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* a := calc - _ ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ - _ ⇒* if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ a := calc + _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ + _ ↠ if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m+1) (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc - _ ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ - _ ⇒* if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc + _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ + _ ↠ if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa @@ -314,15 +314,15 @@ theorem RFind_correct (fNat : Nat → Nat) (f : SKI) def AddPoly : SKI.Polynomial 2 := &0 ⬝' SKI.Succ ⬝' &1 /-- A term representing addition on church numerals -/ protected def Add : SKI := AddPoly.toSKI -theorem add_def (a b : SKI) : SKI.Add ⬝ a ⬝ b ⇒* a ⬝ SKI.Succ ⬝ b := +theorem add_def (a b : SKI) : SKI.Add ⬝ a ⬝ b ↠ a ⬝ SKI.Succ ⬝ b := AddPoly.toSKI_correct [a, b] (by simp) theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : IsChurch (n+m) (SKI.Add ⬝ a ⬝ b) := by refine isChurch_trans (n+m) (a' := Church n SKI.Succ b) ?_ ?_ · calc - _ ⇒* a ⬝ SKI.Succ ⬝ b := add_def a b - _ ⇒* Church n SKI.Succ b := ha SKI.Succ b + _ ↠ a ⬝ SKI.Succ ⬝ b := add_def a b + _ ↠ Church n SKI.Succ b := ha SKI.Succ b · clear ha induction n with | zero => simp_rw [Nat.zero_add, Church]; exact hb @@ -334,7 +334,7 @@ theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def MulPoly : SKI.Polynomial 2 := &0 ⬝' (SKI.Add ⬝' &1) ⬝' SKI.Zero /-- A term representing multiplication on church numerals -/ protected def Mul : SKI := MulPoly.toSKI -theorem mul_def (a b : SKI) : SKI.Mul ⬝ a ⬝ b ⇒* a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := +theorem mul_def (a b : SKI) : SKI.Mul ⬝ a ⬝ b ↠ a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := MulPoly.toSKI_correct [a, b] (by simp) theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m b) : @@ -352,15 +352,15 @@ theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m def SubPoly : SKI.Polynomial 2 := &1 ⬝' Pred ⬝' &0 /-- A term representing subtraction on church numerals -/ protected def Sub : SKI := SubPoly.toSKI -theorem sub_def (a b : SKI) : SKI.Sub ⬝ a ⬝ b ⇒* b ⬝ Pred ⬝ a := +theorem sub_def (a b : SKI) : SKI.Sub ⬝ a ⬝ b ↠ b ⬝ Pred ⬝ a := SubPoly.toSKI_correct [a, b] (by simp) theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : IsChurch (n-m) (SKI.Sub ⬝ a ⬝ b) := by refine isChurch_trans (n-m) (a' := Church m Pred a) ?_ ?_ · calc - _ ⇒* b ⬝ Pred ⬝ a := sub_def a b - _ ⇒* Church m Pred a := hb Pred a + _ ↠ b ⬝ Pred ⬝ a := sub_def a b + _ ↠ Church m Pred a := hb Pred a · clear hb induction m with | zero => simp_rw [Nat.sub_zero, Church]; exact ha @@ -372,7 +372,7 @@ theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def LEPoly : SKI.Polynomial 2 := IsZero ⬝' (SKI.Sub ⬝' &0 ⬝' &1) /-- A term representing comparison on church numerals -/ protected def LE : SKI := LEPoly.toSKI -theorem le_def (a b : SKI) : SKI.LE ⬝ a ⬝ b ⇒* IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := +theorem le_def (a b : SKI) : SKI.LE ⬝ a ⬝ b ↠ IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := LEPoly.toSKI_correct [a, b] (by simp) theorem le_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : diff --git a/Cslib/Semantics/ReductionSystem/Basic.lean b/Cslib/Semantics/ReductionSystem/Basic.lean index f5c277d1..00e7aac0 100644 --- a/Cslib/Semantics/ReductionSystem/Basic.lean +++ b/Cslib/Semantics/ReductionSystem/Basic.lean @@ -52,15 +52,19 @@ instance {α} (R : α → α → Prop) : Trans R (ReflTransGen R) (ReflTransGen instance {α} (R : α → α → Prop) : Trans (ReflTransGen R) R (ReflTransGen R) where trans := tail -instance (rs : ReductionSystem Term) : Trans rs.Red rs.Red rs.MRed := by infer_instance +-- instance (rs : ReductionSystem Term) : Trans rs.Red rs.Red rs.MRed := by infer_instance instance (rs : ReductionSystem Term) : Trans rs.Red rs.MRed rs.MRed := by infer_instance instance (rs : ReductionSystem Term) : Trans rs.MRed rs.Red rs.MRed := by infer_instance +instance (rs : ReductionSystem Term) : IsTrans Term rs.MRed := by infer_instance +instance (rs : ReductionSystem Term) : Transitive rs.MRed := transitive_of_trans rs.MRed +instance (rs : ReductionSystem Term) : Trans rs.MRed rs.MRed rs.MRed := instTransOfIsTrans + end MultiStep open Lean Elab Meta Command Term --- thank you to Kyle Miller for this: +-- thank you to Kyle Miller for this: -- https://leanprover.zulipchat.com/#narrow/channel/239415-metaprogramming-.2F-tactics/topic/Working.20with.20variables.20in.20a.20command/near/529324084 /-- A command to create a `ReductionSystem` from a relation, robust to use of `variable `-/ @@ -91,10 +95,10 @@ elab "create_reduction_sys" rel:ident name:ident : command => do } addTermInfo' name (.const name.getId params) (isBinder := true) addDeclarationRangesFromSyntax name.getId name - -/-- + +/-- This command adds notations for a `ReductionSystem.Red` and `ReductionSystem.MRed`. This should - not usually be called directly, but from the `reduction_sys` attribute. + not usually be called directly, but from the `reduction_sys` attribute. As an example `reduction_notation foo "β"` will add the notations "⭢β" and "↠β". @@ -104,19 +108,19 @@ elab "create_reduction_sys" rel:ident name:ident : command => do -/ syntax "reduction_notation" ident (Lean.Parser.Command.notationItem)? : command macro_rules - | `(reduction_notation $rs $sym) => + | `(reduction_notation $rs $sym) => `( notation:39 t " ⭢"$sym t' => (ReductionSystem.Red $rs) t t' notation:39 t " ↠"$sym t' => (ReductionSystem.MRed $rs) t t' ) - | `(reduction_notation $rs) => + | `(reduction_notation $rs) => `( notation:39 t " ⭢" t' => (ReductionSystem.Red $rs) t t' notation:39 t " ↠" t' => (ReductionSystem.MRed $rs) t t' ) -/-- +/-- This attribute calls the `reduction_notation` command for the annotated declaration, such as in: ``` From 62b1776d1089fcf9e7f1ac3611fd069b72068980 Mon Sep 17 00:00:00 2001 From: twwar Date: Tue, 22 Jul 2025 12:58:25 +0200 Subject: [PATCH 02/19] evaluation results --- .../Computability/CombinatoryLogic/Defs.lean | 7 + .../CombinatoryLogic/Evaluation.lean | 179 ++++++++++++++++++ 2 files changed, 186 insertions(+) create mode 100644 Cslib/Computability/CombinatoryLogic/Evaluation.lean diff --git a/Cslib/Computability/CombinatoryLogic/Defs.lean b/Cslib/Computability/CombinatoryLogic/Defs.lean index ca4c8b09..168a92f0 100644 --- a/Cslib/Computability/CombinatoryLogic/Defs.lean +++ b/Cslib/Computability/CombinatoryLogic/Defs.lean @@ -77,6 +77,13 @@ inductive Red : SKI → SKI → Prop where open Red ReductionSystem +lemma Red.ne {x y : SKI} : (x ⭢ y) → x ≠ y + | red_S _ _ _, h => by cases h + | red_K _ _, h => by cases h + | red_I _, h => by cases h + | red_head _ _ _ h', h => Red.ne h' (SKI.app.inj h).1 + | red_tail _ _ _ h', h => Red.ne h' (SKI.app.inj h).2 + theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) := MRed.single RedSKI <| red_S .. theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. diff --git a/Cslib/Computability/CombinatoryLogic/Evaluation.lean b/Cslib/Computability/CombinatoryLogic/Evaluation.lean new file mode 100644 index 00000000..7c73c3a9 --- /dev/null +++ b/Cslib/Computability/CombinatoryLogic/Evaluation.lean @@ -0,0 +1,179 @@ +/- +Copyright (c) 2025 Thomas Waring. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Waring +-/ +import Cslib.Computability.CombinatoryLogic.Defs +import Cslib.Computability.CombinatoryLogic.Confluence + +/-! +# Evaluation results + +This file draws heavily from +-/ + +open SKI Red + +/-- The predicate that a term has no reducible sub-terms. -/ +def RedexFree : SKI → Prop + | S => True + | K => True + | I => True + | S ⬝ x => RedexFree x + | K ⬝ x => RedexFree x + | I ⬝ _ => False + | S ⬝ x ⬝ y => RedexFree x ∧ RedexFree y + | K ⬝ _ ⬝ _ => False + | I ⬝ _ ⬝ _ => False + | S ⬝ _ ⬝ _ ⬝ _ => False + | K ⬝ _ ⬝ _ ⬝ _ => False + | I ⬝ _ ⬝ _ ⬝ _ => False + | a ⬝ b ⬝ c ⬝ d ⬝ e => RedexFree (a ⬝ b ⬝ c ⬝ d) ∧ RedexFree e + +/-- +One-step evaluation as a function: either it returns a term that has been reduced by one step, +or a proof that the term is redex free. Uses normal-order reduction. +-/ +def evalStep : (x : SKI) → PLift (RedexFree x) ⊕ SKI + | S => Sum.inl (PLift.up trivial) + | K => Sum.inl (PLift.up trivial) + | I => Sum.inl (PLift.up trivial) + | S ⬝ x => match evalStep x with + | Sum.inl h => Sum.inl h + | Sum.inr x' => Sum.inr (S ⬝ x') + | K ⬝ x => match evalStep x with + | Sum.inl h => Sum.inl h + | Sum.inr x' => Sum.inr (K ⬝ x') + | I ⬝ x => Sum.inr x + | S ⬝ x ⬝ y => match evalStep x, evalStep y with + | Sum.inl h1, Sum.inl h2 => Sum.inl (.up ⟨h1.down, h2.down⟩) + | Sum.inl _, Sum.inr y' => Sum.inr (S ⬝ x ⬝ y') + | Sum.inr x', _ => Sum.inr (S ⬝ x' ⬝ y) + | K ⬝ x ⬝ _ => Sum.inr x + | I ⬝ x ⬝ y => Sum.inr (x ⬝ y) + | S ⬝ x ⬝ y ⬝ z => Sum.inr (x ⬝ z ⬝ (y ⬝ z)) + | K ⬝ x ⬝ _ ⬝ z => Sum.inr (x ⬝ z) + | I ⬝ x ⬝ y ⬝ z => Sum.inr (x ⬝ y ⬝ z) + | a ⬝ b ⬝ c ⬝ d ⬝ e => match evalStep (a ⬝ b ⬝ c ⬝ d), evalStep e with + | Sum.inl h1, Sum.inl h2 => Sum.inl (.up ⟨h1.down, h2.down⟩) + | Sum.inl _, Sum.inr e' => Sum.inr (a ⬝ b ⬝ c ⬝ d ⬝ e') + | Sum.inr abcd', _ => Sum.inr (abcd' ⬝ e) + +/-- The normal-order reduction implemented by `evalStep` indeed computes a one-step reduction. -/ +theorem evalStep_right_correct : (x y : SKI) → (evalStep x = Sum.inr y) → x ⭢ y + | S ⬝ x, a, h => + match hx : evalStep x with + | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h + | Sum.inr x' => by + simp only [evalStep, hx, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) + | K ⬝ x, a, h => + match hx : evalStep x with + | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h + | Sum.inr x' => by + simp only [evalStep, hx, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) + | I ⬝ x, a, h => Sum.inr.inj h ▸ red_I _ + | S ⬝ x ⬝ y, a, h => + match hx : evalStep x, hy : evalStep y with + | Sum.inl _, Sum.inl _ => by simp only [hx, hy, evalStep, reduceCtorEq] at h + | Sum.inl _, Sum.inr y' => by + simp only [hx, hy, evalStep, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ <| evalStep_right_correct _ _ hy + | Sum.inr x', _ => by + simp only [hx, hy, evalStep, Sum.inr.injEq] at h + rw [←h] + exact .red_head _ _ _ <| .red_tail _ _ _ <| evalStep_right_correct _ _ hx + | K ⬝ x ⬝ y, a, h => Sum.inr.inj h ▸ red_K .. + | I ⬝ x ⬝ y, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_I _) + | S ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ red_S .. + | K ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_K ..) + | I ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_head _ _ _ <| red_I _) + | a ⬝ b ⬝ c ⬝ d ⬝ e, x, h => + match habcd : evalStep (a ⬝ b ⬝ c ⬝ d), he : evalStep e with + | Sum.inl _, Sum.inl _ => by simp only [habcd, he, evalStep, reduceCtorEq] at h + | Sum.inl _, Sum.inr e' => by + simp only [habcd, he, evalStep, Sum.inr.injEq] at h + rw [←h] + exact red_tail _ _ _ <| evalStep_right_correct _ _ he + | Sum.inr abcd', _ => by + simp only [habcd, he, evalStep, Sum.inr.injEq] at h + rw [←h] + exact red_head _ _ _ <| evalStep_right_correct _ _ habcd + +theorem redexFree_of_no_red {x : SKI} (h : ∀ y, ¬ (x ⭢ y)) : RedexFree x := by + match hx : evalStep x with + | Sum.inl h' => exact h'.down + | Sum.inr y => cases h _ (evalStep_right_correct x y hx) + +theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) +| S ⬝ x, hx, S ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' +| K ⬝ x, hx, K ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' +| S ⬝ _ ⬝ _, ⟨hx, _⟩, S ⬝ _ ⬝ _, red_head _ _ _ (red_tail _ _ _ h3) => hx.no_red _ h3 +| S ⬝ _ ⬝ _, ⟨_, hy⟩, S ⬝ _ ⬝ _, red_tail _ _ _ h3 => hy.no_red _ h3 +| _ ⬝ _ ⬝ _ ⬝ _ ⬝ _, ⟨hx, _⟩, _ ⬝ _, red_head _ _ _ hq => hx.no_red _ hq +| _ ⬝ _ ⬝ _ ⬝ _ ⬝ _, ⟨_, hy⟩, _ ⬝ _, red_tail _ _ _ he => hy.no_red _ he + +/-- A term is redex free iff it has no one-step reductions. -/ +theorem redexFree_iff {x : SKI} : RedexFree x ↔ ∀ y, ¬ (x ⭢ y) := + ⟨RedexFree.no_red, redexFree_of_no_red⟩ + +theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by + constructor + case mp => + intro h + match hx : evalStep x with + | Sum.inl h' => exact rfl + | Sum.inr y => cases h.no_red _ (evalStep_right_correct _ _ hx) + case mpr => + intro h + match hx : evalStep x with + | Sum.inl h' => exact h'.down + | Sum.inr y => rw [hx] at h; cases h + +instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_onceEval + +/-- A term is redex free iff its only many-step reduction is itself. -/ +theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := by + constructor + case mp => + intro h y + constructor + case mp => + intro h' + cases h'.cases_head + case inl => assumption + case inr h' => + obtain ⟨z, hz, _⟩ := h' + cases h.no_red _ hz + case mpr => + intro h + rw [h] + case mpr => + intro h + rw [redexFree_iff] + intro y hy + specialize h y + exact Red.ne hy (h.1 (Relation.ReflTransGen.single hy)) + +/-- If a term has a common reduct with a normal term, it in fact reduces to that term. -/ +theorem commonReduct_redexFree {x y : SKI} (hy : RedexFree y) (h : CommonReduct x y) : x ↠ y := + let ⟨w, hyw, hzw⟩ := h + (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw + +/-- If `x` reduces to both `y` and `z`, and `z` is not reducible, then `y` reduces to `z`. -/ +lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : RedexFree z) : y ↠ z := + let ⟨w, hyw, hzw⟩ := MRed.diamond x y z hxy hxz + (redexFree_iff'.1 hz _ |>.1 hzw : z = w) ▸ hyw + +/-- +If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. +-/ +lemma unique_normal_form {x y z : SKI} + (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := + (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) + +#check From 86c8407487c56d559b0b64a66243fd09f184447e Mon Sep 17 00:00:00 2001 From: twwar Date: Fri, 5 Sep 2025 16:46:47 +0200 Subject: [PATCH 03/19] rice's theorem --- .../Computability/CombinatoryLogic/Defs.lean | 12 ++ .../CombinatoryLogic/Evaluation.lean | 115 +++++++++++++++++- 2 files changed, 126 insertions(+), 1 deletion(-) diff --git a/Cslib/Computability/CombinatoryLogic/Defs.lean b/Cslib/Computability/CombinatoryLogic/Defs.lean index 168a92f0..1199f83e 100644 --- a/Cslib/Computability/CombinatoryLogic/Defs.lean +++ b/Cslib/Computability/CombinatoryLogic/Defs.lean @@ -57,6 +57,12 @@ lemma applyList_concat (f : SKI) (ys : List SKI) (z : SKI) : f.applyList (ys ++ [z]) = f.applyList ys ⬝ z := by simp [applyList] +/-- The size of an SKI term is its number of combinators. -/ +def size : SKI → Nat + | S => 1 + | K => 1 + | I => 1 + | x ⬝ y => size x + size y /-! ### Reduction relations between SKI terms -/ @@ -136,3 +142,9 @@ lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join theorem reflexive_commonReduct : Reflexive CommonReduct := λ x => by refine ⟨x,?_,?_⟩ <;> rfl + +theorem commonReduct_head {x x' : SKI} (y : SKI) : CommonReduct x x' → CommonReduct (x ⬝ y) (x' ⬝ y) + | ⟨z, hz, hz'⟩ => ⟨z ⬝ y, MRed.head y hz, MRed.head y hz'⟩ + +theorem commonReduct_tail (x : SKI) {y y' : SKI} : CommonReduct y y' → CommonReduct (x ⬝ y) (x ⬝ y') + | ⟨z, hz, hz'⟩ => ⟨x ⬝ z, MRed.tail x hz, MRed.tail x hz'⟩ diff --git a/Cslib/Computability/CombinatoryLogic/Evaluation.lean b/Cslib/Computability/CombinatoryLogic/Evaluation.lean index 7c73c3a9..6e7bc7f8 100644 --- a/Cslib/Computability/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Computability/CombinatoryLogic/Evaluation.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Waring -/ import Cslib.Computability.CombinatoryLogic.Defs +import Cslib.Computability.CombinatoryLogic.Basic import Cslib.Computability.CombinatoryLogic.Confluence +import Cslib.Computability.CombinatoryLogic.Recursion +import Mathlib.Tactic.Common /-! # Evaluation results @@ -176,4 +179,114 @@ lemma unique_normal_form {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) -#check +lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) + (hx : RedexFree x) (hy : RedexFree y) : x = y := + (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) + +/-! ### Injectivity for datatypes -/ + +lemma sk_nequiv : ¬ CommonReduct S K := by + intro ⟨z, hsz, hkz⟩ + have hS : RedexFree S := by simp [RedexFree] + have hK : RedexFree K := by simp [RedexFree] + cases (redexFree_iff'.1 hS z).1 hsz + cases (redexFree_iff'.1 hK _).1 hkz + +/-- Injectivity for booleans. -/ +theorem isBool_injective (x y : SKI) (u v : Bool) (hx : IsBool u x) (hy : IsBool v y) + (hxy : CommonReduct x y) : u = v := by + have h : CommonReduct (if u then S else K) (if v then S else K) := by + apply commonReduct_equivalence.trans (y := x ⬝ S ⬝ K) + · apply commonReduct_equivalence.symm + apply commonReduct_of_single + exact hx S K + · apply commonReduct_equivalence.trans (y := y ⬝ S ⬝ K) + · exact commonReduct_head K <| commonReduct_head S hxy + · apply commonReduct_of_single + exact hy S K + by_cases u + case pos hu => + by_cases v + case pos hv => + rw [hu, hv] + case neg hv => + simp_rw [hu, hv, Bool.false_eq_true, reduceIte] at h + exact False.elim <| sk_nequiv h + case neg hu => + by_cases v + case pos hv => + simp_rw [hu, hv, Bool.false_eq_true, reduceIte] at h + exact False.elim <| sk_nequiv (commonReduct_equivalence.symm h) + case neg hv => + simp_rw [hu, hv] + +lemma TF_nequiv : ¬ CommonReduct TT FF := fun h => + (Bool.eq_not_self true).mp <| isBool_injective TT FF true false TT_correct FF_correct h + +/-- A specialisation of `Church : Nat → SKI`. -/ +def churchK : Nat → SKI + | 0 => K + | n+1 => K ⬝ (churchK n) + +lemma churchK_church : (n : Nat) → churchK n = Church n K K + | 0 => rfl + | n+1 => by simp [churchK, Church, churchK_church n] + +lemma churchK_redexFree : (n : Nat) → RedexFree (churchK n) + | 0 => trivial + | n+1 => churchK_redexFree n + +@[simp] +lemma churchK_size : (n : Nat) → (churchK n).size = n+1 + | 0 => rfl + | n + 1 => by rw [churchK, size, size, churchK_size, Nat.add_comm] + +lemma churchK_injective : Function.Injective churchK := + fun n m h => by simpa using congrArg SKI.size h + +theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsChurch m y) + (hxy : CommonReduct x y) : n = m := by + suffices CommonReduct (churchK n) (churchK m) by + apply churchK_injective + exact unique_normal_form' this (churchK_redexFree n) (churchK_redexFree m) + apply commonReduct_equivalence.trans (y := x ⬝ K ⬝ K) + · simp_rw [churchK_church] + exact commonReduct_equivalence.symm <| commonReduct_of_single (hx K K) + · apply commonReduct_equivalence.trans (y := y ⬝ K ⬝ K) + · apply commonReduct_head; apply commonReduct_head; assumption + · simp_rw [churchK_church] + exact commonReduct_of_single (hy K K) + +/-- **Rice's theorem**: no SKI term is a non-trivial predicate. -/ +theorem rice {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) + (hxt : ∃ x : SKI, P ⬝ x ↠ TT) (hxf : ∃ x : SKI, P ⬝ x ↠ FF) : False := by + obtain ⟨a, ha⟩ := hxt + obtain ⟨b, hb⟩ := hxf + let Neg : SKI := P ⬝' &0 ⬝' b ⬝' a |>.toSKI (n := 1) + let Abs : SKI := Neg.fixedPoint + have Neg_app : ∀ x : SKI, Neg ⬝ x ↠ P ⬝ x ⬝ b ⬝ a := + fun x => (P ⬝' &0 ⬝' b ⬝' a) |>.toSKI_correct (n := 1) [x] (by simp) + cases hP Abs + case inl h => + have : P ⬝ Abs ↠ FF := calc + _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct + _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app + _ ↠ P ⬝ (TT ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h + _ ↠ P ⬝ b := by apply MRed.tail; apply TT_correct + _ ↠ FF := hb + exact TF_nequiv <| MRed.diamond _ _ _ h this + case inr h => + have : P ⬝ Abs ↠ TT := calc + _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct + _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app + _ ↠ P ⬝ (FF ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h + _ ↠ P ⬝ a := by apply MRed.tail; apply FF_correct + _ ↠ TT := ha + exact TF_nequiv <| MRed.diamond _ _ _ this h + +/-- **Rice's theorem**: any SKI predicate is trivial. -/ +theorem rice' {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) : + (∀ x : SKI, P ⬝ x ↠ TT) ∨ (∀ x : SKI, P ⬝ x ↠ FF) := by + by_contra! h + obtain ⟨⟨a, ha⟩, b, hb⟩ := h + exact rice hP ⟨b, (hP _).resolve_right hb⟩ ⟨a, (hP _).resolve_left ha⟩ From 7df191082d000b60e0a1733da81a14e85649cd41 Mon Sep 17 00:00:00 2001 From: twwar Date: Fri, 5 Sep 2025 16:58:50 +0200 Subject: [PATCH 04/19] documentation --- .../CombinatoryLogic/Evaluation.lean | 24 +++++++++++++++++-- 1 file changed, 22 insertions(+), 2 deletions(-) diff --git a/Cslib/Computability/CombinatoryLogic/Evaluation.lean b/Cslib/Computability/CombinatoryLogic/Evaluation.lean index 6e7bc7f8..4a232f2c 100644 --- a/Cslib/Computability/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Computability/CombinatoryLogic/Evaluation.lean @@ -12,6 +12,24 @@ import Mathlib.Tactic.Common /-! # Evaluation results +This file formalises evaluation and normal forms of SKI terms. + +## Main definitions + +- `RedexFree` : a predicate expressing that a term has no redexes +- `evalStep` : SKI-reduction as a function + +## Main results + +- `evalStep_right_correct` : correctness for `evalStep` +- `redexFree_iff` and `redexFree_iff'` : a term is redex free if and only if it has (respectively) +no one-step reductions, or if its only many step reduction is itself. +- `unique_normal_form` : if `x` reduces to both `y` and `z`, and both `y` and `z` are in normal +form, then they are equal. +- **Rice's theorem**: no SKI term is a non-trivial predicate. + +## References + This file draws heavily from -/ @@ -124,7 +142,7 @@ theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) theorem redexFree_iff {x : SKI} : RedexFree x ↔ ∀ y, ¬ (x ⭢ y) := ⟨RedexFree.no_red, redexFree_of_no_red⟩ -theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by +theorem redexFree_iff_evalStep {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by constructor case mp => intro h @@ -137,7 +155,7 @@ theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = | Sum.inl h' => exact h'.down | Sum.inr y => rw [hx] at h; cases h -instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_onceEval +instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_evalStep /-- A term is redex free iff its only many-step reduction is itself. -/ theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := by @@ -179,6 +197,7 @@ lemma unique_normal_form {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) +/-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) (hx : RedexFree x) (hy : RedexFree y) : x = y := (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) @@ -244,6 +263,7 @@ lemma churchK_size : (n : Nat) → (churchK n).size = n+1 lemma churchK_injective : Function.Injective churchK := fun n m h => by simpa using congrArg SKI.size h +/-- Injectivity for Church numerals-/ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsChurch m y) (hxy : CommonReduct x y) : n = m := by suffices CommonReduct (churchK n) (churchK m) by From 5b2ba412c4e8e5867ddec142b2005a46fd9e6eb3 Mon Sep 17 00:00:00 2001 From: Clark Barrett Date: Fri, 5 Sep 2025 16:11:03 -0700 Subject: [PATCH 05/19] Added Boole directory --- Cslib/Languages/Boole/README.md | 1 + 1 file changed, 1 insertion(+) create mode 100644 Cslib/Languages/Boole/README.md diff --git a/Cslib/Languages/Boole/README.md b/Cslib/Languages/Boole/README.md new file mode 100644 index 00000000..78e4c7aa --- /dev/null +++ b/Cslib/Languages/Boole/README.md @@ -0,0 +1 @@ +Placeholder for the Boole language. From 33c03a9cbef437b048349c29703fb15c29292350 Mon Sep 17 00:00:00 2001 From: Fabrizio Montesi Date: Wed, 3 Sep 2025 14:07:07 +0200 Subject: [PATCH 06/19] Fix CLL exchange, add HasSize, add cut' --- Cslib/Logics/LinearLogic/CLL/Basic.lean | 5 +++ .../LinearLogic/CLL/CutElimination.lean | 31 +++++++++++++++++-- 2 files changed, 34 insertions(+), 2 deletions(-) diff --git a/Cslib/Logics/LinearLogic/CLL/Basic.lean b/Cslib/Logics/LinearLogic/CLL/Basic.lean index fa016912..8ee52020 100644 --- a/Cslib/Logics/LinearLogic/CLL/Basic.lean +++ b/Cslib/Logics/LinearLogic/CLL/Basic.lean @@ -161,9 +161,14 @@ inductive Proof : Sequent Atom → Prop where scoped notation "⊢" Γ:90 => Proof Γ +/-- The axiom, but where the order of propositions is reversed. -/ theorem Proof.ax' {a : Proposition Atom} : Proof [a⫠, a] := Proof.exchange (List.Perm.swap ..) Proof.ax +/-- Cut, but where the premises are reversed. -/ +theorem Proof.cut' (p : ⊢(a⫠ :: Γ)) (q : ⊢(a :: Δ)) : ⊢(Γ ++ Δ) := by + grind [Proof.cut, Proposition.dual.involution] + section LogicalEquiv /-! ## Logical equivalences -/ diff --git a/Cslib/Logics/LinearLogic/CLL/CutElimination.lean b/Cslib/Logics/LinearLogic/CLL/CutElimination.lean index dccd2786..bb666574 100644 --- a/Cslib/Logics/LinearLogic/CLL/CutElimination.lean +++ b/Cslib/Logics/LinearLogic/CLL/CutElimination.lean @@ -18,7 +18,7 @@ inductive Proof.CutFree : {Γ : Sequent Atom} → ⊢Γ → Prop where | ax : Proof.ax.CutFree | one : Proof.one.CutFree | bot (p : ⊢Γ) (hp : CutFree p) : p.bot.CutFree - | exchange (hperm : Γ.Perm Δ) (p : ⊢Γ) : (Proof.exchange hperm p).CutFree + | exchange (hperm : Γ.Perm Δ) (p : ⊢Γ) : p.CutFree → (Proof.exchange hperm p).CutFree | parr (p : ⊢(a :: b :: Γ)) : p.CutFree → p.parr.CutFree | tensor (p : ⊢(a :: Γ)) (q : ⊢(b :: Δ)) : p.CutFree → q.CutFree → (Proof.tensor p q).CutFree @@ -33,9 +33,36 @@ inductive Proof.CutFree : {Γ : Sequent Atom} → ⊢Γ → Prop where | bang (hqs: Sequent.allQuest Γ) (p : ⊢(a :: Γ)) : p.CutFree → (p.bang hqs).CutFree -- No rule for cut. +/-- Size of a `Proof`. -/ +inductive Proof.HasSize : {Γ : Sequent Atom} → ⊢Γ → Nat → Prop where + | ax : Proof.ax.HasSize 1 + | one : Proof.one.HasSize 1 + | bot (p : ⊢Γ) (n : Nat) (hp : p.HasSize n) : p.bot.HasSize (n + 1) + | exchange (hperm : Γ.Perm Δ) (p : ⊢Γ) (n : Nat) (hp : p.HasSize n) : + (Proof.exchange hperm p).HasSize n + | parr (p : ⊢(a :: b :: Γ)) (n : Nat) (hp : p.HasSize n) : p.parr.HasSize (n + 1) + | tensor (p : ⊢(a :: Γ)) (q : ⊢(b :: Δ)) (np nq : Nat) + (hp : p.HasSize np) (hq : q.HasSize nq) : + (Proof.tensor p q).HasSize (np + nq + 1) + | oplus₁ (p : ⊢(a :: Γ)) (n : Nat) (hp : p.HasSize n) : p.oplus₁.HasSize (n + 1) + | oplus₂ (p : ⊢(b :: Γ)) (n : Nat) (hp : p.HasSize n) : p.oplus₂.HasSize (n + 1) + | with (p : ⊢(a :: Γ)) (q : ⊢(b :: Γ)) (np nq : Nat) + (hp : p.HasSize np) (hq : q.HasSize nq) : + (Proof.with p q).HasSize (np + nq + 1) + | top : Proof.top.HasSize 1 + | quest (p : ⊢(a :: Γ)) (n : Nat) (hp : p.HasSize n) : p.quest.HasSize (n + 1) + | weaken (p : ⊢Γ) (n : Nat) (hp : p.HasSize n) : p.weaken.HasSize (n + 1) + | contract (p : ⊢(ʔa :: ʔa :: Γ)) (n : Nat) (hp : p.HasSize n) : + p.contract.HasSize (n + 1) + | bang (hqs: Sequent.allQuest Γ) (p : ⊢(a :: Γ)) (n : Nat) (hp : p.HasSize n) : + (p.bang hqs).HasSize (n + 1) + | cut (p : ⊢(a :: Γ)) (q : ⊢(a.dual :: Δ)) (np nq : Nat) + (hp : p.HasSize np) (hq : q.HasSize nq) : + (Proof.cut p q).HasSize (np + nq + 1) + /-- Cut is admissible. -/ proof_wanted Proof.cut_admissible - (p : ⊢(a :: Γ)) (q : ⊢(a.dual :: Δ)) (hp : p.CutFree) (hq : q.CutFree) : + {a : Proposition Atom} (p : ⊢(a :: Γ)) (q : ⊢(a.dual :: Δ)) (hp : p.CutFree) (hq : q.CutFree) : ∃ r : ⊢(Γ ++ Δ), r.CutFree /-- Cut elimination: for any sequent Γ, if there is a proof of Γ, then there exists a cut-free From 05ae340eeaa03b883b5879f75e3fff7ae1757588 Mon Sep 17 00:00:00 2001 From: Fabrizio Montesi Date: Mon, 8 Sep 2025 09:27:56 +0200 Subject: [PATCH 07/19] Bisimulation -> IsBisimulation and lots of grinding --- Cslib/Foundations/Data/Relation.lean | 2 +- .../Semantics/Lts/Bisimulation.lean | 419 ++++++++---------- Cslib/Languages/CCS/BehaviouralTheory.lean | 61 +-- 3 files changed, 205 insertions(+), 277 deletions(-) diff --git a/Cslib/Foundations/Data/Relation.lean b/Cslib/Foundations/Data/Relation.lean index 9bc36b69..e3bd649e 100644 --- a/Cslib/Foundations/Data/Relation.lean +++ b/Cslib/Foundations/Data/Relation.lean @@ -13,7 +13,7 @@ universe u v section Relation /-- The relation `r` 'up to' the relation `s`. -/ -def Relation.upTo (r s : α → α → Prop) : α → α → Prop := Relation.Comp s (Relation.Comp r s) +def Relation.UpTo (r s : α → α → Prop) : α → α → Prop := Relation.Comp s (Relation.Comp r s) /-- A relation has the diamond property when all reductions with a common origin are joinable -/ abbrev Diamond (R : α → α → Prop) := ∀ {A B C : α}, R A B → R A C → (∃ D, R B D ∧ R C D) diff --git a/Cslib/Foundations/Semantics/Lts/Bisimulation.lean b/Cslib/Foundations/Semantics/Lts/Bisimulation.lean index c5a2796d..2067fc2f 100644 --- a/Cslib/Foundations/Semantics/Lts/Bisimulation.lean +++ b/Cslib/Foundations/Semantics/Lts/Bisimulation.lean @@ -27,16 +27,17 @@ For an introduction to theory of bisimulation, we refer to [Sangiorgi2011]. ## Main definitions -- `Bisimulation lts r`: the relation `r` on the states of the Lts `lts` is a bisimulation. +- `lts.IsBisimulation r`: the relation `r` is a bisimulation for the LTS `lts`. - `Bisimilarity lts` is the binary relation on the states of `lts` that relates any two states related by some bisimulation on `lts`. -- `BisimulationUpTo lts r`: the relation `r` is a bisimulation up to bisimilarity (this is known as -one of the 'up to' techniques for bisimulation). +- `lts.IsBisimulationUpTo r`: the relation `r` is a bisimulation up to bisimilarity (this is known +as one of the 'up to' techniques for bisimulation). -- `WeakBisimulation lts r`: the relation `r` on the states of the Lts `lts` is a weak bisimulation. +- `lts.IsWeakBisimulation r`: the relation `r` on the states of the Lts `lts` is a weak +bisimulation. - `WeakBisimilarity lts` is the binary relation on the states of `lts` that relates any two states related by some weak bisimulation on `lts`. -- `SWBisimulation lts` is a more convenient definition for establishing weak bisimulations, which +- `lts.IsSWBisimulation` is a more convenient definition for establishing weak bisimulations, which we prove to be sound and complete. - `SWBisimilarity lts` is the binary relation on the states of `lts` that relates any two states related by some sw-bisimulation on `lts`. @@ -70,36 +71,47 @@ universe u v section Bisimulation -variable {State : Type u} {Label : Type v} (lts : Lts State Label) +variable {State : Type u} {Label : Type v} {lts : Lts State Label} /-- A relation is a bisimulation if, whenever it relates two states in an lts, the transitions originating from these states mimic each other and the reached derivatives are themselves related. -/ -def Bisimulation (lts : Lts State Label) (r : State → State → Prop) : Prop := - ∀ s1 s2, r s1 s2 → ∀ μ, ( +@[grind] +def Lts.IsBisimulation (lts : Lts State Label) (r : State → State → Prop) : Prop := + ∀ ⦃s1 s2⦄, r s1 s2 → ∀ μ, ( (∀ s1', lts.Tr s1 μ s1' → ∃ s2', lts.Tr s2 μ s2' ∧ r s1' s2') ∧ (∀ s2', lts.Tr s2 μ s2' → ∃ s1', lts.Tr s1 μ s1' ∧ r s1' s2') ) -/-- Helper for following a transition using the first component of a `Bisimulation`. -/ -def Bisimulation.follow_fst - {lts : Lts State Label} {r : State → State → Prop} - {s1 s2 : State} {μ : Label} {s1' : State} - (hb : Bisimulation lts r) (hr : r s1 s2) (htr : lts.Tr s1 μ s1') := - (hb _ _ hr μ).1 _ htr - -/-- Helper for following a transition using the second component of a `Bisimulation`. -/ -def Bisimulation.follow_snd - {lts : Lts State Label} {r : State → State → Prop} - {s1 s2 : State} {μ : Label} {s2' : State} - (hb : Bisimulation lts r) (hr : r s1 s2) (htr : lts.Tr s2 μ s2') := - (hb _ _ hr μ).2 _ htr +/- Semi-bundled version of `Lts.IsBisimulation`. -/ +-- @[grind ext] +-- structure Bisimulation (lts : Lts State Label) where +-- /-- The relation on the states of the lts. -/ +-- rel : State → State → Prop +-- /-- Proof that the relation is a bisimulation. -/ +-- is_bisimulation : lts.IsBisimulation rel + +/- Any `Bisimulation` can be coerced into a relation. -/ +-- instance : CoeFun (lts.IsBisimulation) (fun _ => State → State → Prop) where +-- coe := fun bisim => bisim.rel + +/-- Helper for following a transition by the first state in a pair of a `Bisimulation`. -/ +theorem Lts.IsBisimulation.follow_fst + (hb : lts.IsBisimulation r) (hr : r s1 s2) (htr : lts.Tr s1 μ s1') : + ∃ s2', lts.Tr s2 μ s2' ∧ r s1' s2' := + (hb hr μ).1 _ htr + +/-- Helper for following a transition by the second state in a pair of a `Bisimulation`. -/ +theorem Lts.IsBisimulation.follow_snd + (hb : lts.IsBisimulation r) (hr : r s1 s2) (htr : lts.Tr s2 μ s2') : + ∃ s1', lts.Tr s1 μ s1' ∧ r s1' s2' := + (hb hr μ).2 _ htr /-- Two states are bisimilar if they are related by some bisimulation. -/ +@[grind] def Bisimilarity (lts : Lts State Label) : State → State → Prop := - fun s1 s2 => - ∃ r : State → State → Prop, r s1 s2 ∧ Bisimulation lts r + fun s1 s2 => ∃ r : State → State → Prop, r s1 s2 ∧ lts.IsBisimulation r /-- Notation for bisimilarity. @@ -110,105 +122,57 @@ explicitly. notation s:max " ~[" lts "] " s':max => Bisimilarity lts s s' /-- Bisimilarity is reflexive. -/ +@[grind, refl] theorem Bisimilarity.refl (s : State) : s ~[lts] s := by exists Eq - grind [Bisimulation] + grind /-- The inverse of a bisimulation is a bisimulation. -/ -theorem Bisimulation.inv (h : Bisimulation lts r) : - Bisimulation lts (flip r) := by - simp only [Bisimulation] at h - simp only [Bisimulation] - intro s1 s2 hrinv μ - constructor - case left => - intro s1' htr - specialize h s2 s1 hrinv μ - have h' := h.2 s1' htr - obtain ⟨ s2', h' ⟩ := h' - exists s2' - case right => - intro s2' htr - specialize h s2 s1 hrinv μ - have h' := h.1 s2' htr - obtain ⟨ s1', h' ⟩ := h' - exists s1' +@[grind] +theorem Bisimulation.inv (h : lts.IsBisimulation r) : + lts.IsBisimulation (flip r) := by grind [flip] /-- Bisimilarity is symmetric. -/ +@[grind, symm] theorem Bisimilarity.symm {s1 s2 : State} (h : s1 ~[lts] s2) : s2 ~[lts] s1 := by - obtain ⟨r, hr, hb⟩ := h + obtain ⟨r, _, _⟩ := h exists (flip r) - constructor - case left => - exact hr - case right => - apply Bisimulation.inv - exact hb + grind [flip] /-- The composition of two bisimulations is a bisimulation. -/ +@[grind] theorem Bisimulation.comp - (r1 r2 : State → State → Prop) (h1 : Bisimulation lts r1) (h2 : Bisimulation lts r2) : - Bisimulation lts (Relation.Comp r1 r2) := by - simp_all only [Bisimulation] - intro s1 s2 hrc μ - constructor - case left => - intro s1' htr - rcases hrc with ⟨sb, hr1, hr2⟩ - specialize h1 s1 sb hr1 μ - specialize h2 sb s2 hr2 μ - have h1' := h1.1 s1' htr - obtain ⟨s1'', h1'tr, h1'⟩ := h1' - have h2' := h2.1 s1'' h1'tr - obtain ⟨s2'', h2'tr, h2'⟩ := h2' - exists s2'' - constructor - · exact h2'tr - · exists s1'' - case right => - intro s2' htr - rcases hrc with ⟨sb, hr1, hr2⟩ - specialize h1 s1 sb hr1 μ - specialize h2 sb s2 hr2 μ - have h2' := h2.2 s2' htr - obtain ⟨s2'', h2'tr, h2'⟩ := h2' - have h1' := h1.2 s2'' h2'tr - obtain ⟨s1'', h1'tr, h1'⟩ := h1' - exists s1'' - constructor - · exact h1'tr - · exists s2'' + (h1 : lts.IsBisimulation r1) (h2 : lts.IsBisimulation r2) : + lts.IsBisimulation (Relation.Comp r1 r2) := by grind [Relation.Comp] /-- Bisimilarity is transitive. -/ +@[grind] theorem Bisimilarity.trans - {s1 s2 s3 : State} (h1 : s1 ~[lts] s2) (h2 : s2 ~[lts] s3) : + (h1 : s1 ~[lts] s2) (h2 : s2 ~[lts] s3) : s1 ~[lts] s3 := by - obtain ⟨r1, hr1, hr1b⟩ := h1 - obtain ⟨r2, hr2, hr2b⟩ := h2 + obtain ⟨r1, _, _⟩ := h1 + obtain ⟨r2, _, _⟩ := h2 exists Relation.Comp r1 r2 - constructor - case left => - exists s2 - case right => - apply Bisimulation.comp lts r1 r2 hr1b hr2b + grind [Relation.Comp] /-- Bisimilarity is an equivalence relation. -/ -theorem Bisimilarity.eqv (lts : Lts State Label) : +theorem Bisimilarity.eqv : Equivalence (Bisimilarity lts) := { - refl := Bisimilarity.refl lts - symm := Bisimilarity.symm lts - trans := Bisimilarity.trans lts + refl := Bisimilarity.refl + symm := Bisimilarity.symm + trans := Bisimilarity.trans } /-- The union of two bisimulations is a bisimulation. -/ -theorem Bisimulation.union (hrb : Bisimulation lts r) (hsb : Bisimulation lts s) : - Bisimulation lts (r ⊔ s) := by +@[grind] +theorem Bisimulation.union (hrb : lts.IsBisimulation r) (hsb : lts.IsBisimulation s) : + lts.IsBisimulation (r ⊔ s) := by intro s1 s2 hrs μ cases hrs case inl h => constructor · intro s1' htr - obtain ⟨s2', htr', hr'⟩ := Bisimulation.follow_fst hrb h htr + obtain ⟨s2', htr', hr'⟩ := hrb.follow_fst h htr exists s2' constructor · assumption @@ -216,7 +180,7 @@ theorem Bisimulation.union (hrb : Bisimulation lts r) (hsb : Bisimulation lts s) left exact hr' · intro s2' htr - obtain ⟨s1', htr', hr'⟩ := Bisimulation.follow_snd hrb h htr + obtain ⟨s1', htr', hr'⟩ := hrb.follow_snd h htr exists s1' constructor · assumption @@ -226,7 +190,7 @@ theorem Bisimulation.union (hrb : Bisimulation lts r) (hsb : Bisimulation lts s) case inr h => constructor · intro s1' htr - obtain ⟨s2', htr', hs'⟩ := Bisimulation.follow_fst hsb h htr + obtain ⟨s2', htr', hs'⟩ := hsb.follow_fst h htr exists s2' constructor · assumption @@ -234,7 +198,7 @@ theorem Bisimulation.union (hrb : Bisimulation lts r) (hsb : Bisimulation lts s) right exact hs' · intro s2' htr - obtain ⟨s1', htr', hs'⟩ := Bisimulation.follow_snd hsb h htr + obtain ⟨s1', htr', hs'⟩ := hsb.follow_snd h htr exists s1' constructor · assumption @@ -243,65 +207,38 @@ theorem Bisimulation.union (hrb : Bisimulation lts r) (hsb : Bisimulation lts s) exact hs' /-- Bisimilarity is a bisimulation. -/ -theorem Bisimilarity.is_bisimulation : Bisimulation lts (Bisimilarity lts) := by - simp only [Bisimulation] - intro s1 s2 h μ - obtain ⟨r, hr, hb⟩ := h - have hrBisim := hb - specialize hb s1 s2 - constructor - case left => - intro s1' htr - specialize hb hr μ - obtain ⟨hb1, hb2⟩ := hb - specialize hb1 s1' htr - obtain ⟨s2', htr2, hr2⟩ := hb1 - exists s2' - constructor - case left => - exact htr2 - case right => - exists r - case right => - intro s2' htr - specialize hb hr μ - obtain ⟨hb1, hb2⟩ := hb - specialize hb2 s2' htr - obtain ⟨s1', htr1, hr1⟩ := hb2 - exists s1' - constructor - case left => - exact htr1 - case right => - exists r +@[grind] +theorem Bisimilarity.is_bisimulation : lts.IsBisimulation (Bisimilarity lts) := by grind /-- Bisimilarity is the largest bisimulation. -/ +@[grind] theorem Bisimilarity.largest_bisimulation - (h : Bisimulation lts r) : + (h : lts.IsBisimulation r) : Subrelation r (Bisimilarity lts) := by intro s1 s2 hr exists r /-- The union of bisimilarity with any bisimulation is bisimilarity. -/ -theorem Bisimilarity.gfp (r : State → State → Prop) (h : Bisimulation lts r) : +@[grind, simp] +theorem Bisimilarity.gfp (r : State → State → Prop) (h : lts.IsBisimulation r) : (Bisimilarity lts) ⊔ r = Bisimilarity lts := by funext s1 s2 simp only [max, SemilatticeSup.sup, eq_iff_iff, or_iff_left_iff_imp] - apply Bisimilarity.largest_bisimulation lts h + apply Bisimilarity.largest_bisimulation h /-- `calc` support for bisimilarity. -/ instance : Trans (Bisimilarity lts) (Bisimilarity lts) (Bisimilarity lts) where - trans := Bisimilarity.trans lts + trans := Bisimilarity.trans section Order /-! ## Order properties -/ -noncomputable instance : Max {r // Bisimulation lts r} where - max r s := ⟨r.1 ⊔ s.1, Bisimulation.union lts r.2 s.2⟩ +noncomputable instance : Max {r // lts.IsBisimulation r} where + max r s := ⟨r.1 ⊔ s.1, Bisimulation.union r.2 s.2⟩ /-- Bisimulations equipped with union form a join-semilattice. -/ -noncomputable instance : SemilatticeSup {r // Bisimulation lts r} where +noncomputable instance : SemilatticeSup {r // lts.IsBisimulation r} where sup r s := r ⊔ s le_sup_left r s := by simp only [LE.le] @@ -327,7 +264,8 @@ noncomputable instance : SemilatticeSup {r // Bisimulation lts r} where apply h2 _ _ h /-- The empty relation is a bisimulation. -/ -theorem Bisimulation.emptyRelation_bisimulation : Bisimulation lts emptyRelation := by +@[grind] +theorem Bisimulation.emptyRelation_bisimulation : lts.IsBisimulation emptyRelation := by intro s1 s2 hr cases hr @@ -336,13 +274,13 @@ theorem Bisimulation.emptyRelation_bisimulation : Bisimulation lts emptyRelation - The empty relation is the bottom element. - Bisimilarity is the top element. -/ -instance : BoundedOrder {r // Bisimulation lts r} where - top := ⟨Bisimilarity lts, Bisimilarity.is_bisimulation lts⟩ - bot := ⟨emptyRelation, Bisimulation.emptyRelation_bisimulation lts⟩ +instance : BoundedOrder {r // lts.IsBisimulation r} where + top := ⟨Bisimilarity lts, Bisimilarity.is_bisimulation⟩ + bot := ⟨emptyRelation, Bisimulation.emptyRelation_bisimulation⟩ le_top r := by intro s1 s2 simp only [LE.le] - apply Bisimilarity.largest_bisimulation lts r.2 + apply Bisimilarity.largest_bisimulation r.2 bot_le r := by intro s1 s2 simp only [LE.le] @@ -356,16 +294,18 @@ end Order /-- A relation `r` is a bisimulation up to bisimilarity if, whenever it relates two states in an lts, the transitions originating from these states mimic each other and the reached derivatives are themselves related by `r` up to bisimilarity. -/ -def BisimulationUpTo (lts : Lts State Label) (r : State → State → Prop) : Prop := - ∀ s1 s2, r s1 s2 → ∀ μ, ( - (∀ s1', lts.Tr s1 μ s1' → ∃ s2', lts.Tr s2 μ s2' ∧ Relation.upTo r (Bisimilarity lts) s1' s2') +@[grind] +def Lts.IsBisimulationUpTo (lts : Lts State Label) (r : State → State → Prop) : Prop := + ∀ ⦃s1 s2⦄, r s1 s2 → ∀ μ, ( + (∀ s1', lts.Tr s1 μ s1' → ∃ s2', lts.Tr s2 μ s2' ∧ Relation.UpTo r (Bisimilarity lts) s1' s2') ∧ - (∀ s2', lts.Tr s2 μ s2' → ∃ s1', lts.Tr s1 μ s1' ∧ Relation.upTo r (Bisimilarity lts) s1' s2') + (∀ s2', lts.Tr s2 μ s2' → ∃ s1', lts.Tr s1 μ s1' ∧ Relation.UpTo r (Bisimilarity lts) s1' s2') ) /-- Any bisimulation up to bisimilarity is a bisimulation. -/ -theorem Bisimulation.upTo_bisimulation (r : State → State → Prop) (h : BisimulationUpTo lts r) : - Bisimulation lts (Relation.upTo r (Bisimilarity lts)) := by +@[grind] +theorem Bisimulation.upTo_bisimulation (h : lts.IsBisimulationUpTo r) : + lts.IsBisimulation (Relation.UpTo r (Bisimilarity lts)) := by intro s1 s2 hr μ rcases hr with ⟨s1b, hr1b, s2b, hrb, hr2b⟩ obtain ⟨r1, hr1, hr1b⟩ := hr1b @@ -373,9 +313,9 @@ theorem Bisimulation.upTo_bisimulation (r : State → State → Prop) (h : Bisim constructor case left => intro s1' htr1 - obtain ⟨s1b', hs1b'tr, hs1b'r⟩ := (hr1b _ _ hr1 μ).1 s1' htr1 - obtain ⟨s2b', hs2b'tr, hs2b'r⟩ := (h s1b s2b hrb μ).1 s1b' hs1b'tr - obtain ⟨s2', hs2btr, hs2br⟩ := (hr2b _ _ hr2 μ).1 _ hs2b'tr + obtain ⟨s1b', hs1b'tr, hs1b'r⟩ := (hr1b hr1 μ).1 s1' htr1 + obtain ⟨s2b', hs2b'tr, hs2b'r⟩ := (h hrb μ).1 s1b' hs1b'tr + obtain ⟨s2', hs2btr, hs2br⟩ := (hr2b hr2 μ).1 _ hs2b'tr exists s2' constructor case left => @@ -384,18 +324,18 @@ theorem Bisimulation.upTo_bisimulation (r : State → State → Prop) (h : Bisim obtain ⟨smid1, hsmidb, smid2, hsmidr, hsmidrb⟩ := hs2b'r constructor constructor - · apply Bisimilarity.trans lts (Bisimilarity.largest_bisimulation lts hr1b hs1b'r) + · apply Bisimilarity.trans (Bisimilarity.largest_bisimulation hr1b hs1b'r) hsmidb · exists smid2 constructor · exact hsmidr - · apply Bisimilarity.trans lts hsmidrb - apply Bisimilarity.largest_bisimulation lts hr2b hs2br + · apply Bisimilarity.trans hsmidrb + apply Bisimilarity.largest_bisimulation hr2b hs2br case right => intro s2' htr2 - obtain ⟨s2b', hs2b'tr, hs2b'r⟩ := (hr2b _ _ hr2 μ).2 s2' htr2 - obtain ⟨s1b', hs1b'tr, hs1b'r⟩ := (h s1b s2b hrb μ).2 s2b' hs2b'tr - obtain ⟨s1', hs1btr, hs1br⟩ := (hr1b _ _ hr1 μ).2 _ hs1b'tr + obtain ⟨s2b', hs2b'tr, hs2b'r⟩ := (hr2b hr2 μ).2 s2' htr2 + obtain ⟨s1b', hs1b'tr, hs1b'r⟩ := (h hrb μ).2 s2b' hs2b'tr + obtain ⟨s1', hs1btr, hs1br⟩ := (hr1b hr1 μ).2 _ hs1b'tr exists s1' constructor case left => @@ -404,19 +344,19 @@ theorem Bisimulation.upTo_bisimulation (r : State → State → Prop) (h : Bisim obtain ⟨smid1, hsmidb, smid2, hsmidr, hsmidrb⟩ := hs1b'r constructor constructor - · apply Bisimilarity.trans lts (Bisimilarity.largest_bisimulation lts hr1b _) hsmidb + · apply Bisimilarity.trans (Bisimilarity.largest_bisimulation hr1b _) hsmidb · exact hs1br · exists smid2 constructor · exact hsmidr - · apply Bisimilarity.trans lts hsmidrb - apply Bisimilarity.largest_bisimulation lts hr2b _ + · apply Bisimilarity.trans hsmidrb + apply Bisimilarity.largest_bisimulation hr2b _ exact hs2b'r /-- If two states are related by a bisimulation, they can mimic each other's multi-step transitions. -/ theorem Bisimulation.bisim_trace - (s1 s2 : State) (r : State → State → Prop) (hb : Bisimulation lts r) (hr : r s1 s2) : + (hb : lts.IsBisimulation r) (hr : r s1 s2) : ∀ μs s1', lts.MTr s1 μs s1' → ∃ s2', lts.MTr s2 μs s2' ∧ r s1' s2' := by intro μs induction μs generalizing s1 s2 @@ -431,10 +371,10 @@ theorem Bisimulation.bisim_trace intro s1' hmtr1 cases hmtr1 case stepL s1'' htr hmtr => - specialize hb s1 s2 hr μ + specialize hb hr μ have hf := hb.1 s1'' htr obtain ⟨s2'', htr2, hb2⟩ := hf - specialize ih s1'' s2'' hb2 s1' hmtr + specialize ih hb2 s1' hmtr obtain ⟨s2', hmtr2, hr'⟩ := ih exists s2' constructor @@ -448,8 +388,9 @@ theorem Bisimulation.bisim_trace /-! ## Relation to trace equivalence -/ /-- Any bisimulation implies trace equivalence. -/ +@[grind] theorem Bisimulation.bisim_traceEq - (hb : Bisimulation lts r) (hr : r s1 s2) : + (hb : lts.IsBisimulation r) (hr : r s1 s2) : s1 ~tr[lts] s2 := by funext μs simp only [eq_iff_iff] @@ -457,22 +398,23 @@ theorem Bisimulation.bisim_traceEq case mp => intro h obtain ⟨s1', h⟩ := h - obtain ⟨s2', hmtr⟩ := @Bisimulation.bisim_trace State Label lts s1 s2 r hb hr μs s1' h + obtain ⟨s2', hmtr⟩ := Bisimulation.bisim_trace hb hr μs s1' h exists s2' exact hmtr.1 case mpr => intro h obtain ⟨s2', h⟩ := h have hinv := @Bisimulation.inv State Label lts r hb - obtain ⟨s1', hmtr⟩ := @Bisimulation.bisim_trace State Label lts s2 s1 (flip r) hinv hr μs s2' h + obtain ⟨s1', hmtr⟩ := Bisimulation.bisim_trace hinv hr μs s2' h exists s1' exact hmtr.1 /-- Bisimilarity is included in trace equivalence. -/ +@[grind] theorem Bisimilarity.le_traceEq : Bisimilarity lts ≤ TraceEq lts := by intro s1 s2 h obtain ⟨r, hr, hb⟩ := h - apply Bisimulation.bisim_traceEq lts hb hr + apply Bisimulation.bisim_traceEq hb hr /- One of the standard motivating examples for bisimulation: `1` and `5` are trace equivalent, but not bisimilar. -/ @@ -490,13 +432,13 @@ private inductive BisimMotTr : ℕ → Char → ℕ → Prop where /-- In general, trace equivalence is not a bisimulation (extra conditions are needed, see for example `Bisimulation.deterministic_trace_eq_is_bisim`). -/ theorem Bisimulation.traceEq_not_bisim : - ∃ (State : Type) (Label : Type) (lts : Lts State Label), ¬(Bisimulation lts (TraceEq lts)) := by + ∃ (State : Type) (Label : Type) (lts : Lts State Label), ¬(lts.IsBisimulation (TraceEq lts)) := by exists ℕ exists Char let lts := Lts.mk BisimMotTr exists lts intro h - specialize h 1 5 + -- specialize h 1 5 have htreq : (1 ~tr[lts] 5) := by simp [TraceEq] have htraces1 : lts.traces 1 = {[], ['a'], ['a', 'b'], ['a', 'c']} := by @@ -724,7 +666,7 @@ theorem Bisimulation.traceEq_not_bisim : cases htr cases hmtr case stepL μ sb μs' htr hmtr => cases htr - simp + simp case mpr => intro h cases h @@ -753,15 +695,15 @@ theorem Bisimilarity.bisimilarity_neq_traceEq : obtain ⟨State, Label, lts, h⟩ := Bisimulation.traceEq_not_bisim exists State; exists Label; exists lts intro heq - have hb := Bisimilarity.is_bisimulation lts + have hb := Bisimilarity.is_bisimulation (lts := lts) rw [heq] at hb contradiction /-- In any deterministic Lts, trace equivalence is a bisimulation. -/ theorem Bisimulation.deterministic_traceEq_is_bisim - (lts : Lts State Label) (hdet : lts.Deterministic) : - (Bisimulation lts (TraceEq lts)) := by - simp only [Bisimulation] + (hdet : lts.Deterministic) : + (lts.IsBisimulation (TraceEq lts)) := by + simp only [Lts.IsBisimulation] intro s1 s2 hteq μ constructor case left => @@ -780,14 +722,14 @@ theorem Bisimulation.deterministic_traceEq_is_bisim /-- In any deterministic Lts, trace equivalence implies bisimilarity. -/ theorem Bisimilarity.deterministic_traceEq_bisim - (lts : Lts State Label) (hdet : lts.Deterministic) (s1 s2 : State) (h : s1 ~tr[lts] s2) : + (hdet : lts.Deterministic) (h : s1 ~tr[lts] s2) : (s1 ~[lts] s2) := by exists TraceEq lts constructor case left => exact h case right => - apply Bisimulation.deterministic_traceEq_is_bisim lts hdet + apply Bisimulation.deterministic_traceEq_is_bisim hdet /-- In any deterministic Lts, bisimilarity and trace equivalence coincide. -/ theorem Bisimilarity.deterministic_bisim_eq_traceEq @@ -799,21 +741,21 @@ theorem Bisimilarity.deterministic_bisim_eq_traceEq case mp => apply Bisimilarity.le_traceEq case mpr => - apply Bisimilarity.deterministic_traceEq_bisim lts hdet + apply Bisimilarity.deterministic_traceEq_bisim hdet /-! ## Relation to simulation -/ /-- Any bisimulation is also a simulation. -/ theorem Bisimulation.is_simulation (lts : Lts State Label) (r : State → State → Prop) : - Bisimulation lts r → Simulation lts r := by - grind [Bisimulation, Simulation] + lts.IsBisimulation r → Simulation lts r := by + grind [Simulation] /-- A relation is a bisimulation iff both it and its inverse are simulations. -/ theorem Bisimulation.simulation_iff (lts : Lts State Label) (r : State → State → Prop) : - Bisimulation lts r ↔ (Simulation lts r ∧ Simulation lts (flip r)) := by + lts.IsBisimulation r ↔ (Simulation lts r ∧ Simulation lts (flip r)) := by constructor - case mp => grind [Bisimulation, Simulation, flip] - case mpr => aesop (add simp [Bisimulation]) + case mp => grind [Simulation, flip] + case mpr => aesop (add simp [Lts.IsBisimulation]) end Bisimulation @@ -823,13 +765,13 @@ section WeakBisimulation /-- A weak bisimulation is similar to a `Bisimulation`, but allows for the related processes to do internal work. Technically, this is defined as a `Bisimulation` on the saturation of the Lts. -/ -def WeakBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) := - Bisimulation (lts.saturate) r +def Lts.IsWeakBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) := + lts.saturate.IsBisimulation r /-- Two states are weakly bisimilar if they are related by some weak bisimulation. -/ def WeakBisimilarity [HasTau Label] (lts : Lts State Label) : State → State → Prop := fun s1 s2 => - ∃ r : State → State → Prop, r s1 s2 ∧ WeakBisimulation lts r + ∃ r : State → State → Prop, r s1 s2 ∧ lts.IsWeakBisimulation r /-- Notation for weak bisimilarity. -/ notation s:max " ≈[" lts "] " s':max => WeakBisimilarity lts s s' @@ -837,8 +779,8 @@ notation s:max " ≈[" lts "] " s':max => WeakBisimilarity lts s s' /-- An `SWBisimulation` is a more convenient definition of weak bisimulation, because the challenge is a single transition. We prove later that this technique is sound, following a strategy inspired by [Sangiorgi2011]. -/ -def SWBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) : Prop := - ∀ s1 s2, r s1 s2 → ∀ μ, ( +def Lts.IsSWBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) : Prop := + ∀ ⦃s1 s2⦄, r s1 s2 → ∀ μ, ( (∀ s1', lts.Tr s1 μ s1' → ∃ s2', lts.STr s2 μ s2' ∧ r s1' s2') ∧ (∀ s2', lts.Tr s2 μ s2' → ∃ s1', lts.STr s1 μ s1' ∧ r s1' s2') @@ -847,7 +789,7 @@ def SWBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State /-- Two states are sw-bisimilar if they are related by some sw-bisimulation. -/ def SWBisimilarity [HasTau Label] (lts : Lts State Label) : State → State → Prop := fun s1 s2 => - ∃ r : State → State → Prop, r s1 s2 ∧ SWBisimulation lts r + ∃ r : State → State → Prop, r s1 s2 ∧ lts.IsSWBisimulation r /-- Notation for swbisimilarity. -/ notation s:max " ≈sw[" lts "] " s':max => SWBisimilarity lts s s' @@ -855,8 +797,8 @@ notation s:max " ≈sw[" lts "] " s':max => SWBisimilarity lts s s' /-- Utility theorem for 'following' internal transitions using an `SWBisimulation` (first component, weighted version). -/ theorem SWBisimulation.follow_internal_fst_n - [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) - (hswb : SWBisimulation lts r) (hr : r s1 s2) (hstrN : lts.strN n s1 HasTau.τ s1') : + [HasTau Label] {lts : Lts State Label} + (hswb : lts.IsSWBisimulation r) (hr : r s1 s2) (hstrN : lts.strN n s1 HasTau.τ s1') : ∃ s2', lts.STr s2 HasTau.τ s2' ∧ r s1' s2' := by cases n case zero => @@ -868,11 +810,11 @@ theorem SWBisimulation.follow_internal_fst_n cases hstrN rename_i n1 sb sb' n2 hstrN1 htr hstrN2 let hswb_m := hswb - have ih1 := SWBisimulation.follow_internal_fst_n lts r hswb hr hstrN1 + have ih1 := SWBisimulation.follow_internal_fst_n hswb hr hstrN1 obtain ⟨sb2, hstrs2, hrsb⟩ := ih1 - have h := (hswb sb sb2 hrsb HasTau.τ).1 sb' htr + have h := (hswb hrsb HasTau.τ).1 sb' htr obtain ⟨sb2', hstrsb2, hrsb2⟩ := h - have ih2 := SWBisimulation.follow_internal_fst_n lts r hswb hrsb2 hstrN2 + have ih2 := SWBisimulation.follow_internal_fst_n hswb hrsb2 hstrN2 obtain ⟨s2', hstrs2', hrs2⟩ := ih2 exists s2' constructor @@ -882,8 +824,8 @@ theorem SWBisimulation.follow_internal_fst_n /-- Utility theorem for 'following' internal transitions using an `SWBisimulation` (second component, weighted version). -/ theorem SWBisimulation.follow_internal_snd_n - [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) - (hswb : SWBisimulation lts r) (hr : r s1 s2) (hstrN : lts.strN n s2 HasTau.τ s2') : + [HasTau Label] {lts : Lts State Label} + (hswb : lts.IsSWBisimulation r) (hr : r s1 s2) (hstrN : lts.strN n s2 HasTau.τ s2') : ∃ s1', lts.STr s1 HasTau.τ s1' ∧ r s1' s2' := by cases n case zero => @@ -895,11 +837,11 @@ theorem SWBisimulation.follow_internal_snd_n cases hstrN rename_i n1 sb sb' n2 hstrN1 htr hstrN2 let hswb_m := hswb - have ih1 := SWBisimulation.follow_internal_snd_n lts r hswb hr hstrN1 + have ih1 := SWBisimulation.follow_internal_snd_n hswb hr hstrN1 obtain ⟨sb1, hstrs1, hrsb⟩ := ih1 - have h := (hswb sb1 sb hrsb HasTau.τ).2 sb' htr + have h := (hswb hrsb HasTau.τ).2 sb' htr obtain ⟨sb2', hstrsb2, hrsb2⟩ := h - have ih2 := SWBisimulation.follow_internal_snd_n lts r hswb hrsb2 hstrN2 + have ih2 := SWBisimulation.follow_internal_snd_n hswb hrsb2 hstrN2 obtain ⟨s2', hstrs2', hrs2⟩ := ih2 exists s2' constructor @@ -909,26 +851,26 @@ theorem SWBisimulation.follow_internal_snd_n /-- Utility theorem for 'following' internal transitions using an `SWBisimulation` (first component). -/ theorem SWBisimulation.follow_internal_fst - [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) - (hswb : SWBisimulation lts r) (hr : r s1 s2) (hstr : lts.STr s1 HasTau.τ s1') : + [HasTau Label] {lts : Lts State Label} + (hswb : lts.IsSWBisimulation r) (hr : r s1 s2) (hstr : lts.STr s1 HasTau.τ s1') : ∃ s2', lts.STr s2 HasTau.τ s2' ∧ r s1' s2' := by obtain ⟨n, hstrN⟩ := (Lts.str_strN lts).1 hstr - apply SWBisimulation.follow_internal_fst_n lts r hswb hr hstrN + apply SWBisimulation.follow_internal_fst_n hswb hr hstrN /-- Utility theorem for 'following' internal transitions using an `SWBisimulation` (second component). -/ theorem SWBisimulation.follow_internal_snd - [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) - (hswb : SWBisimulation lts r) (hr : r s1 s2) (hstr : lts.STr s2 HasTau.τ s2') : + [HasTau Label] {lts : Lts State Label} + (hswb : lts.IsSWBisimulation r) (hr : r s1 s2) (hstr : lts.STr s2 HasTau.τ s2') : ∃ s1', lts.STr s1 HasTau.τ s1' ∧ r s1' s2' := by obtain ⟨n, hstrN⟩ := (Lts.str_strN lts).1 hstr - apply SWBisimulation.follow_internal_snd_n lts r hswb hr hstrN + apply SWBisimulation.follow_internal_snd_n hswb hr hstrN /-- We can now prove that any relation is a `WeakBisimulation` iff it is an `SWBisimulation`. This formalises lemma 4.2.10 in [Sangiorgi2011]. -/ -theorem WeakBisimulation.iff_swBisimulation - [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) : - WeakBisimulation lts r ↔ SWBisimulation lts r := by +theorem Lts.isWeakBisimulation_iff_isSWBisimulation + [HasTau Label] {lts : Lts State Label} : + lts.IsWeakBisimulation r ↔ lts.IsSWBisimulation r := by apply Iff.intro case mp => intro h @@ -936,13 +878,13 @@ theorem WeakBisimulation.iff_swBisimulation apply And.intro case left => intro s1' htr - specialize h s1 s2 hr μ + specialize h hr μ have h' := h.1 s1' (Lts.STr.single lts htr) obtain ⟨s2', htr2, hr2⟩ := h' exists s2' case right => intro s2' htr - specialize h s1 s2 hr μ + specialize h hr μ have h' := h.2 s2' (Lts.STr.single lts htr) obtain ⟨s1', htr1, hr1⟩ := h' exists s1' @@ -958,9 +900,9 @@ theorem WeakBisimulation.iff_swBisimulation constructor; constructor exact hr case tr sb sb' hstr1 htr hstr2 => - obtain ⟨sb2, hstr2b, hrb⟩ := SWBisimulation.follow_internal_fst lts r h hr hstr1 - obtain ⟨sb2', hstr2b', hrb'⟩ := (h sb sb2 hrb μ).1 _ htr - obtain ⟨s2', hstr2', hrb2⟩ := SWBisimulation.follow_internal_fst lts r h hrb' hstr2 + obtain ⟨sb2, hstr2b, hrb⟩ := SWBisimulation.follow_internal_fst h hr hstr1 + obtain ⟨sb2', hstr2b', hrb'⟩ := (h hrb μ).1 _ htr + obtain ⟨s2', hstr2', hrb2⟩ := SWBisimulation.follow_internal_fst h hrb' hstr2 exists s2' constructor · exact Lts.STr.comp lts hstr2b hstr2b' hstr2' @@ -973,32 +915,33 @@ theorem WeakBisimulation.iff_swBisimulation constructor; constructor exact hr case tr sb sb' hstr1 htr hstr2 => - obtain ⟨sb1, hstr1b, hrb⟩ := SWBisimulation.follow_internal_snd lts r h hr hstr1 - obtain ⟨sb2', hstr1b', hrb'⟩ := (h sb1 sb hrb μ).2 _ htr - obtain ⟨s1', hstr1', hrb2⟩ := SWBisimulation.follow_internal_snd lts r h hrb' hstr2 + obtain ⟨sb1, hstr1b, hrb⟩ := SWBisimulation.follow_internal_snd h hr hstr1 + obtain ⟨sb2', hstr1b', hrb'⟩ := (h hrb μ).2 _ htr + obtain ⟨s1', hstr1', hrb2⟩ := SWBisimulation.follow_internal_snd h hrb' hstr2 exists s1' constructor · exact Lts.STr.comp lts hstr1b hstr1b' hstr1' · exact hrb2 theorem WeakBisimulation.toSwBisimulation - [HasTau Label] {lts : Lts State Label} {r : State → State → Prop} (h : WeakBisimulation lts r) : - SWBisimulation lts r := (WeakBisimulation.iff_swBisimulation lts r).1 h + [HasTau Label] {lts : Lts State Label} {r : State → State → Prop} (h : lts.IsWeakBisimulation r) : + lts.IsSWBisimulation r := Lts.isWeakBisimulation_iff_isSWBisimulation.1 h theorem SWBisimulation.toWeakBisimulation - [HasTau Label] {lts : Lts State Label} {r : State → State → Prop} (h : SWBisimulation lts r) : - WeakBisimulation lts r := (WeakBisimulation.iff_swBisimulation lts r).2 h + [HasTau Label] {lts : Lts State Label} {r : State → State → Prop} (h : lts.IsSWBisimulation r) : + lts.IsWeakBisimulation r := Lts.isWeakBisimulation_iff_isSWBisimulation.2 h /-- If two states are related by an `SWBisimulation`, then they are weakly bisimilar. -/ theorem WeakBisimilarity.by_swBisimulation [HasTau Label] (lts : Lts State Label) (r : State → State → Prop) - (hb : SWBisimulation lts r) (hr : r s1 s2) : s1 ≈[lts] s2 := by + (hb : lts.IsSWBisimulation r) (hr : r s1 s2) : s1 ≈[lts] s2 := by exists r constructor · exact hr - apply (WeakBisimulation.iff_swBisimulation lts r).2 hb + apply Lts.isWeakBisimulation_iff_isSWBisimulation.2 hb /-- Weak bisimilarity and sw-bisimilarity coincide for all Ltss. -/ +@[grind _=_] theorem WeakBisimilarity.weakBisim_eq_swBisim [HasTau Label] (lts : Lts State Label) : WeakBisimilarity lts = SWBisimilarity lts := by funext s1 s2 @@ -1010,21 +953,21 @@ theorem WeakBisimilarity.weakBisim_eq_swBisim [HasTau Label] (lts : Lts State La exists r constructor · exact hr - apply (WeakBisimulation.iff_swBisimulation lts r).1 hrh + apply Lts.isWeakBisimulation_iff_isSWBisimulation.1 hrh case mpr => intro h obtain ⟨r, hr, hrh⟩ := h exists r constructor · exact hr - apply (WeakBisimulation.iff_swBisimulation lts r).2 hrh + apply Lts.isWeakBisimulation_iff_isSWBisimulation.2 hrh /-- sw-bisimilarity is reflexive. -/ theorem SWBisimilarity.refl [HasTau Label] (lts : Lts State Label) (s : State) : s ≈sw[lts] s := by exists Eq constructor · rfl - simp only [SWBisimulation] + simp only [Lts.IsSWBisimulation] intro s1 s2 hr μ cases hr constructor @@ -1048,29 +991,29 @@ theorem WeakBisimilarity.refl [HasTau Label] (lts : Lts State Label) (s : State) /-- The inverse of an sw-bisimulation is an sw-bisimulation. -/ theorem SWBisimulation.inv [HasTau Label] (lts : Lts State Label) - (r : State → State → Prop) (h : SWBisimulation lts r) : - SWBisimulation lts (flip r) := by - simp only [SWBisimulation] at h - simp only [SWBisimulation] + (r : State → State → Prop) (h : lts.IsSWBisimulation r) : + lts.IsSWBisimulation (flip r) := by + simp only [Lts.IsSWBisimulation] at h + simp only [Lts.IsSWBisimulation] intro s1 s2 hrinv μ constructor case left => intro s1' htr - specialize h s2 s1 hrinv μ + specialize h hrinv μ have h' := h.2 s1' htr obtain ⟨ s2', h' ⟩ := h' exists s2' case right => intro s2' htr - specialize h s2 s1 hrinv μ + specialize h hrinv μ have h' := h.1 s2' htr obtain ⟨ s1', h' ⟩ := h' exists s1' /-- The inverse of a weak bisimulation is a weak bisimulation. -/ theorem WeakBisimulation.inv [HasTau Label] (lts : Lts State Label) - (r : State → State → Prop) (h : WeakBisimulation lts r) : - WeakBisimulation lts (flip r) := by + (r : State → State → Prop) (h : lts.IsWeakBisimulation r) : + lts.IsWeakBisimulation (flip r) := by apply WeakBisimulation.toSwBisimulation at h have h' := SWBisimulation.inv lts r h apply SWBisimulation.toWeakBisimulation at h' @@ -1099,20 +1042,20 @@ theorem WeakBisimilarity.symm [HasTau Label] (lts : Lts State Label) (h : s1 ≈ theorem WeakBisimulation.comp [HasTau Label] (lts : Lts State Label) - (r1 r2 : State → State → Prop) (h1 : WeakBisimulation lts r1) (h2 : WeakBisimulation lts r2) : - WeakBisimulation lts (Relation.Comp r1 r2) := by - simp_all only [WeakBisimulation] - exact Bisimulation.comp lts.saturate r1 r2 h1 h2 + (r1 r2 : State → State → Prop) (h1 : lts.IsWeakBisimulation r1) (h2 : lts.IsWeakBisimulation r2) : + lts.IsWeakBisimulation (Relation.Comp r1 r2) := by + simp_all only [Lts.IsWeakBisimulation] + exact Bisimulation.comp h1 h2 /-- The composition of two sw-bisimulations is an sw-bisimulation. -/ theorem SWBisimulation.comp [HasTau Label] (lts : Lts State Label) - (r1 r2 : State → State → Prop) (h1 : SWBisimulation lts r1) (h2 : SWBisimulation lts r2) : - SWBisimulation lts (Relation.Comp r1 r2) := by + (r1 r2 : State → State → Prop) (h1 : lts.IsSWBisimulation r1) (h2 : lts.IsSWBisimulation r2) : + lts.IsSWBisimulation (Relation.Comp r1 r2) := by apply SWBisimulation.toWeakBisimulation at h1 apply SWBisimulation.toWeakBisimulation at h2 - apply (WeakBisimulation.iff_swBisimulation lts (Relation.Comp r1 r2)).1 + apply Lts.isWeakBisimulation_iff_isSWBisimulation.1 apply WeakBisimulation.comp lts r1 r2 h1 h2 /-- Weak bisimilarity is transitive. -/ diff --git a/Cslib/Languages/CCS/BehaviouralTheory.lean b/Cslib/Languages/CCS/BehaviouralTheory.lean index c2ae6221..0d7d97cc 100644 --- a/Cslib/Languages/CCS/BehaviouralTheory.lean +++ b/Cslib/Languages/CCS/BehaviouralTheory.lean @@ -61,7 +61,7 @@ theorem bisimilarity_par_comm : (par p q) ~[lts (defs := defs)] (par q p) := by case left => constructor case right => - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 hr μ cases hr case parComm p q => @@ -154,7 +154,7 @@ private inductive ChoiceComm : (Process Name Constant) → (Process Name Constan theorem bisimilarity_choice_comm : (choice p q) ~[lts (defs := defs)] (choice q p) := by exists @ChoiceComm Name Constant defs repeat constructor - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 hr μ cases hr case choiceComm => @@ -174,24 +174,9 @@ theorem bisimilarity_choice_comm : (choice p q) ~[lts (defs := defs)] (choice q constructor · unfold lts cases htr with grind [Tr.choiceR, Tr.choiceL] - · constructor - grind [Bisimilarity.refl] + · grind [ChoiceComm] case bisim h => - constructor - case left => - intro s1' htr - have hb := Bisimulation.follow_fst (Bisimilarity.is_bisimulation lts) h htr - obtain ⟨s2', htr2, hr2⟩ := hb - exists s2' - apply And.intro htr2 - constructor; assumption - case right => - intro s2' htr - have hb := Bisimulation.follow_snd (Bisimilarity.is_bisimulation lts) h htr - obtain ⟨s1', htr1, hr1⟩ := hb - exists s1' - apply And.intro htr1 - constructor; assumption + grind [ChoiceComm] /-- P + (Q + R) ~ (P + Q) + R -/ proof_wanted bisimilarity_choice_assoc : @@ -208,7 +193,7 @@ theorem bisimilarity_congr_pre : exists @PreBisim _ _ defs constructor · constructor; assumption - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 hr μ' cases hr case pre => @@ -236,19 +221,19 @@ theorem bisimilarity_congr_pre : exists s2' apply And.intro htr2 constructor - apply Bisimilarity.largest_bisimulation _ hbisim hr2 + apply Bisimilarity.largest_bisimulation hbisim hr2 case right => intro s2' htr obtain ⟨r, hr, hb⟩ := hbis let hbisim := hb - specialize hb _ _ hr μ' + specialize hb hr μ' obtain ⟨hb1, hb2⟩ := hb specialize hb2 _ htr obtain ⟨s1', htr1, hr1⟩ := hb2 exists s1' apply And.intro htr1 constructor - apply Bisimilarity.largest_bisimulation _ hbisim hr1 + apply Bisimilarity.largest_bisimulation hbisim hr1 private inductive ResBisim : (Process Name Constant) → (Process Name Constant) → Prop where | res : (p ~[lts (defs := defs)] q) → ResBisim (res a p) (res a q) @@ -261,7 +246,7 @@ theorem bisimilarity_congr_res : exists @ResBisim _ _ defs constructor · constructor; assumption - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 hr μ' cases hr rename_i p q a h @@ -270,7 +255,7 @@ theorem bisimilarity_congr_res : intro s1' htr cases htr rename_i p' h1 h2 htr - have h := Bisimulation.follow_fst (Bisimilarity.is_bisimulation lts) h htr + have h := Bisimilarity.is_bisimulation.follow_fst h htr obtain ⟨q', htrq, h⟩ := h exists (res a q') constructor; constructor; repeat assumption @@ -279,7 +264,7 @@ theorem bisimilarity_congr_res : intro s2' htr cases htr rename_i q' h1 h2 htr - have h := Bisimulation.follow_snd (Bisimilarity.is_bisimulation lts) h htr + have h := Bisimilarity.is_bisimulation.follow_snd h htr obtain ⟨p', htrq, h⟩ := h exists (res a p') constructor; constructor; repeat assumption @@ -296,7 +281,7 @@ theorem bisimilarity_congr_choice : exists @ChoiceBisim _ _ defs constructor · constructor; assumption - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 r μ constructor case left => @@ -311,7 +296,7 @@ theorem bisimilarity_congr_choice : constructor · apply Tr.choiceL htr2 · constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 case choiceR a b c htr => exists s1' constructor @@ -325,7 +310,7 @@ theorem bisimilarity_congr_choice : constructor · assumption constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 case right => intro s2' htr cases r @@ -338,7 +323,7 @@ theorem bisimilarity_congr_choice : constructor · apply Tr.choiceL htr1 · constructor - apply Bisimilarity.largest_bisimulation _ hb hr1 + apply Bisimilarity.largest_bisimulation hb hr1 case choiceR a b c htr => exists s2' constructor @@ -352,7 +337,7 @@ theorem bisimilarity_congr_choice : constructor · assumption · constructor - apply Bisimilarity.largest_bisimulation _ hb hr1 + apply Bisimilarity.largest_bisimulation hb hr1 private inductive ParBisim : (Process Name Constant) → (Process Name Constant) → Prop where | par : (p ~[lts (defs := defs)] q) → ParBisim (par p r) (par q r) @@ -364,7 +349,7 @@ theorem bisimilarity_congr_par : exists @ParBisim _ _ defs constructor · constructor; assumption - simp only [Bisimulation] + simp only [Lts.IsBisimulation] intro s1 s2 r μ constructor case left => @@ -379,20 +364,20 @@ theorem bisimilarity_congr_par : constructor · apply Tr.parL htr2 · constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 case parR _ _ r' htr => exists (par q r') constructor · apply Tr.parR htr · constructor - apply Bisimilarity.largest_bisimulation _ hb hr + apply Bisimilarity.largest_bisimulation hb hr case com μ' p' r' htrp htrr => obtain ⟨q', htr2, hr2⟩ := hb.follow_fst hr htrp exists (par q' r') constructor · apply Tr.com htr2 htrr · constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 case right => intro s2' htr cases r @@ -405,20 +390,20 @@ theorem bisimilarity_congr_par : constructor · apply Tr.parL htr2 · constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 case parR _ _ r' htr => exists (par p r') constructor · apply Tr.parR htr · constructor - apply Bisimilarity.largest_bisimulation _ hb hr + apply Bisimilarity.largest_bisimulation hb hr case com μ' p' r' htrq htrr => obtain ⟨q', htr2, hr2⟩ := hb.follow_snd hr htrq exists (par q' r') constructor · apply Tr.com htr2 htrr · constructor - apply Bisimilarity.largest_bisimulation _ hb hr2 + apply Bisimilarity.largest_bisimulation hb hr2 /-- Bisimilarity is a congruence in CCS. -/ theorem bisimilarity_congr From 0ca486847c91f7611a75688ea3457515080d96aa Mon Sep 17 00:00:00 2001 From: Fabrizio Montesi Date: Mon, 8 Sep 2025 10:47:26 +0200 Subject: [PATCH 08/19] Better support for dot-notation in IsBisimulation --- .../Semantics/Lts/Bisimulation.lean | 23 +++++++++---------- 1 file changed, 11 insertions(+), 12 deletions(-) diff --git a/Cslib/Foundations/Semantics/Lts/Bisimulation.lean b/Cslib/Foundations/Semantics/Lts/Bisimulation.lean index 2067fc2f..393818a5 100644 --- a/Cslib/Foundations/Semantics/Lts/Bisimulation.lean +++ b/Cslib/Foundations/Semantics/Lts/Bisimulation.lean @@ -50,13 +50,13 @@ related by some sw-bisimulation on `lts`. ## Main statements -- `Bisimulation.inv`: the inverse of a bisimulation is a bisimulation. +- `Lts.IsBisimulation.inv`: the inverse of a bisimulation is a bisimulation. - `Bisimilarity.eqv`: bisimilarity is an equivalence relation (see `Equivalence`). -- `Bisimilarity.is_bisimulation`: bisimilarity is itself a bisimulation. +- `Bisimilarity.isBisimulation`: bisimilarity is itself a bisimulation. - `Bisimilarity.largest_bisimulation`: bisimilarity is the largest bisimulation. - `Bisimilarity.gfp`: the union of bisimilarity and any bisimulation is equal to bisimilarity. -- `Bisimulation.upTo_bisimulation`: any bisimulation up to bisimilarity is a bisimulation. -- `Bisimulation.bisim_traceEq`: any bisimulation that relates two states implies that they are +- `Lts.IsBisimulationUpTo.isBisimulation`: any bisimulation up to bisimilarity is a bisimulation. +- `Lts.IsBisimulation.traceEq`: any bisimulation that relates two states implies that they are trace equivalent (see `TraceEq`). - `Bisimilarity.deterministic_bisim_eq_traceEq`: in a deterministic Lts, bisimilarity and trace equivalence coincide. @@ -129,7 +129,7 @@ theorem Bisimilarity.refl (s : State) : s ~[lts] s := by /-- The inverse of a bisimulation is a bisimulation. -/ @[grind] -theorem Bisimulation.inv (h : lts.IsBisimulation r) : +theorem Lts.IsBisimulation.inv (h : lts.IsBisimulation r) : lts.IsBisimulation (flip r) := by grind [flip] /-- Bisimilarity is symmetric. -/ @@ -141,7 +141,7 @@ theorem Bisimilarity.symm {s1 s2 : State} (h : s1 ~[lts] s2) : s2 ~[lts] s1 := b /-- The composition of two bisimulations is a bisimulation. -/ @[grind] -theorem Bisimulation.comp +theorem Lts.IsBisimulation.comp (h1 : lts.IsBisimulation r1) (h2 : lts.IsBisimulation r2) : lts.IsBisimulation (Relation.Comp r1 r2) := by grind [Relation.Comp] @@ -304,7 +304,7 @@ def Lts.IsBisimulationUpTo (lts : Lts State Label) (r : State → State → Prop /-- Any bisimulation up to bisimilarity is a bisimulation. -/ @[grind] -theorem Bisimulation.upTo_bisimulation (h : lts.IsBisimulationUpTo r) : +theorem Lts.IsBisimulationUpTo.isBisimulation (h : lts.IsBisimulationUpTo r) : lts.IsBisimulation (Relation.UpTo r (Bisimilarity lts)) := by intro s1 s2 hr μ rcases hr with ⟨s1b, hr1b, s2b, hrb, hr2b⟩ @@ -389,7 +389,7 @@ theorem Bisimulation.bisim_trace /-- Any bisimulation implies trace equivalence. -/ @[grind] -theorem Bisimulation.bisim_traceEq +theorem Lts.IsBisimulation.traceEq (hb : lts.IsBisimulation r) (hr : r s1 s2) : s1 ~tr[lts] s2 := by funext μs @@ -404,8 +404,7 @@ theorem Bisimulation.bisim_traceEq case mpr => intro h obtain ⟨s2', h⟩ := h - have hinv := @Bisimulation.inv State Label lts r hb - obtain ⟨s1', hmtr⟩ := Bisimulation.bisim_trace hinv hr μs s2' h + obtain ⟨s1', hmtr⟩ := Bisimulation.bisim_trace hb.inv hr μs s2' h exists s1' exact hmtr.1 @@ -414,7 +413,7 @@ theorem Bisimulation.bisim_traceEq theorem Bisimilarity.le_traceEq : Bisimilarity lts ≤ TraceEq lts := by intro s1 s2 h obtain ⟨r, hr, hb⟩ := h - apply Bisimulation.bisim_traceEq hb hr + apply hb.traceEq hr /- One of the standard motivating examples for bisimulation: `1` and `5` are trace equivalent, but not bisimilar. -/ @@ -1045,7 +1044,7 @@ theorem WeakBisimulation.comp (r1 r2 : State → State → Prop) (h1 : lts.IsWeakBisimulation r1) (h2 : lts.IsWeakBisimulation r2) : lts.IsWeakBisimulation (Relation.Comp r1 r2) := by simp_all only [Lts.IsWeakBisimulation] - exact Bisimulation.comp h1 h2 + exact h1.comp h2 /-- The composition of two sw-bisimulations is an sw-bisimulation. -/ theorem SWBisimulation.comp From 9de7e5f464b001d250df0fb13f8990fff4d5a237 Mon Sep 17 00:00:00 2001 From: twwar Date: Tue, 22 Jul 2025 10:56:30 +0200 Subject: [PATCH 09/19] reduction system attribute --- .../Semantics/ReductionSystem/Basic.lean | 18 +-- Cslib/Languages/CombinatoryLogic/Basic.lean | 62 +++++------ .../CombinatoryLogic/Confluence.lean | 28 ++--- Cslib/Languages/CombinatoryLogic/Defs.lean | 64 +++++------ .../Languages/CombinatoryLogic/Recursion.lean | 103 +++++++++--------- 5 files changed, 132 insertions(+), 143 deletions(-) diff --git a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean index a3742eeb..dcf7e9ee 100644 --- a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean +++ b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean @@ -46,11 +46,15 @@ open Relation Relation.ReflTransGen instance (rs : ReductionSystem Term) : Trans rs.Red rs.MRed rs.MRed := by infer_instance instance (rs : ReductionSystem Term) : Trans rs.MRed rs.Red rs.MRed := by infer_instance +instance (rs : ReductionSystem Term) : IsTrans Term rs.MRed := by infer_instance +instance (rs : ReductionSystem Term) : Transitive rs.MRed := transitive_of_trans rs.MRed +instance (rs : ReductionSystem Term) : Trans rs.MRed rs.MRed rs.MRed := instTransOfIsTrans + end MultiStep open Lean Elab Meta Command Term --- thank you to Kyle Miller for this: +-- thank you to Kyle Miller for this: -- https://leanprover.zulipchat.com/#narrow/channel/239415-metaprogramming-.2F-tactics/topic/Working.20with.20variables.20in.20a.20command/near/529324084 /-- A command to create a `ReductionSystem` from a relation, robust to use of `variable `-/ @@ -81,10 +85,10 @@ elab "create_reduction_sys" rel:ident name:ident : command => do } addTermInfo' name (.const name.getId params) (isBinder := true) addDeclarationRangesFromSyntax name.getId name - -/-- + +/-- This command adds notations for a `ReductionSystem.Red` and `ReductionSystem.MRed`. This should - not usually be called directly, but from the `reduction_sys` attribute. + not usually be called directly, but from the `reduction_sys` attribute. As an example `reduction_notation foo "β"` will add the notations "⭢β" and "↠β". @@ -94,19 +98,19 @@ elab "create_reduction_sys" rel:ident name:ident : command => do -/ syntax "reduction_notation" ident (str)? : command macro_rules - | `(reduction_notation $rs $sym) => + | `(reduction_notation $rs $sym) => `( notation3 t:39 " ⭢" $sym:str t':39 => (ReductionSystem.Red $rs) t t' notation3 t:39 " ↠" $sym:str t':39 => (ReductionSystem.MRed $rs) t t' ) - | `(reduction_notation $rs) => + | `(reduction_notation $rs) => `( notation3 t:39 " ⭢" t':39 => (ReductionSystem.Red $rs) t t' notation3 t:39 " ↠" t':39 => (ReductionSystem.MRed $rs) t t' ) -/-- +/-- This attribute calls the `reduction_notation` command for the annotated declaration, such as in: ``` diff --git a/Cslib/Languages/CombinatoryLogic/Basic.lean b/Cslib/Languages/CombinatoryLogic/Basic.lean index c37aa4bd..4dc3b2cb 100644 --- a/Cslib/Languages/CombinatoryLogic/Basic.lean +++ b/Cslib/Languages/CombinatoryLogic/Basic.lean @@ -61,7 +61,7 @@ def Polynomial.eval {n : Nat} (Γ : SKI.Polynomial n) (l : List SKI) (hl : List. def Polynomial.varFreeToSKI (Γ : SKI.Polynomial 0) : SKI := Γ.eval [] (by trivial) /-- Inductively define a polynomial `Γ'` so that (up to the fact that we haven't -defined reduction on polynomials) `Γ' ⬝ t ⇒* Γ[xₙ ← t]`. -/ +defined reduction on polynomials) `Γ' ⬝ t ↠ Γ[xₙ ← t]`. -/ def Polynomial.elimVar {n : Nat} : SKI.Polynomial (n+1) → SKI.Polynomial n /- The K-combinator leaves plain terms unchanged by substitution `K ⬝ x ⬝ t ⇒ x` -/ | SKI.Polynomial.term x => K ⬝' x @@ -83,7 +83,7 @@ for the inner variables. -/ theorem Polynomial.elimVar_correct {n : Nat} (Γ : SKI.Polynomial (n + 1)) {ys : List SKI} (hys : ys.length = n) (z : SKI) : - Γ.elimVar.eval ys hys ⬝ z ⇒* Γ.eval (ys ++ [z]) + Γ.elimVar.eval ys hys ⬝ z ↠ Γ.eval (ys ++ [z]) (by rw [List.length_append, hys, List.length_singleton]) := by match n, Γ with @@ -125,13 +125,13 @@ def Polynomial.toSKI {n : Nat} (Γ : SKI.Polynomial n) : SKI := /-- Correctness for the toSKI (bracket abstraction) algorithm. -/ theorem Polynomial.toSKI_correct {n : Nat} (Γ : SKI.Polynomial n) (xs : List SKI) - (hxs : xs.length = n) : Γ.toSKI.applyList xs ⇒* Γ.eval xs hxs := by + (hxs : xs.length = n) : Γ.toSKI.applyList xs ↠ Γ.eval xs hxs := by match n with | 0 => unfold toSKI varFreeToSKI applyList rw [List.length_eq_zero_iff] at hxs simp_rw [hxs, List.foldl_nil] - apply MRed.refl + rfl | n+1 => -- show that xs = ys + [z] have : xs ≠ [] := List.ne_nil_of_length_eq_add_one hxs @@ -165,7 +165,7 @@ choose a descriptive name. def RPoly : SKI.Polynomial 2 := &1 ⬝' &0 /-- A SKI term representing R -/ def R : SKI := RPoly.toSKI -theorem R_def (x y : SKI) : R ⬝ x ⬝ y ⇒* y ⬝ x := +theorem R_def (x y : SKI) : R ⬝ x ⬝ y ↠ y ⬝ x := RPoly.toSKI_correct [x, y] (by simp) @@ -173,7 +173,7 @@ theorem R_def (x y : SKI) : R ⬝ x ⬝ y ⇒* y ⬝ x := def BPoly : SKI.Polynomial 3 := &0 ⬝' (&1 ⬝' &2) /-- A SKI term representing B -/ def B : SKI := BPoly.toSKI -theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ⇒* f ⬝ (g ⬝ x) := +theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ↠ f ⬝ (g ⬝ x) := BPoly.toSKI_correct [f, g, x] (by simp) @@ -181,7 +181,7 @@ theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ⇒* f ⬝ (g ⬝ x) := def CPoly : SKI.Polynomial 3 := &0 ⬝' &2 ⬝' &1 /-- A SKI term representing C -/ def C : SKI := CPoly.toSKI -theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ⇒* f ⬝ y ⬝ x := +theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ↠ f ⬝ y ⬝ x := CPoly.toSKI_correct [f, x, y] (by simp) @@ -189,7 +189,7 @@ theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ⇒* f ⬝ y ⬝ x := def RotRPoly : SKI.Polynomial 3 := &2 ⬝' &0 ⬝' &1 /-- A SKI term representing RotR -/ def RotR : SKI := RotRPoly.toSKI -theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ⇒* z ⬝ x ⬝ y := +theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ↠ z ⬝ x ⬝ y := RotRPoly.toSKI_correct [x, y, z] (by simp) @@ -197,7 +197,7 @@ theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ⇒* z ⬝ x ⬝ y := def RotLPoly : SKI.Polynomial 3 := &1 ⬝' &2 ⬝' &0 /-- A SKI term representing RotL -/ def RotL : SKI := RotLPoly.toSKI -theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ⇒* y ⬝ z ⬝ x := +theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ↠ y ⬝ z ⬝ x := RotLPoly.toSKI_correct [x, y, z] (by simp) @@ -205,7 +205,7 @@ theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ⇒* y ⬝ z ⬝ x := def δPoly : SKI.Polynomial 1 := &0 ⬝' &0 /-- A SKI term representing δ -/ def δ : SKI := δPoly.toSKI -theorem δ_def (x : SKI) : δ ⬝ x ⇒* x ⬝ x := +theorem δ_def (x : SKI) : δ ⬝ x ↠ x ⬝ x := δPoly.toSKI_correct [x] (by simp) @@ -213,7 +213,7 @@ theorem δ_def (x : SKI) : δ ⬝ x ⇒* x ⬝ x := def HPoly : SKI.Polynomial 2 := &0 ⬝' (&1 ⬝' &1) /-- A SKI term representing H -/ def H : SKI := HPoly.toSKI -theorem H_def (f x : SKI) : H ⬝ f ⬝ x ⇒* f ⬝ (x ⬝ x) := +theorem H_def (f x : SKI) : H ⬝ f ⬝ x ↠ f ⬝ (x ⬝ x) := HPoly.toSKI_correct [f, x] (by simp) @@ -221,7 +221,7 @@ theorem H_def (f x : SKI) : H ⬝ f ⬝ x ⇒* f ⬝ (x ⬝ x) := def YPoly : SKI.Polynomial 1 := H ⬝' &0 ⬝' (H ⬝' &0) /-- A SKI term representing Y -/ def Y : SKI := YPoly.toSKI -theorem Y_def (f : SKI) : Y ⬝ f ⇒* H ⬝ f ⬝ (H ⬝ f) := +theorem Y_def (f : SKI) : Y ⬝ f ↠ H ⬝ f ⬝ (H ⬝ f) := YPoly.toSKI_correct [f] (by simp) @@ -239,29 +239,29 @@ rather than up to a common reduct. An alternative is to use Turing's fixed-point (defined below). -/ def fixedPoint (f : SKI) : SKI := H ⬝ f ⬝ (H ⬝ f) -theorem fixedPoint_correct (f : SKI) : f.fixedPoint ⇒* f ⬝ f.fixedPoint := H_def f (H ⬝ f) +theorem fixedPoint_correct (f : SKI) : f.fixedPoint ↠ f ⬝ f.fixedPoint := H_def f (H ⬝ f) /-- Auxiliary definition for Turing's fixed-point combinator: ΘAux := λ x y. y (x x y) -/ def ΘAuxPoly : SKI.Polynomial 2 := &1 ⬝' (&0 ⬝' &0 ⬝' &1) /-- A term representing ΘAux -/ def ΘAux : SKI := ΘAuxPoly.toSKI -theorem ΘAux_def (x y : SKI) : ΘAux ⬝ x ⬝ y ⇒* y ⬝ (x ⬝ x ⬝ y) := +theorem ΘAux_def (x y : SKI) : ΘAux ⬝ x ⬝ y ↠ y ⬝ (x ⬝ x ⬝ y) := ΘAuxPoly.toSKI_correct [x, y] (by simp) /-- Turing's fixed-point combinator: Θ := (λ x y. y (x x y)) (λ x y. y (x x y)) -/ def Θ : SKI := ΘAux ⬝ ΘAux /-- A SKI term representing Θ -/ -theorem Θ_correct (f : SKI) : Θ ⬝ f ⇒* f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f +theorem Θ_correct (f : SKI) : Θ ⬝ f ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f /-! ### Church Booleans -/ /-- A term a represents the boolean value u if it is βη-equivalent to a standard Church boolean. -/ def IsBool (u : Bool) (a : SKI) : Prop := - ∀ x y : SKI, a ⬝ x ⬝ y ⇒* (if u then x else y) + ∀ x y : SKI, a ⬝ x ⬝ y ↠ (if u then x else y) -theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ⇒* a') (ha' : IsBool u a') : +theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠ a') (ha' : IsBool u a') : IsBool u a := by intro x y trans a' ⬝ x ⬝ y @@ -278,13 +278,13 @@ theorem TT_correct : IsBool true TT := fun x y ↦ MRed.K x y def FF : SKI := K ⬝ I theorem FF_correct : IsBool false FF := fun x y ↦ calc - FF ⬝ x ⬝ y ⇒ I ⬝ y := by apply red_head; exact red_K I x - _ ⇒ y := red_I y + FF ⬝ x ⬝ y ⭢ I ⬝ y := by apply red_head; exact red_K I x + _ ⭢ y := red_I y /-- Conditional: Cond x y b := if b then x else y -/ protected def Cond : SKI := RotR theorem cond_correct (a x y : SKI) (u : Bool) (h : IsBool u a) : - SKI.Cond ⬝ x ⬝ y ⬝ a ⇒* if u then x else y := by + SKI.Cond ⬝ x ⬝ y ⬝ a ↠ if u then x else y := by trans a ⬝ x ⬝ y · exact rotR_def x y a · exact h x y @@ -302,7 +302,7 @@ theorem neg_correct (a : SKI) (ua : Bool) (h : IsBool ua a) : IsBool (¬ ua) (SK def AndPoly : SKI.Polynomial 2 := SKI.Cond ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' FF ⬝' &0 /-- A SKI term representing And -/ protected def And : SKI := AndPoly.toSKI -theorem and_def (a b : SKI) : SKI.And ⬝ a ⬝ b ⇒* SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := +theorem and_def (a b : SKI) : SKI.And ⬝ a ⬝ b ↠ SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := AndPoly.toSKI_correct [a, b] (by simp) theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -321,7 +321,7 @@ theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool u def OrPoly : SKI.Polynomial 2 := SKI.Cond ⬝' TT ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' &0 /-- A SKI term representing Or -/ protected def Or : SKI := OrPoly.toSKI -theorem or_def (a b : SKI) : SKI.Or ⬝ a ⬝ b ⇒* SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := +theorem or_def (a b : SKI) : SKI.Or ⬝ a ⬝ b ↠ SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := OrPoly.toSKI_correct [a, b] (by simp) theorem or_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -350,22 +350,22 @@ def Fst : SKI := R ⬝ TT /-- Second projection -/ def Snd : SKI := R ⬝ FF -theorem fst_correct (a b : SKI) : Fst ⬝ (MkPair ⬝ a ⬝ b) ⇒* a := by calc - _ ⇒* SKI.Cond ⬝ a ⬝ b ⬝ TT := R_def _ _ - _ ⇒* a := cond_correct TT a b true TT_correct +theorem fst_correct (a b : SKI) : Fst ⬝ (MkPair ⬝ a ⬝ b) ↠ a := by calc + _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ TT := R_def _ _ + _ ↠ a := cond_correct TT a b true TT_correct -theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ⇒* b := by calc - _ ⇒* SKI.Cond ⬝ a ⬝ b ⬝ FF := R_def _ _ - _ ⇒* b := cond_correct FF a b false FF_correct +theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ↠ b := by calc + _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ FF := R_def _ _ + _ ↠ b := cond_correct FF a b false FF_correct /-- Unpaired f ⟨x, y⟩ := f x y, cf `Nat.unparied`. -/ def UnpairedPoly : SKI.Polynomial 2 := &0 ⬝' (Fst ⬝' &1) ⬝' (Snd ⬝' &1) /-- A term representing Unpaired -/ protected def Unpaired : SKI := UnpairedPoly.toSKI -theorem unpaired_def (f p : SKI) : SKI.Unpaired ⬝ f ⬝ p ⇒* f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := +theorem unpaired_def (f p : SKI) : SKI.Unpaired ⬝ f ⬝ p ↠ f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := UnpairedPoly.toSKI_correct [f, p] (by simp) -theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y) ⇒* f ⬝ x ⬝ y := by +theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y) ↠ f ⬝ x ⬝ y := by trans f ⬝ (Fst ⬝ (MkPair ⬝ x ⬝ y)) ⬝ (Snd ⬝ (MkPair ⬝ x ⬝ y)) · exact unpaired_def f _ · apply parallel_mRed @@ -377,7 +377,7 @@ theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x def PairPoly : SKI.Polynomial 3 := MkPair ⬝' (&0 ⬝' &2) ⬝' (&1 ⬝' &2) /-- A SKI term representing Pair -/ protected def Pair : SKI := PairPoly.toSKI -theorem pair_def (f g x : SKI) : SKI.Pair ⬝ f ⬝ g ⬝ x ⇒* MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := +theorem pair_def (f g x : SKI) : SKI.Pair ⬝ f ⬝ g ⬝ x ↠ MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := PairPoly.toSKI_correct [f, g, x] (by simp) end SKI diff --git a/Cslib/Languages/CombinatoryLogic/Confluence.lean b/Cslib/Languages/CombinatoryLogic/Confluence.lean index c015c975..d26e4cf9 100644 --- a/Cslib/Languages/CombinatoryLogic/Confluence.lean +++ b/Cslib/Languages/CombinatoryLogic/Confluence.lean @@ -9,11 +9,11 @@ import Cslib.Foundations.Data.Relation /-! # SKI reduction is confluent -This file proves the **Church-Rosser** theorem for the SKI calculus, that is, if `a ⇒* b` and -`a ⇒* c`, `b ⇒* d` and `c ⇒* d` for some term `d`. More strongly (though equivalently), we show +This file proves the **Church-Rosser** theorem for the SKI calculus, that is, if `a ↠ b` and +`a ↠ c`, `b ↠ d` and `c ↠ d` for some term `d`. More strongly (though equivalently), we show that the relation of having a common reduct is transitive — in the above situation, `a` and `b`, and `a` and `c` have common reducts, so the result implies the same of `b` and `c`. Note that -`CommonReduct` is symmetric (trivially) and reflexive (since `⇒*` is), so we in fact show that +`CommonReduct` is symmetric (trivially) and reflexive (since `↠` is), so we in fact show that `CommonReduct` is an equivalence. Our proof @@ -23,7 +23,7 @@ Chapter 4 of Peter Selinger's notes: ## Main definitions -- `ParallelReduction` : a relation `⇒ₚ` on terms such that `⇒ ⊆ ⇒ₚ ⊆ ⇒*`, allowing simultaneous +- `ParallelReduction` : a relation `⇒ₚ` on terms such that `⇒ ⊆ ⇒ₚ ⊆ ↠`, allowing simultaneous reduction on the head and tail of a term. ## Main results @@ -31,13 +31,13 @@ reduction on the head and tail of a term. - `parallelReduction_diamond` : parallel reduction satisfies the diamond property, that is, it is confluent in a single step. - `commonReduct_equivalence` : by a general result, the diamond property for `⇒ₚ` implies the same -for its reflexive-transitive closure. This closure is exactly `⇒*`, which implies the +for its reflexive-transitive closure. This closure is exactly `↠`, which implies the **Church-Rosser** theorem as sketched above. -/ namespace SKI -open Red MRed +open Red MRed ReductionSystem /-- A reduction step allowing simultaneous reduction of disjoint redexes -/ inductive ParallelReduction : SKI → SKI → Prop @@ -55,8 +55,8 @@ inductive ParallelReduction : SKI → SKI → Prop /-- Notation for parallel reduction -/ scoped infix:90 " ⇒ₚ " => ParallelReduction -/-- The inclusion `⇒ₚ ⊆ ⇒*` -/ -theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ⇒* a' := by +/-- The inclusion `⇒ₚ ⊆ ↠` -/ +theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ↠ a' := by cases h case refl => exact Relation.ReflTransGen.refl case par a a' b b' ha hb => @@ -68,7 +68,7 @@ theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ⇒* a' := case red_S a b c => exact Relation.ReflTransGen.single (red_S a b c) /-- The inclusion `⇒ ⊆ ⇒ₚ` -/ -theorem parallelReduction_of_red {a a' : SKI} (h : a ⇒ a') : a ⇒ₚ a' := by +theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢ a') : a ⇒ₚ a' := by cases h case red_S => apply ParallelReduction.red_S case red_K => apply ParallelReduction.red_K @@ -86,12 +86,12 @@ theorem parallelReduction_of_red {a a' : SKI} (h : a ⇒ a') : a ⇒ₚ a' := by `parallelReduction_of_red` imply that `⇒` and `⇒ₚ` have the same reflexive-transitive closure. -/ theorem reflTransGen_parallelReduction_mRed : - Relation.ReflTransGen ParallelReduction = MRed := by + Relation.ReflTransGen ParallelReduction = RedSKI.MRed := by ext a b constructor · apply Relation.reflTransGen_minimal - · exact MRed.reflexive - · exact MRed.transitive + · exact λ _ => by rfl + · exact instTransitiveMRed RedSKI · exact @mRed_of_parallelReduction · apply Relation.reflTransGen_minimal · exact Relation.reflexive_reflTransGen @@ -101,7 +101,7 @@ theorem reflTransGen_parallelReduction_mRed : /-! Irreducibility for the (partially applied) primitive combinators. -TODO: possibly these should be proven more generally (in another file) for `⇒*`. +TODO: possibly these should be proven more generally (in another file) for `↠`. -/ lemma I_irreducible (a : SKI) (h : I ⇒ₚ a) : a = I := by @@ -234,7 +234,7 @@ theorem commonReduct_equivalence : Equivalence CommonReduct := by exact join_parallelReduction_equivalence /-- The **Church-Rosser** theorem in the form it is usually stated. -/ -theorem MRed.diamond (a b c : SKI) (hab : a ⇒* b) (hac : a ⇒* c) : CommonReduct b c := by +theorem MRed.diamond (a b c : SKI) (hab : a ↠ b) (hac : a ↠ c) : CommonReduct b c := by apply commonReduct_equivalence.trans (y := a) · exact commonReduct_equivalence.symm (commonReduct_of_single hab) · exact commonReduct_of_single hac diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index 29a565ba..9b3bea37 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Waring -/ import Mathlib.Logic.Relation -import Cslib.Foundations.Data.Relation +import Cslib.Utils.Relation +import Cslib.Semantics.ReductionSystem.Basic /-! # SKI Combinatory Logic @@ -23,7 +24,7 @@ using the SKI basis. - `⬝` : application between SKI terms, - `⇒` : single-step reduction, -- `⇒*` : multi-step reduction, +- `↠` : multi-step reduction, ## References @@ -60,6 +61,7 @@ lemma applyList_concat (f : SKI) (ys : List SKI) (z : SKI) : /-! ### Reduction relations between SKI terms -/ /-- Single-step reduction of SKI terms -/ +@[reduction_sys RedSKI] inductive Red : SKI → SKI → Prop where /-- The operational semantics of the `S`, -/ | red_S (x y z : SKI) : Red (S ⬝ x ⬝ y ⬝ z) (x ⬝ z ⬝ (y ⬝ z)) @@ -72,60 +74,47 @@ inductive Red : SKI → SKI → Prop where /-- and tail of an SKI term. -/ | red_tail (x y y' : SKI) (_ : Red y y') : Red (x ⬝ y) (x ⬝ y') -/-- Notation for single-step reduction -/ -scoped infix:90 " ⇒ " => Red -/-- Multi-step reduction of SKI terms -/ -def MRed : SKI → SKI → Prop := Relation.ReflTransGen Red +open Red ReductionSystem -/-- Notation for multi-step reduction (by analogy with the Kleene star) -/ -scoped infix:90 " ⇒* " => MRed +theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) := MRed.single RedSKI <| red_S .. +theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. +theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. -open Red - -@[refl] -theorem MRed.refl (a : SKI) : a ⇒* a := Relation.ReflTransGen.refl - -theorem MRed.single {a b : SKI} (h : a ⇒ b) : a ⇒* b := Relation.ReflTransGen.single h - -theorem MRed.S (x y z : SKI) : MRed (S ⬝ x ⬝ y ⬝ z) (x ⬝ z ⬝ (y ⬝ z)) := MRed.single <| red_S .. -theorem MRed.K (x y : SKI) : MRed (K ⬝ x ⬝ y) x := MRed.single <| red_K .. -theorem MRed.I (x : SKI) : MRed (I ⬝ x) x := MRed.single <| red_I .. - -theorem MRed.head {a a' : SKI} (b : SKI) (h : a ⇒* a') : (a ⬝ b) ⇒* (a' ⬝ b) := by +theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠ a') : (a ⬝ b) ↠ (a' ⬝ b) := by induction h with | refl => apply MRed.refl | @tail a' a'' _ ha'' ih => apply Relation.ReflTransGen.tail (b := a' ⬝ b) ih exact Red.red_head a' a'' b ha'' -theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ⇒* b') : (a ⬝ b) ⇒* (a ⬝ b') := by +theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠ b') : (a ⬝ b) ↠ (a ⬝ b') := by induction h with | refl => apply MRed.refl | @tail b' b'' _ hb'' ih => apply Relation.ReflTransGen.tail (b := a ⬝ b') ih exact Red.red_tail a b' b'' hb'' -instance MRed.instTrans : IsTrans SKI MRed := Relation.instIsTransReflTransGen -theorem MRed.transitive : Transitive MRed := transitive_of_trans MRed +-- instance MRed.instTrans : IsTrans SKI MRed := Relation.instIsTransReflTransGen +-- theorem MRed.transitive : Transitive MRed := transitive_of_trans MRed -instance MRed.instIsRefl : IsRefl SKI MRed := Relation.instIsReflReflTransGen -theorem MRed.reflexive : Reflexive MRed := IsRefl.reflexive +-- instance MRed.instIsRefl : IsRefl SKI MRed := Relation.instIsReflReflTransGen +-- theorem MRed.reflexive : Reflexive MRed := IsRefl.reflexive -instance MRedTrans : Trans Red MRed MRed := - ⟨fun hab => Relation.ReflTransGen.trans (MRed.single hab)⟩ +-- instance MRedTrans : Trans Red MRed MRed := +-- ⟨fun hab => Relation.ReflTransGen.trans (MRed.single hab)⟩ -instance MRedRedTrans : Trans MRed Red MRed := - ⟨fun hab hbc => Relation.ReflTransGen.trans hab (MRed.single hbc)⟩ +-- instance MRedRedTrans : Trans MRed Red MRed := +-- ⟨fun hab hbc => Relation.ReflTransGen.trans hab (MRed.single hbc)⟩ -instance RedMRedTrans : Trans Red Red MRed := - ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ +-- instance RedMRedTrans : Trans Red Red MRed := +-- ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ -lemma parallel_mRed {a a' b b' : SKI} (ha : a ⇒* a') (hb : b ⇒* b') : - (a ⬝ b) ⇒* (a' ⬝ b') := +lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠ a') (hb : b ↠ b') : + (a ⬝ b) ↠ (a' ⬝ b') := Trans.simple (MRed.head b ha) (MRed.tail a' hb) -lemma parallel_red {a a' b b' : SKI} (ha : a ⇒ a') (hb : b ⇒ b') : (a ⬝ b) ⇒* (a' ⬝ b') := by +lemma parallel_red {a a' b b' : SKI} (ha : a ⭢ a') (hb : b ⭢ b') : (a ⬝ b) ↠ (a' ⬝ b') := by trans a' ⬝ b all_goals apply MRed.single · exact Red.red_head a a' b ha @@ -133,13 +122,10 @@ lemma parallel_red {a a' b b' : SKI} (ha : a ⇒ a') (hb : b ⇒ b') : (a ⬝ b) /-- Express that two terms have a reduce to a common term. -/ -def CommonReduct : SKI → SKI → Prop := Relation.Join MRed +def CommonReduct : SKI → SKI → Prop := Relation.Join RedSKI.MRed -lemma commonReduct_of_single {a b : SKI} (h : a ⇒* b) : CommonReduct a b := by - refine Relation.join_of_single MRed.reflexive h +lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨b, h, by rfl⟩ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join theorem reflexive_commonReduct : Reflexive CommonReduct := Relation.reflexive_join MRed.reflexive - -end SKI diff --git a/Cslib/Languages/CombinatoryLogic/Recursion.lean b/Cslib/Languages/CombinatoryLogic/Recursion.lean index 1c5e3ef8..bbe5f533 100644 --- a/Cslib/Languages/CombinatoryLogic/Recursion.lean +++ b/Cslib/Languages/CombinatoryLogic/Recursion.lean @@ -15,17 +15,16 @@ formalisation of `Mathlib.Computability.Partrec`. Since composition (`B`-combina what remains are the following definitions and proofs of their correctness. - Church numerals : a predicate `IsChurch n a` expressing that the term `a` is βη-equivalent to -the standard church numeral `n` — that is, `a ⬝ f ⬝ x ⇒* f ⬝ (f ⬝ ... ⬝ (f ⬝ x)))`. +the standard church numeral `n` — that is, `a ⬝ f ⬝ x ↠ f ⬝ (f ⬝ ... ⬝ (f ⬝ x)))`. - SKI numerals : `Zero` and `Succ`, corresponding to `Partrec.zero` and `Partrec.succ`, and correctness proofs `zero_correct` and `succ_correct`. - Predecessor : a term `Pred` so that (`pred_correct`) `IsChurch n a → IsChurch n.pred (Pred ⬝ a)`. - Primitive recursion : a term `Rec` so that (`rec_correct_succ`) `IsChurch (n+1) a` implies -`Rec ⬝ x ⬝ g ⬝ a ⇒* g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` -implies -`Rec ⬝ x ⬝ g ⬝ a ⇒* x`. +`Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` implies +`Rec ⬝ x ⬝ g ⬝ a ↠ x`. - Unbounded root finding (μ-recursion) : given a term `f` representing a function `fℕ: Nat → Nat`, -which takes on the value 0 a term `RFind` such that (`rFind_correct`) `RFind ⬝ f ⇒* a` such that +which takes on the value 0 a term `RFind` such that (`rFind_correct`) `RFind ⬝ f ↠ a` such that `IsChurch n a` for `n` the smallest root of `fℕ`. ## References @@ -50,7 +49,7 @@ sense of `Mathlib.Data.Part` (as used in `Mathlib.Computability.Partrec`). namespace SKI -open Red MRed +open Red MRed ReductionSystem /-- Function form of the church numerals. -/ def Church (n : Nat) (f x : SKI) : SKI := @@ -59,22 +58,22 @@ match n with | n+1 => f ⬝ (Church n f x) /-- `church` commutes with reduction. -/ -lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ⇒* f') (hx : x ⇒* x') : - Church n f x ⇒* Church n f' x' := by +lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠ f') (hx : x ↠ x') : + Church n f x ↠ Church n f' x' := by induction n with | zero => exact hx | succ n ih => exact parallel_mRed hf ih /-- The term `a` is βη-equivalent to a standard church numeral. -/ -def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x : SKI, a ⬝ f ⬝ x ⇒* Church n f x +def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x : SKI, a ⬝ f ⬝ x ↠ Church n f x /-- To show `IsChurch n a` it suffices to show the same for a reduct of `a`. -/ -theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ⇒* a') : IsChurch n a' → IsChurch n a := by +theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠ a') : IsChurch n a' → IsChurch n a := by simp_rw [IsChurch] intro ha' f x calc - _ ⇒* a' ⬝ f ⬝ x := by apply MRed.head; apply MRed.head; exact h - _ ⇒* Church n f x := by apply ha' + _ ↠ a' ⬝ f ⬝ x := by apply MRed.head; apply MRed.head; exact h + _ ↠ Church n f x := by apply ha' /-! ### Church numeral basics -/ @@ -85,24 +84,24 @@ theorem zero_correct : IsChurch 0 SKI.Zero := by unfold IsChurch SKI.Zero Church intro f x calc - _ ⇒ I ⬝ x := by apply red_head; apply red_K - _ ⇒ x := by apply red_I + _ ⭢ I ⬝ x := by apply red_head; apply red_K + _ ⭢ x := by apply red_I /-- Church one := λ f x. f x -/ protected def One : SKI := I theorem one_correct : IsChurch 1 SKI.One := by intro f x apply head - exact single (red_I f) + exact MRed.single RedSKI (red_I f) /-- Church succ := λ a f x. f (a f x) ~ λ a f. B f (a f) ~ λ a. S B a ~ S B -/ protected def Succ : SKI := S ⬝ B theorem succ_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch (n+1) (SKI.Succ ⬝ a) := by intro f x calc - _ ⇒ B ⬝ f ⬝ (a ⬝ f) ⬝ x := by apply red_head; apply red_S - _ ⇒* f ⬝ (a ⬝ f ⬝ x) := by apply B_def - _ ⇒* f ⬝ (Church n f x) := by apply MRed.tail; exact h f x + _ ⭢ B ⬝ f ⬝ (a ⬝ f) ⬝ x := by apply red_head; apply red_S + _ ↠ f ⬝ (a ⬝ f ⬝ x) := by apply B_def + _ ↠ f ⬝ (Church n f x) := by apply MRed.tail; exact h f x /-- To define the predecessor, iterate the function `PredAux` ⟨i, j⟩ ↦ ⟨j, j+1⟩ on ⟨0,0⟩, then take @@ -111,14 +110,14 @@ the first component. def PredAuxPoly : SKI.Polynomial 1 := MkPair ⬝' (Snd ⬝' &0) ⬝' (SKI.Succ ⬝' (Snd ⬝' &0)) /-- A term representing PredAux -/ def PredAux : SKI := PredAuxPoly.toSKI -theorem predAux_def (p : SKI) : PredAux ⬝ p ⇒* MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := +theorem predAux_def (p : SKI) : PredAux ⬝ p ↠ MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := PredAuxPoly.toSKI_correct [p] (by simp) /-- Useful auxiliary definition expressing that `p` represents ns ∈ Nat × Nat. -/ def IsChurchPair (ns : Nat × Nat) (x : SKI) : Prop := IsChurch ns.1 (Fst ⬝ x) ∧ IsChurch ns.2 (Snd ⬝ x) -theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ⇒* a') : +theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠ a') : IsChurchPair ns a' → IsChurchPair ns a := by simp_rw [IsChurchPair] intro ⟨ha₁,ha₂⟩ @@ -156,7 +155,7 @@ theorem predAux_correct' (n : Nat) : def PredPoly : SKI.Polynomial 1 := Fst ⬝' (&0 ⬝' PredAux ⬝' (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) /-- A term representing Pred -/ def Pred : SKI := PredPoly.toSKI -theorem pred_def (a : SKI) : Pred ⬝ a ⇒* Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := +theorem pred_def (a : SKI) : Pred ⬝ a ↠ Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := PredPoly.toSKI_correct [a] (by simp) theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (Pred ⬝ a) := by @@ -174,7 +173,7 @@ theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (P def IsZeroPoly : SKI.Polynomial 1 := &0 ⬝' (K ⬝ FF) ⬝' TT /-- A term representing IsZero -/ def IsZero : SKI := IsZeroPoly.toSKI -theorem isZero_def (a : SKI) : IsZero ⬝ a ⇒* a ⬝ (K ⬝ FF) ⬝ TT := +theorem isZero_def (a : SKI) : IsZero ⬝ a ↠ a ⬝ (K ⬝ FF) ⬝ TT := IsZeroPoly.toSKI_correct [a] (by simp) theorem isZero_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsBool (n = 0) (IsZero ⬝ a) := by @@ -191,8 +190,8 @@ theorem isZero_correct (n : Nat) (a : SKI) (h : IsChurch n a) : rw [hk] at h apply isBool_trans (ha' := FF_correct) calc - _ ⇒* (K ⬝ FF) ⬝ Church k (K ⬝ FF) TT := h _ _ - _ ⇒ FF := red_K _ _ + _ ↠ (K ⬝ FF) ⬝ Church k (K ⬝ FF) TT := h _ _ + _ ⭢ FF := red_K _ _ /-- @@ -204,33 +203,33 @@ def RecAuxPoly : SKI.Polynomial 4 := /-- A term representing RecAux -/ def RecAux : SKI := RecAuxPoly.toSKI theorem recAux_def (R₀ x g a : SKI) : - RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := + RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := RecAuxPoly.toSKI_correct [R₀, x, g, a] (by simp) /-- We define Rec so that -`Rec ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a)` +`Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a)` -/ def Rec : SKI := fixedPoint RecAux theorem rec_def (x g a : SKI) : - Rec ⬝ x ⬝ g ⬝ a ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc - _ ⇒* RecAux ⬝ Rec ⬝ x ⬝ g ⬝ a := by + Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc + _ ↠ RecAux ⬝ Rec ⬝ x ⬝ g ⬝ a := by apply MRed.head; apply MRed.head; apply MRed.head apply fixedPoint_correct - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := recAux_def Rec x g a + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := recAux_def Rec x g a -theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ⇒* x := by +theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ↠ x := by calc - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ - _ ⇒* if (Nat.beq 0 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ + _ ↠ if (Nat.beq 0 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by apply cond_correct exact isZero_correct 0 a ha -theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n + 1) a) : - Rec ⬝ x ⬝ g ⬝ a ⇒* g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by +theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n+1) a) : + Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by calc - _ ⇒* SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ - _ ⇒* if (Nat.beq (n+1) 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by + _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ + _ ↠ if (Nat.beq (n+1) 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by apply cond_correct exact isZero_correct (n+1) a ha @@ -247,19 +246,19 @@ def RFindAboveAuxPoly : SKI.Polynomial 3 := /-- A term representing RFindAboveAux -/ def RFindAboveAux : SKI := RFindAboveAuxPoly.toSKI lemma rfindAboveAux_def (R₀ f a : SKI) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := RFindAboveAuxPoly.toSKI_correct [R₀, a, f] (by trivial) theorem rfindAboveAux_base (R₀ f a : SKI) (hfa : IsChurch 0 (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* a := calc - _ ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ - _ ⇒* if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ a := calc + _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ + _ ↠ if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa -theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m + 1) (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ⇒* R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc - _ ⇒* SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ - _ ⇒* if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by +theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m+1) (f ⬝ a)) : + RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc + _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ + _ ↠ if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa @@ -311,15 +310,15 @@ theorem RFind_correct (fNat : Nat → Nat) (f : SKI) def AddPoly : SKI.Polynomial 2 := &0 ⬝' SKI.Succ ⬝' &1 /-- A term representing addition on church numerals -/ protected def Add : SKI := AddPoly.toSKI -theorem add_def (a b : SKI) : SKI.Add ⬝ a ⬝ b ⇒* a ⬝ SKI.Succ ⬝ b := +theorem add_def (a b : SKI) : SKI.Add ⬝ a ⬝ b ↠ a ⬝ SKI.Succ ⬝ b := AddPoly.toSKI_correct [a, b] (by simp) theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : IsChurch (n+m) (SKI.Add ⬝ a ⬝ b) := by refine isChurch_trans (n+m) (a' := Church n SKI.Succ b) ?_ ?_ · calc - _ ⇒* a ⬝ SKI.Succ ⬝ b := add_def a b - _ ⇒* Church n SKI.Succ b := ha SKI.Succ b + _ ↠ a ⬝ SKI.Succ ⬝ b := add_def a b + _ ↠ Church n SKI.Succ b := ha SKI.Succ b · clear ha induction n with | zero => simp_rw [Nat.zero_add, Church]; exact hb @@ -331,7 +330,7 @@ theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def MulPoly : SKI.Polynomial 2 := &0 ⬝' (SKI.Add ⬝' &1) ⬝' SKI.Zero /-- A term representing multiplication on church numerals -/ protected def Mul : SKI := MulPoly.toSKI -theorem mul_def (a b : SKI) : SKI.Mul ⬝ a ⬝ b ⇒* a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := +theorem mul_def (a b : SKI) : SKI.Mul ⬝ a ⬝ b ↠ a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := MulPoly.toSKI_correct [a, b] (by simp) theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m b) : @@ -349,15 +348,15 @@ theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m def SubPoly : SKI.Polynomial 2 := &1 ⬝' Pred ⬝' &0 /-- A term representing subtraction on church numerals -/ protected def Sub : SKI := SubPoly.toSKI -theorem sub_def (a b : SKI) : SKI.Sub ⬝ a ⬝ b ⇒* b ⬝ Pred ⬝ a := +theorem sub_def (a b : SKI) : SKI.Sub ⬝ a ⬝ b ↠ b ⬝ Pred ⬝ a := SubPoly.toSKI_correct [a, b] (by simp) theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : IsChurch (n-m) (SKI.Sub ⬝ a ⬝ b) := by refine isChurch_trans (n-m) (a' := Church m Pred a) ?_ ?_ · calc - _ ⇒* b ⬝ Pred ⬝ a := sub_def a b - _ ⇒* Church m Pred a := hb Pred a + _ ↠ b ⬝ Pred ⬝ a := sub_def a b + _ ↠ Church m Pred a := hb Pred a · clear hb induction m with | zero => simp_rw [Nat.sub_zero, Church]; exact ha @@ -369,7 +368,7 @@ theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def LEPoly : SKI.Polynomial 2 := IsZero ⬝' (SKI.Sub ⬝' &0 ⬝' &1) /-- A term representing comparison on church numerals -/ protected def LE : SKI := LEPoly.toSKI -theorem le_def (a b : SKI) : SKI.LE ⬝ a ⬝ b ⇒* IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := +theorem le_def (a b : SKI) : SKI.LE ⬝ a ⬝ b ↠ IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := LEPoly.toSKI_correct [a, b] (by simp) theorem le_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : From 9625db3a738d5d398d29c369be279166cad5e64d Mon Sep 17 00:00:00 2001 From: twwar Date: Tue, 22 Jul 2025 12:58:25 +0200 Subject: [PATCH 10/19] evaluation results --- Cslib/Languages/CombinatoryLogic/Defs.lean | 7 + .../CombinatoryLogic/Evaluation.lean | 179 ++++++++++++++++++ 2 files changed, 186 insertions(+) create mode 100644 Cslib/Languages/CombinatoryLogic/Evaluation.lean diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index 9b3bea37..fb8e4fa5 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -77,6 +77,13 @@ inductive Red : SKI → SKI → Prop where open Red ReductionSystem +lemma Red.ne {x y : SKI} : (x ⭢ y) → x ≠ y + | red_S _ _ _, h => by cases h + | red_K _ _, h => by cases h + | red_I _, h => by cases h + | red_head _ _ _ h', h => Red.ne h' (SKI.app.inj h).1 + | red_tail _ _ _ h', h => Red.ne h' (SKI.app.inj h).2 + theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) := MRed.single RedSKI <| red_S .. theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean new file mode 100644 index 00000000..7c73c3a9 --- /dev/null +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -0,0 +1,179 @@ +/- +Copyright (c) 2025 Thomas Waring. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Waring +-/ +import Cslib.Computability.CombinatoryLogic.Defs +import Cslib.Computability.CombinatoryLogic.Confluence + +/-! +# Evaluation results + +This file draws heavily from +-/ + +open SKI Red + +/-- The predicate that a term has no reducible sub-terms. -/ +def RedexFree : SKI → Prop + | S => True + | K => True + | I => True + | S ⬝ x => RedexFree x + | K ⬝ x => RedexFree x + | I ⬝ _ => False + | S ⬝ x ⬝ y => RedexFree x ∧ RedexFree y + | K ⬝ _ ⬝ _ => False + | I ⬝ _ ⬝ _ => False + | S ⬝ _ ⬝ _ ⬝ _ => False + | K ⬝ _ ⬝ _ ⬝ _ => False + | I ⬝ _ ⬝ _ ⬝ _ => False + | a ⬝ b ⬝ c ⬝ d ⬝ e => RedexFree (a ⬝ b ⬝ c ⬝ d) ∧ RedexFree e + +/-- +One-step evaluation as a function: either it returns a term that has been reduced by one step, +or a proof that the term is redex free. Uses normal-order reduction. +-/ +def evalStep : (x : SKI) → PLift (RedexFree x) ⊕ SKI + | S => Sum.inl (PLift.up trivial) + | K => Sum.inl (PLift.up trivial) + | I => Sum.inl (PLift.up trivial) + | S ⬝ x => match evalStep x with + | Sum.inl h => Sum.inl h + | Sum.inr x' => Sum.inr (S ⬝ x') + | K ⬝ x => match evalStep x with + | Sum.inl h => Sum.inl h + | Sum.inr x' => Sum.inr (K ⬝ x') + | I ⬝ x => Sum.inr x + | S ⬝ x ⬝ y => match evalStep x, evalStep y with + | Sum.inl h1, Sum.inl h2 => Sum.inl (.up ⟨h1.down, h2.down⟩) + | Sum.inl _, Sum.inr y' => Sum.inr (S ⬝ x ⬝ y') + | Sum.inr x', _ => Sum.inr (S ⬝ x' ⬝ y) + | K ⬝ x ⬝ _ => Sum.inr x + | I ⬝ x ⬝ y => Sum.inr (x ⬝ y) + | S ⬝ x ⬝ y ⬝ z => Sum.inr (x ⬝ z ⬝ (y ⬝ z)) + | K ⬝ x ⬝ _ ⬝ z => Sum.inr (x ⬝ z) + | I ⬝ x ⬝ y ⬝ z => Sum.inr (x ⬝ y ⬝ z) + | a ⬝ b ⬝ c ⬝ d ⬝ e => match evalStep (a ⬝ b ⬝ c ⬝ d), evalStep e with + | Sum.inl h1, Sum.inl h2 => Sum.inl (.up ⟨h1.down, h2.down⟩) + | Sum.inl _, Sum.inr e' => Sum.inr (a ⬝ b ⬝ c ⬝ d ⬝ e') + | Sum.inr abcd', _ => Sum.inr (abcd' ⬝ e) + +/-- The normal-order reduction implemented by `evalStep` indeed computes a one-step reduction. -/ +theorem evalStep_right_correct : (x y : SKI) → (evalStep x = Sum.inr y) → x ⭢ y + | S ⬝ x, a, h => + match hx : evalStep x with + | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h + | Sum.inr x' => by + simp only [evalStep, hx, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) + | K ⬝ x, a, h => + match hx : evalStep x with + | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h + | Sum.inr x' => by + simp only [evalStep, hx, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) + | I ⬝ x, a, h => Sum.inr.inj h ▸ red_I _ + | S ⬝ x ⬝ y, a, h => + match hx : evalStep x, hy : evalStep y with + | Sum.inl _, Sum.inl _ => by simp only [hx, hy, evalStep, reduceCtorEq] at h + | Sum.inl _, Sum.inr y' => by + simp only [hx, hy, evalStep, Sum.inr.injEq] at h + rw [←h] + exact .red_tail _ _ _ <| evalStep_right_correct _ _ hy + | Sum.inr x', _ => by + simp only [hx, hy, evalStep, Sum.inr.injEq] at h + rw [←h] + exact .red_head _ _ _ <| .red_tail _ _ _ <| evalStep_right_correct _ _ hx + | K ⬝ x ⬝ y, a, h => Sum.inr.inj h ▸ red_K .. + | I ⬝ x ⬝ y, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_I _) + | S ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ red_S .. + | K ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_K ..) + | I ⬝ x ⬝ y ⬝ z, a, h => Sum.inr.inj h ▸ (red_head _ _ _ <| red_head _ _ _ <| red_I _) + | a ⬝ b ⬝ c ⬝ d ⬝ e, x, h => + match habcd : evalStep (a ⬝ b ⬝ c ⬝ d), he : evalStep e with + | Sum.inl _, Sum.inl _ => by simp only [habcd, he, evalStep, reduceCtorEq] at h + | Sum.inl _, Sum.inr e' => by + simp only [habcd, he, evalStep, Sum.inr.injEq] at h + rw [←h] + exact red_tail _ _ _ <| evalStep_right_correct _ _ he + | Sum.inr abcd', _ => by + simp only [habcd, he, evalStep, Sum.inr.injEq] at h + rw [←h] + exact red_head _ _ _ <| evalStep_right_correct _ _ habcd + +theorem redexFree_of_no_red {x : SKI} (h : ∀ y, ¬ (x ⭢ y)) : RedexFree x := by + match hx : evalStep x with + | Sum.inl h' => exact h'.down + | Sum.inr y => cases h _ (evalStep_right_correct x y hx) + +theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) +| S ⬝ x, hx, S ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' +| K ⬝ x, hx, K ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' +| S ⬝ _ ⬝ _, ⟨hx, _⟩, S ⬝ _ ⬝ _, red_head _ _ _ (red_tail _ _ _ h3) => hx.no_red _ h3 +| S ⬝ _ ⬝ _, ⟨_, hy⟩, S ⬝ _ ⬝ _, red_tail _ _ _ h3 => hy.no_red _ h3 +| _ ⬝ _ ⬝ _ ⬝ _ ⬝ _, ⟨hx, _⟩, _ ⬝ _, red_head _ _ _ hq => hx.no_red _ hq +| _ ⬝ _ ⬝ _ ⬝ _ ⬝ _, ⟨_, hy⟩, _ ⬝ _, red_tail _ _ _ he => hy.no_red _ he + +/-- A term is redex free iff it has no one-step reductions. -/ +theorem redexFree_iff {x : SKI} : RedexFree x ↔ ∀ y, ¬ (x ⭢ y) := + ⟨RedexFree.no_red, redexFree_of_no_red⟩ + +theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by + constructor + case mp => + intro h + match hx : evalStep x with + | Sum.inl h' => exact rfl + | Sum.inr y => cases h.no_red _ (evalStep_right_correct _ _ hx) + case mpr => + intro h + match hx : evalStep x with + | Sum.inl h' => exact h'.down + | Sum.inr y => rw [hx] at h; cases h + +instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_onceEval + +/-- A term is redex free iff its only many-step reduction is itself. -/ +theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := by + constructor + case mp => + intro h y + constructor + case mp => + intro h' + cases h'.cases_head + case inl => assumption + case inr h' => + obtain ⟨z, hz, _⟩ := h' + cases h.no_red _ hz + case mpr => + intro h + rw [h] + case mpr => + intro h + rw [redexFree_iff] + intro y hy + specialize h y + exact Red.ne hy (h.1 (Relation.ReflTransGen.single hy)) + +/-- If a term has a common reduct with a normal term, it in fact reduces to that term. -/ +theorem commonReduct_redexFree {x y : SKI} (hy : RedexFree y) (h : CommonReduct x y) : x ↠ y := + let ⟨w, hyw, hzw⟩ := h + (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw + +/-- If `x` reduces to both `y` and `z`, and `z` is not reducible, then `y` reduces to `z`. -/ +lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : RedexFree z) : y ↠ z := + let ⟨w, hyw, hzw⟩ := MRed.diamond x y z hxy hxz + (redexFree_iff'.1 hz _ |>.1 hzw : z = w) ▸ hyw + +/-- +If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. +-/ +lemma unique_normal_form {x y z : SKI} + (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := + (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) + +#check From f63529b2197b5ef946a913cfbe39ebf335da6385 Mon Sep 17 00:00:00 2001 From: twwar Date: Fri, 5 Sep 2025 16:46:47 +0200 Subject: [PATCH 11/19] rice's theorem --- Cslib/Languages/CombinatoryLogic/Defs.lean | 16 ++- .../CombinatoryLogic/Evaluation.lean | 115 +++++++++++++++++- 2 files changed, 128 insertions(+), 3 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index fb8e4fa5..583bf495 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -57,6 +57,12 @@ lemma applyList_concat (f : SKI) (ys : List SKI) (z : SKI) : f.applyList (ys ++ [z]) = f.applyList ys ⬝ z := by simp [applyList] +/-- The size of an SKI term is its number of combinators. -/ +def size : SKI → Nat + | S => 1 + | K => 1 + | I => 1 + | x ⬝ y => size x + size y /-! ### Reduction relations between SKI terms -/ @@ -134,5 +140,11 @@ def CommonReduct : SKI → SKI → Prop := Relation.Join RedSKI.MRed lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨b, h, by rfl⟩ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join -theorem reflexive_commonReduct : Reflexive CommonReduct := - Relation.reflexive_join MRed.reflexive +theorem reflexive_commonReduct : Reflexive CommonReduct := λ x => by + refine ⟨x,?_,?_⟩ <;> rfl + +theorem commonReduct_head {x x' : SKI} (y : SKI) : CommonReduct x x' → CommonReduct (x ⬝ y) (x' ⬝ y) + | ⟨z, hz, hz'⟩ => ⟨z ⬝ y, MRed.head y hz, MRed.head y hz'⟩ + +theorem commonReduct_tail (x : SKI) {y y' : SKI} : CommonReduct y y' → CommonReduct (x ⬝ y) (x ⬝ y') + | ⟨z, hz, hz'⟩ => ⟨x ⬝ z, MRed.tail x hz, MRed.tail x hz'⟩ diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 7c73c3a9..6e7bc7f8 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Waring -/ import Cslib.Computability.CombinatoryLogic.Defs +import Cslib.Computability.CombinatoryLogic.Basic import Cslib.Computability.CombinatoryLogic.Confluence +import Cslib.Computability.CombinatoryLogic.Recursion +import Mathlib.Tactic.Common /-! # Evaluation results @@ -176,4 +179,114 @@ lemma unique_normal_form {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) -#check +lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) + (hx : RedexFree x) (hy : RedexFree y) : x = y := + (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) + +/-! ### Injectivity for datatypes -/ + +lemma sk_nequiv : ¬ CommonReduct S K := by + intro ⟨z, hsz, hkz⟩ + have hS : RedexFree S := by simp [RedexFree] + have hK : RedexFree K := by simp [RedexFree] + cases (redexFree_iff'.1 hS z).1 hsz + cases (redexFree_iff'.1 hK _).1 hkz + +/-- Injectivity for booleans. -/ +theorem isBool_injective (x y : SKI) (u v : Bool) (hx : IsBool u x) (hy : IsBool v y) + (hxy : CommonReduct x y) : u = v := by + have h : CommonReduct (if u then S else K) (if v then S else K) := by + apply commonReduct_equivalence.trans (y := x ⬝ S ⬝ K) + · apply commonReduct_equivalence.symm + apply commonReduct_of_single + exact hx S K + · apply commonReduct_equivalence.trans (y := y ⬝ S ⬝ K) + · exact commonReduct_head K <| commonReduct_head S hxy + · apply commonReduct_of_single + exact hy S K + by_cases u + case pos hu => + by_cases v + case pos hv => + rw [hu, hv] + case neg hv => + simp_rw [hu, hv, Bool.false_eq_true, reduceIte] at h + exact False.elim <| sk_nequiv h + case neg hu => + by_cases v + case pos hv => + simp_rw [hu, hv, Bool.false_eq_true, reduceIte] at h + exact False.elim <| sk_nequiv (commonReduct_equivalence.symm h) + case neg hv => + simp_rw [hu, hv] + +lemma TF_nequiv : ¬ CommonReduct TT FF := fun h => + (Bool.eq_not_self true).mp <| isBool_injective TT FF true false TT_correct FF_correct h + +/-- A specialisation of `Church : Nat → SKI`. -/ +def churchK : Nat → SKI + | 0 => K + | n+1 => K ⬝ (churchK n) + +lemma churchK_church : (n : Nat) → churchK n = Church n K K + | 0 => rfl + | n+1 => by simp [churchK, Church, churchK_church n] + +lemma churchK_redexFree : (n : Nat) → RedexFree (churchK n) + | 0 => trivial + | n+1 => churchK_redexFree n + +@[simp] +lemma churchK_size : (n : Nat) → (churchK n).size = n+1 + | 0 => rfl + | n + 1 => by rw [churchK, size, size, churchK_size, Nat.add_comm] + +lemma churchK_injective : Function.Injective churchK := + fun n m h => by simpa using congrArg SKI.size h + +theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsChurch m y) + (hxy : CommonReduct x y) : n = m := by + suffices CommonReduct (churchK n) (churchK m) by + apply churchK_injective + exact unique_normal_form' this (churchK_redexFree n) (churchK_redexFree m) + apply commonReduct_equivalence.trans (y := x ⬝ K ⬝ K) + · simp_rw [churchK_church] + exact commonReduct_equivalence.symm <| commonReduct_of_single (hx K K) + · apply commonReduct_equivalence.trans (y := y ⬝ K ⬝ K) + · apply commonReduct_head; apply commonReduct_head; assumption + · simp_rw [churchK_church] + exact commonReduct_of_single (hy K K) + +/-- **Rice's theorem**: no SKI term is a non-trivial predicate. -/ +theorem rice {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) + (hxt : ∃ x : SKI, P ⬝ x ↠ TT) (hxf : ∃ x : SKI, P ⬝ x ↠ FF) : False := by + obtain ⟨a, ha⟩ := hxt + obtain ⟨b, hb⟩ := hxf + let Neg : SKI := P ⬝' &0 ⬝' b ⬝' a |>.toSKI (n := 1) + let Abs : SKI := Neg.fixedPoint + have Neg_app : ∀ x : SKI, Neg ⬝ x ↠ P ⬝ x ⬝ b ⬝ a := + fun x => (P ⬝' &0 ⬝' b ⬝' a) |>.toSKI_correct (n := 1) [x] (by simp) + cases hP Abs + case inl h => + have : P ⬝ Abs ↠ FF := calc + _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct + _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app + _ ↠ P ⬝ (TT ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h + _ ↠ P ⬝ b := by apply MRed.tail; apply TT_correct + _ ↠ FF := hb + exact TF_nequiv <| MRed.diamond _ _ _ h this + case inr h => + have : P ⬝ Abs ↠ TT := calc + _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct + _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app + _ ↠ P ⬝ (FF ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h + _ ↠ P ⬝ a := by apply MRed.tail; apply FF_correct + _ ↠ TT := ha + exact TF_nequiv <| MRed.diamond _ _ _ this h + +/-- **Rice's theorem**: any SKI predicate is trivial. -/ +theorem rice' {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) : + (∀ x : SKI, P ⬝ x ↠ TT) ∨ (∀ x : SKI, P ⬝ x ↠ FF) := by + by_contra! h + obtain ⟨⟨a, ha⟩, b, hb⟩ := h + exact rice hP ⟨b, (hP _).resolve_right hb⟩ ⟨a, (hP _).resolve_left ha⟩ From 99b4d1935f20d6e61c6060a9421049ab351910d7 Mon Sep 17 00:00:00 2001 From: twwar Date: Fri, 5 Sep 2025 16:58:50 +0200 Subject: [PATCH 12/19] documentation --- .../CombinatoryLogic/Evaluation.lean | 24 +++++++++++++++++-- 1 file changed, 22 insertions(+), 2 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 6e7bc7f8..4a232f2c 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -12,6 +12,24 @@ import Mathlib.Tactic.Common /-! # Evaluation results +This file formalises evaluation and normal forms of SKI terms. + +## Main definitions + +- `RedexFree` : a predicate expressing that a term has no redexes +- `evalStep` : SKI-reduction as a function + +## Main results + +- `evalStep_right_correct` : correctness for `evalStep` +- `redexFree_iff` and `redexFree_iff'` : a term is redex free if and only if it has (respectively) +no one-step reductions, or if its only many step reduction is itself. +- `unique_normal_form` : if `x` reduces to both `y` and `z`, and both `y` and `z` are in normal +form, then they are equal. +- **Rice's theorem**: no SKI term is a non-trivial predicate. + +## References + This file draws heavily from -/ @@ -124,7 +142,7 @@ theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) theorem redexFree_iff {x : SKI} : RedexFree x ↔ ∀ y, ¬ (x ⭢ y) := ⟨RedexFree.no_red, redexFree_of_no_red⟩ -theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by +theorem redexFree_iff_evalStep {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by constructor case mp => intro h @@ -137,7 +155,7 @@ theorem redexFree_iff_onceEval {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = | Sum.inl h' => exact h'.down | Sum.inr y => rw [hx] at h; cases h -instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_onceEval +instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_evalStep /-- A term is redex free iff its only many-step reduction is itself. -/ theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := by @@ -179,6 +197,7 @@ lemma unique_normal_form {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) +/-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) (hx : RedexFree x) (hy : RedexFree y) : x = y := (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) @@ -244,6 +263,7 @@ lemma churchK_size : (n : Nat) → (churchK n).size = n+1 lemma churchK_injective : Function.Injective churchK := fun n m h => by simpa using congrArg SKI.size h +/-- Injectivity for Church numerals-/ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsChurch m y) (hxy : CommonReduct x y) : n = m := by suffices CommonReduct (churchK n) (churchK m) by From 58f5f24ebc8551784a1341f67088fc1990e3a4e4 Mon Sep 17 00:00:00 2001 From: twwar Date: Mon, 8 Sep 2025 11:54:13 +0200 Subject: [PATCH 13/19] fix imports, build --- Cslib/Languages/CombinatoryLogic/Basic.lean | 42 +++++++++---------- Cslib/Languages/CombinatoryLogic/Defs.lean | 4 +- .../CombinatoryLogic/Evaluation.lean | 22 +++++----- .../Languages/CombinatoryLogic/Recursion.lean | 34 ++++++++------- 4 files changed, 52 insertions(+), 50 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Basic.lean b/Cslib/Languages/CombinatoryLogic/Basic.lean index 4dc3b2cb..2e7212f2 100644 --- a/Cslib/Languages/CombinatoryLogic/Basic.lean +++ b/Cslib/Languages/CombinatoryLogic/Basic.lean @@ -83,7 +83,7 @@ for the inner variables. -/ theorem Polynomial.elimVar_correct {n : Nat} (Γ : SKI.Polynomial (n + 1)) {ys : List SKI} (hys : ys.length = n) (z : SKI) : - Γ.elimVar.eval ys hys ⬝ z ↠ Γ.eval (ys ++ [z]) + (Γ.elimVar.eval ys hys ⬝ z) ↠ Γ.eval (ys ++ [z]) (by rw [List.length_append, hys, List.length_singleton]) := by match n, Γ with @@ -165,7 +165,7 @@ choose a descriptive name. def RPoly : SKI.Polynomial 2 := &1 ⬝' &0 /-- A SKI term representing R -/ def R : SKI := RPoly.toSKI -theorem R_def (x y : SKI) : R ⬝ x ⬝ y ↠ y ⬝ x := +theorem R_def (x y : SKI) : (R ⬝ x ⬝ y) ↠ y ⬝ x := RPoly.toSKI_correct [x, y] (by simp) @@ -173,7 +173,7 @@ theorem R_def (x y : SKI) : R ⬝ x ⬝ y ↠ y ⬝ x := def BPoly : SKI.Polynomial 3 := &0 ⬝' (&1 ⬝' &2) /-- A SKI term representing B -/ def B : SKI := BPoly.toSKI -theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ↠ f ⬝ (g ⬝ x) := +theorem B_def (f g x : SKI) : (B ⬝ f ⬝ g ⬝ x) ↠ f ⬝ (g ⬝ x) := BPoly.toSKI_correct [f, g, x] (by simp) @@ -181,7 +181,7 @@ theorem B_def (f g x : SKI) : B ⬝ f ⬝ g ⬝ x ↠ f ⬝ (g ⬝ x) := def CPoly : SKI.Polynomial 3 := &0 ⬝' &2 ⬝' &1 /-- A SKI term representing C -/ def C : SKI := CPoly.toSKI -theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ↠ f ⬝ y ⬝ x := +theorem C_def (f x y : SKI) : (C ⬝ f ⬝ x ⬝ y) ↠ f ⬝ y ⬝ x := CPoly.toSKI_correct [f, x, y] (by simp) @@ -189,7 +189,7 @@ theorem C_def (f x y : SKI) : C ⬝ f ⬝ x ⬝ y ↠ f ⬝ y ⬝ x := def RotRPoly : SKI.Polynomial 3 := &2 ⬝' &0 ⬝' &1 /-- A SKI term representing RotR -/ def RotR : SKI := RotRPoly.toSKI -theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ↠ z ⬝ x ⬝ y := +theorem rotR_def (x y z : SKI) : (RotR ⬝ x ⬝ y ⬝ z) ↠ z ⬝ x ⬝ y := RotRPoly.toSKI_correct [x, y, z] (by simp) @@ -197,7 +197,7 @@ theorem rotR_def (x y z : SKI) : RotR ⬝ x ⬝ y ⬝ z ↠ z ⬝ x ⬝ y := def RotLPoly : SKI.Polynomial 3 := &1 ⬝' &2 ⬝' &0 /-- A SKI term representing RotL -/ def RotL : SKI := RotLPoly.toSKI -theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ↠ y ⬝ z ⬝ x := +theorem rotL_def (x y z : SKI) : (RotL ⬝ x ⬝ y ⬝ z) ↠ y ⬝ z ⬝ x := RotLPoly.toSKI_correct [x, y, z] (by simp) @@ -205,7 +205,7 @@ theorem rotL_def (x y z : SKI) : RotL ⬝ x ⬝ y ⬝ z ↠ y ⬝ z ⬝ x := def δPoly : SKI.Polynomial 1 := &0 ⬝' &0 /-- A SKI term representing δ -/ def δ : SKI := δPoly.toSKI -theorem δ_def (x : SKI) : δ ⬝ x ↠ x ⬝ x := +theorem δ_def (x : SKI) : (δ ⬝ x) ↠ x ⬝ x := δPoly.toSKI_correct [x] (by simp) @@ -213,7 +213,7 @@ theorem δ_def (x : SKI) : δ ⬝ x ↠ x ⬝ x := def HPoly : SKI.Polynomial 2 := &0 ⬝' (&1 ⬝' &1) /-- A SKI term representing H -/ def H : SKI := HPoly.toSKI -theorem H_def (f x : SKI) : H ⬝ f ⬝ x ↠ f ⬝ (x ⬝ x) := +theorem H_def (f x : SKI) : (H ⬝ f ⬝ x) ↠ f ⬝ (x ⬝ x) := HPoly.toSKI_correct [f, x] (by simp) @@ -221,7 +221,7 @@ theorem H_def (f x : SKI) : H ⬝ f ⬝ x ↠ f ⬝ (x ⬝ x) := def YPoly : SKI.Polynomial 1 := H ⬝' &0 ⬝' (H ⬝' &0) /-- A SKI term representing Y -/ def Y : SKI := YPoly.toSKI -theorem Y_def (f : SKI) : Y ⬝ f ↠ H ⬝ f ⬝ (H ⬝ f) := +theorem Y_def (f : SKI) : (Y ⬝ f) ↠ H ⬝ f ⬝ (H ⬝ f) := YPoly.toSKI_correct [f] (by simp) @@ -245,21 +245,21 @@ theorem fixedPoint_correct (f : SKI) : f.fixedPoint ↠ f ⬝ f.fixedPoint := H_ def ΘAuxPoly : SKI.Polynomial 2 := &1 ⬝' (&0 ⬝' &0 ⬝' &1) /-- A term representing ΘAux -/ def ΘAux : SKI := ΘAuxPoly.toSKI -theorem ΘAux_def (x y : SKI) : ΘAux ⬝ x ⬝ y ↠ y ⬝ (x ⬝ x ⬝ y) := +theorem ΘAux_def (x y : SKI) : (ΘAux ⬝ x ⬝ y) ↠ y ⬝ (x ⬝ x ⬝ y) := ΘAuxPoly.toSKI_correct [x, y] (by simp) /-- Turing's fixed-point combinator: Θ := (λ x y. y (x x y)) (λ x y. y (x x y)) -/ def Θ : SKI := ΘAux ⬝ ΘAux /-- A SKI term representing Θ -/ -theorem Θ_correct (f : SKI) : Θ ⬝ f ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f +theorem Θ_correct (f : SKI) : (Θ ⬝ f) ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAux f /-! ### Church Booleans -/ /-- A term a represents the boolean value u if it is βη-equivalent to a standard Church boolean. -/ def IsBool (u : Bool) (a : SKI) : Prop := - ∀ x y : SKI, a ⬝ x ⬝ y ↠ (if u then x else y) + ∀ x y : SKI, (a ⬝ x ⬝ y) ↠ (if u then x else y) theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠ a') (ha' : IsBool u a') : IsBool u a := by @@ -278,13 +278,13 @@ theorem TT_correct : IsBool true TT := fun x y ↦ MRed.K x y def FF : SKI := K ⬝ I theorem FF_correct : IsBool false FF := fun x y ↦ calc - FF ⬝ x ⬝ y ⭢ I ⬝ y := by apply red_head; exact red_K I x + (FF ⬝ x ⬝ y) ↠ I ⬝ y := by apply Relation.ReflTransGen.single; apply red_head; exact red_K I x _ ⭢ y := red_I y /-- Conditional: Cond x y b := if b then x else y -/ protected def Cond : SKI := RotR theorem cond_correct (a x y : SKI) (u : Bool) (h : IsBool u a) : - SKI.Cond ⬝ x ⬝ y ⬝ a ↠ if u then x else y := by + (SKI.Cond ⬝ x ⬝ y ⬝ a) ↠ if u then x else y := by trans a ⬝ x ⬝ y · exact rotR_def x y a · exact h x y @@ -302,7 +302,7 @@ theorem neg_correct (a : SKI) (ua : Bool) (h : IsBool ua a) : IsBool (¬ ua) (SK def AndPoly : SKI.Polynomial 2 := SKI.Cond ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' FF ⬝' &0 /-- A SKI term representing And -/ protected def And : SKI := AndPoly.toSKI -theorem and_def (a b : SKI) : SKI.And ⬝ a ⬝ b ↠ SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := +theorem and_def (a b : SKI) : (SKI.And ⬝ a ⬝ b) ↠ SKI.Cond ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ FF ⬝ a := AndPoly.toSKI_correct [a, b] (by simp) theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -321,7 +321,7 @@ theorem and_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool u def OrPoly : SKI.Polynomial 2 := SKI.Cond ⬝' TT ⬝' (SKI.Cond ⬝ TT ⬝ FF ⬝' &1) ⬝' &0 /-- A SKI term representing Or -/ protected def Or : SKI := OrPoly.toSKI -theorem or_def (a b : SKI) : SKI.Or ⬝ a ⬝ b ↠ SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := +theorem or_def (a b : SKI) : (SKI.Or ⬝ a ⬝ b) ↠ SKI.Cond ⬝ TT ⬝ (SKI.Cond ⬝ TT ⬝ FF ⬝ b) ⬝ a := OrPoly.toSKI_correct [a, b] (by simp) theorem or_correct (a b : SKI) (ua ub : Bool) (ha : IsBool ua a) (hb : IsBool ub b) : @@ -350,11 +350,11 @@ def Fst : SKI := R ⬝ TT /-- Second projection -/ def Snd : SKI := R ⬝ FF -theorem fst_correct (a b : SKI) : Fst ⬝ (MkPair ⬝ a ⬝ b) ↠ a := by calc +theorem fst_correct (a b : SKI) : (Fst ⬝ (MkPair ⬝ a ⬝ b)) ↠ a := by calc _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ TT := R_def _ _ _ ↠ a := cond_correct TT a b true TT_correct -theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ↠ b := by calc +theorem snd_correct (a b : SKI) : (Snd ⬝ (MkPair ⬝ a ⬝ b)) ↠ b := by calc _ ↠ SKI.Cond ⬝ a ⬝ b ⬝ FF := R_def _ _ _ ↠ b := cond_correct FF a b false FF_correct @@ -362,10 +362,10 @@ theorem snd_correct (a b : SKI) : Snd ⬝ (MkPair ⬝ a ⬝ b) ↠ b := by calc def UnpairedPoly : SKI.Polynomial 2 := &0 ⬝' (Fst ⬝' &1) ⬝' (Snd ⬝' &1) /-- A term representing Unpaired -/ protected def Unpaired : SKI := UnpairedPoly.toSKI -theorem unpaired_def (f p : SKI) : SKI.Unpaired ⬝ f ⬝ p ↠ f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := +theorem unpaired_def (f p : SKI) : (SKI.Unpaired ⬝ f ⬝ p) ↠ f ⬝ (Fst ⬝ p) ⬝ (Snd ⬝ p) := UnpairedPoly.toSKI_correct [f, p] (by simp) -theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y) ↠ f ⬝ x ⬝ y := by +theorem unpaired_correct (f x y : SKI) : (SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x ⬝ y)) ↠ f ⬝ x ⬝ y := by trans f ⬝ (Fst ⬝ (MkPair ⬝ x ⬝ y)) ⬝ (Snd ⬝ (MkPair ⬝ x ⬝ y)) · exact unpaired_def f _ · apply parallel_mRed @@ -377,7 +377,7 @@ theorem unpaired_correct (f x y : SKI) : SKI.Unpaired ⬝ f ⬝ (MkPair ⬝ x def PairPoly : SKI.Polynomial 3 := MkPair ⬝' (&0 ⬝' &2) ⬝' (&1 ⬝' &2) /-- A SKI term representing Pair -/ protected def Pair : SKI := PairPoly.toSKI -theorem pair_def (f g x : SKI) : SKI.Pair ⬝ f ⬝ g ⬝ x ↠ MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := +theorem pair_def (f g x : SKI) : (SKI.Pair ⬝ f ⬝ g ⬝ x) ↠ MkPair ⬝ (f ⬝ x) ⬝ (g ⬝ x) := PairPoly.toSKI_correct [f, g, x] (by simp) end SKI diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index 583bf495..a7637f2d 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Waring -/ import Mathlib.Logic.Relation -import Cslib.Utils.Relation -import Cslib.Semantics.ReductionSystem.Basic +import Cslib.Foundations.Data.Relation +import Cslib.Foundations.Semantics.ReductionSystem.Basic /-! # SKI Combinatory Logic diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 4a232f2c..907b36b6 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -3,10 +3,10 @@ Copyright (c) 2025 Thomas Waring. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Waring -/ -import Cslib.Computability.CombinatoryLogic.Defs -import Cslib.Computability.CombinatoryLogic.Basic -import Cslib.Computability.CombinatoryLogic.Confluence -import Cslib.Computability.CombinatoryLogic.Recursion +import Cslib.Languages.CombinatoryLogic.Defs +import Cslib.Languages.CombinatoryLogic.Basic +import Cslib.Languages.CombinatoryLogic.Confluence +import Cslib.Languages.CombinatoryLogic.Recursion import Mathlib.Tactic.Common /-! @@ -278,17 +278,17 @@ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsC exact commonReduct_of_single (hy K K) /-- **Rice's theorem**: no SKI term is a non-trivial predicate. -/ -theorem rice {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) - (hxt : ∃ x : SKI, P ⬝ x ↠ TT) (hxf : ∃ x : SKI, P ⬝ x ↠ FF) : False := by +theorem rice {P : SKI} (hP : ∀ x : SKI, ((P ⬝ x) ↠ TT) ∨ (P ⬝ x) ↠ FF) + (hxt : ∃ x : SKI, (P ⬝ x) ↠ TT) (hxf : ∃ x : SKI, (P ⬝ x) ↠ FF) : False := by obtain ⟨a, ha⟩ := hxt obtain ⟨b, hb⟩ := hxf let Neg : SKI := P ⬝' &0 ⬝' b ⬝' a |>.toSKI (n := 1) let Abs : SKI := Neg.fixedPoint - have Neg_app : ∀ x : SKI, Neg ⬝ x ↠ P ⬝ x ⬝ b ⬝ a := + have Neg_app : ∀ x : SKI, (Neg ⬝ x) ↠ P ⬝ x ⬝ b ⬝ a := fun x => (P ⬝' &0 ⬝' b ⬝' a) |>.toSKI_correct (n := 1) [x] (by simp) cases hP Abs case inl h => - have : P ⬝ Abs ↠ FF := calc + have : (P ⬝ Abs) ↠ FF := calc _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app _ ↠ P ⬝ (TT ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h @@ -296,7 +296,7 @@ theorem rice {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) _ ↠ FF := hb exact TF_nequiv <| MRed.diamond _ _ _ h this case inr h => - have : P ⬝ Abs ↠ TT := calc + have : (P ⬝ Abs) ↠ TT := calc _ ↠ P ⬝ (Neg ⬝ Abs) := by apply MRed.tail; apply fixedPoint_correct _ ↠ P ⬝ (P ⬝ Abs ⬝ b ⬝ a) := by apply MRed.tail; apply Neg_app _ ↠ P ⬝ (FF ⬝ b ⬝ a) := by apply MRed.tail; apply MRed.head; apply MRed.head; exact h @@ -305,8 +305,8 @@ theorem rice {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) exact TF_nequiv <| MRed.diamond _ _ _ this h /-- **Rice's theorem**: any SKI predicate is trivial. -/ -theorem rice' {P : SKI} (hP : ∀ x : SKI, (P ⬝ x ↠ TT) ∨ P ⬝ x ↠ FF) : - (∀ x : SKI, P ⬝ x ↠ TT) ∨ (∀ x : SKI, P ⬝ x ↠ FF) := by +theorem rice' {P : SKI} (hP : ∀ x : SKI, ((P ⬝ x) ↠ TT) ∨ (P ⬝ x) ↠ FF) : + (∀ x : SKI, (P ⬝ x) ↠ TT) ∨ (∀ x : SKI, (P ⬝ x) ↠ FF) := by by_contra! h obtain ⟨⟨a, ha⟩, b, hb⟩ := h exact rice hP ⟨b, (hP _).resolve_right hb⟩ ⟨a, (hP _).resolve_left ha⟩ diff --git a/Cslib/Languages/CombinatoryLogic/Recursion.lean b/Cslib/Languages/CombinatoryLogic/Recursion.lean index bbe5f533..154cdf98 100644 --- a/Cslib/Languages/CombinatoryLogic/Recursion.lean +++ b/Cslib/Languages/CombinatoryLogic/Recursion.lean @@ -65,7 +65,8 @@ lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠ f') (hx : x ↠ x') : | succ n ih => exact parallel_mRed hf ih /-- The term `a` is βη-equivalent to a standard church numeral. -/ -def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x : SKI, a ⬝ f ⬝ x ↠ Church n f x +def IsChurch (n : Nat) (a : SKI) : Prop := + ∀ f x :SKI, (a ⬝ f ⬝ x) ↠ (Church n f x) /-- To show `IsChurch n a` it suffices to show the same for a reduct of `a`. -/ theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠ a') : IsChurch n a' → IsChurch n a := by @@ -84,7 +85,7 @@ theorem zero_correct : IsChurch 0 SKI.Zero := by unfold IsChurch SKI.Zero Church intro f x calc - _ ⭢ I ⬝ x := by apply red_head; apply red_K + _ ↠ I ⬝ x := by apply Relation.ReflTransGen.single; apply red_head; apply red_K _ ⭢ x := by apply red_I /-- Church one := λ f x. f x -/ @@ -110,7 +111,7 @@ the first component. def PredAuxPoly : SKI.Polynomial 1 := MkPair ⬝' (Snd ⬝' &0) ⬝' (SKI.Succ ⬝' (Snd ⬝' &0)) /-- A term representing PredAux -/ def PredAux : SKI := PredAuxPoly.toSKI -theorem predAux_def (p : SKI) : PredAux ⬝ p ↠ MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := +theorem predAux_def (p : SKI) : (PredAux ⬝ p) ↠ MkPair ⬝ (Snd ⬝ p) ⬝ (SKI.Succ ⬝ (Snd ⬝ p)) := PredAuxPoly.toSKI_correct [p] (by simp) /-- Useful auxiliary definition expressing that `p` represents ns ∈ Nat × Nat. -/ @@ -155,7 +156,7 @@ theorem predAux_correct' (n : Nat) : def PredPoly : SKI.Polynomial 1 := Fst ⬝' (&0 ⬝' PredAux ⬝' (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) /-- A term representing Pred -/ def Pred : SKI := PredPoly.toSKI -theorem pred_def (a : SKI) : Pred ⬝ a ↠ Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := +theorem pred_def (a : SKI) : (Pred ⬝ a) ↠ Fst ⬝ (a ⬝ PredAux ⬝ (MkPair ⬝ SKI.Zero ⬝ SKI.Zero)) := PredPoly.toSKI_correct [a] (by simp) theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (Pred ⬝ a) := by @@ -173,7 +174,7 @@ theorem pred_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsChurch n.pred (P def IsZeroPoly : SKI.Polynomial 1 := &0 ⬝' (K ⬝ FF) ⬝' TT /-- A term representing IsZero -/ def IsZero : SKI := IsZeroPoly.toSKI -theorem isZero_def (a : SKI) : IsZero ⬝ a ↠ a ⬝ (K ⬝ FF) ⬝ TT := +theorem isZero_def (a : SKI) : (IsZero ⬝ a) ↠ a ⬝ (K ⬝ FF) ⬝ TT := IsZeroPoly.toSKI_correct [a] (by simp) theorem isZero_correct (n : Nat) (a : SKI) (h : IsChurch n a) : IsBool (n = 0) (IsZero ⬝ a) := by @@ -203,7 +204,8 @@ def RecAuxPoly : SKI.Polynomial 4 := /-- A term representing RecAux -/ def RecAux : SKI := RecAuxPoly.toSKI theorem recAux_def (R₀ x g a : SKI) : - RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := + (RecAux ⬝ R₀ ⬝ x ⬝ g ⬝ a) ↠ + SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (R₀ ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := RecAuxPoly.toSKI_correct [R₀, x, g, a] (by simp) /-- @@ -212,13 +214,13 @@ We define Rec so that -/ def Rec : SKI := fixedPoint RecAux theorem rec_def (x g a : SKI) : - Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc + (Rec ⬝ x ⬝ g ⬝ a) ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := calc _ ↠ RecAux ⬝ Rec ⬝ x ⬝ g ⬝ a := by apply MRed.head; apply MRed.head; apply MRed.head apply fixedPoint_correct _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := recAux_def Rec x g a -theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ↠ x := by +theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : (Rec ⬝ x ⬝ g ⬝ a) ↠ x := by calc _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ _ ↠ if (Nat.beq 0 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by @@ -226,7 +228,7 @@ theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : Rec ⬝ x ⬝ g ⬝ a ↠ x exact isZero_correct 0 a ha theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n+1) a) : - Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by + (Rec ⬝ x ⬝ g ⬝ a) ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by calc _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ _ ↠ if (Nat.beq (n+1) 0) then x else (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) := by @@ -246,17 +248,17 @@ def RFindAboveAuxPoly : SKI.Polynomial 3 := /-- A term representing RFindAboveAux -/ def RFindAboveAux : SKI := RFindAboveAuxPoly.toSKI lemma rfindAboveAux_def (R₀ f a : SKI) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := + (RFindAboveAux ⬝ R₀ ⬝ a ⬝ f) ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := RFindAboveAuxPoly.toSKI_correct [R₀, a, f] (by trivial) theorem rfindAboveAux_base (R₀ f a : SKI) (hfa : IsChurch 0 (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ a := calc + (RFindAboveAux ⬝ R₀ ⬝ a ⬝ f) ↠ a := calc _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ _ ↠ if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m+1) (f ⬝ a)) : - RFindAboveAux ⬝ R₀ ⬝ a ⬝ f ↠ R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc + (RFindAboveAux ⬝ R₀ ⬝ a ⬝ f) ↠ R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ _ ↠ if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct @@ -310,7 +312,7 @@ theorem RFind_correct (fNat : Nat → Nat) (f : SKI) def AddPoly : SKI.Polynomial 2 := &0 ⬝' SKI.Succ ⬝' &1 /-- A term representing addition on church numerals -/ protected def Add : SKI := AddPoly.toSKI -theorem add_def (a b : SKI) : SKI.Add ⬝ a ⬝ b ↠ a ⬝ SKI.Succ ⬝ b := +theorem add_def (a b : SKI) : (SKI.Add ⬝ a ⬝ b) ↠ a ⬝ SKI.Succ ⬝ b := AddPoly.toSKI_correct [a, b] (by simp) theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : @@ -330,7 +332,7 @@ theorem add_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def MulPoly : SKI.Polynomial 2 := &0 ⬝' (SKI.Add ⬝' &1) ⬝' SKI.Zero /-- A term representing multiplication on church numerals -/ protected def Mul : SKI := MulPoly.toSKI -theorem mul_def (a b : SKI) : SKI.Mul ⬝ a ⬝ b ↠ a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := +theorem mul_def (a b : SKI) : (SKI.Mul ⬝ a ⬝ b) ↠ a ⬝ (SKI.Add ⬝ b) ⬝ SKI.Zero := MulPoly.toSKI_correct [a, b] (by simp) theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m b) : @@ -348,7 +350,7 @@ theorem mul_correct {n m : Nat} {a b : SKI} (ha : IsChurch n a) (hb : IsChurch m def SubPoly : SKI.Polynomial 2 := &1 ⬝' Pred ⬝' &0 /-- A term representing subtraction on church numerals -/ protected def Sub : SKI := SubPoly.toSKI -theorem sub_def (a b : SKI) : SKI.Sub ⬝ a ⬝ b ↠ b ⬝ Pred ⬝ a := +theorem sub_def (a b : SKI) : (SKI.Sub ⬝ a ⬝ b) ↠ b ⬝ Pred ⬝ a := SubPoly.toSKI_correct [a, b] (by simp) theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : @@ -368,7 +370,7 @@ theorem sub_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m def LEPoly : SKI.Polynomial 2 := IsZero ⬝' (SKI.Sub ⬝' &0 ⬝' &1) /-- A term representing comparison on church numerals -/ protected def LE : SKI := LEPoly.toSKI -theorem le_def (a b : SKI) : SKI.LE ⬝ a ⬝ b ↠ IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := +theorem le_def (a b : SKI) : (SKI.LE ⬝ a ⬝ b) ↠ IsZero ⬝ (SKI.Sub ⬝ a ⬝ b) := LEPoly.toSKI_correct [a, b] (by simp) theorem le_correct (n m : Nat) (a b : SKI) (ha : IsChurch n a) (hb : IsChurch m b) : From acb0a0167bb56527e2759743ef33de3deffd2b01 Mon Sep 17 00:00:00 2001 From: twwar Date: Mon, 8 Sep 2025 20:54:12 +0200 Subject: [PATCH 14/19] linting --- Cslib/Languages/CombinatoryLogic/Basic.lean | 2 +- Cslib/Languages/CombinatoryLogic/Confluence.lean | 6 +++--- Cslib/Languages/CombinatoryLogic/Defs.lean | 14 ++++++++------ Cslib/Languages/CombinatoryLogic/Evaluation.lean | 6 +++--- Cslib/Languages/CombinatoryLogic/Recursion.lean | 14 +++++++------- 5 files changed, 22 insertions(+), 20 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Basic.lean b/Cslib/Languages/CombinatoryLogic/Basic.lean index 2e7212f2..837a6af4 100644 --- a/Cslib/Languages/CombinatoryLogic/Basic.lean +++ b/Cslib/Languages/CombinatoryLogic/Basic.lean @@ -261,7 +261,7 @@ theorem Θ_correct (f : SKI) : (Θ ⬝ f) ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAu def IsBool (u : Bool) (a : SKI) : Prop := ∀ x y : SKI, (a ⬝ x ⬝ y) ↠ (if u then x else y) -theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠ a') (ha' : IsBool u a') : +theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠a') (ha' : IsBool u a') : IsBool u a := by intro x y trans a' ⬝ x ⬝ y diff --git a/Cslib/Languages/CombinatoryLogic/Confluence.lean b/Cslib/Languages/CombinatoryLogic/Confluence.lean index d26e4cf9..e43c482d 100644 --- a/Cslib/Languages/CombinatoryLogic/Confluence.lean +++ b/Cslib/Languages/CombinatoryLogic/Confluence.lean @@ -68,7 +68,7 @@ theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ↠ a' := b case red_S a b c => exact Relation.ReflTransGen.single (red_S a b c) /-- The inclusion `⇒ ⊆ ⇒ₚ` -/ -theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢ a') : a ⇒ₚ a' := by +theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢a') : a ⇒ₚ a' := by cases h case red_S => apply ParallelReduction.red_S case red_K => apply ParallelReduction.red_K @@ -90,7 +90,7 @@ theorem reflTransGen_parallelReduction_mRed : ext a b constructor · apply Relation.reflTransGen_minimal - · exact λ _ => by rfl + · exact fun _ => by rfl · exact instTransitiveMRed RedSKI · exact @mRed_of_parallelReduction · apply Relation.reflTransGen_minimal @@ -234,7 +234,7 @@ theorem commonReduct_equivalence : Equivalence CommonReduct := by exact join_parallelReduction_equivalence /-- The **Church-Rosser** theorem in the form it is usually stated. -/ -theorem MRed.diamond (a b c : SKI) (hab : a ↠ b) (hac : a ↠ c) : CommonReduct b c := by +theorem MRed.diamond (a b c : SKI) (hab : a ↠b) (hac : a ↠c) : CommonReduct b c := by apply commonReduct_equivalence.trans (y := a) · exact commonReduct_equivalence.symm (commonReduct_of_single hab) · exact commonReduct_of_single hac diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index a7637f2d..f0417d0a 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -94,14 +94,14 @@ theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. -theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠ a') : (a ⬝ b) ↠ (a' ⬝ b) := by +theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠a') : (a ⬝ b) ↠ (a' ⬝ b) := by induction h with | refl => apply MRed.refl | @tail a' a'' _ ha'' ih => apply Relation.ReflTransGen.tail (b := a' ⬝ b) ih exact Red.red_head a' a'' b ha'' -theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠ b') : (a ⬝ b) ↠ (a ⬝ b') := by +theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠b') : (a ⬝ b) ↠ (a ⬝ b') := by induction h with | refl => apply MRed.refl | @tail b' b'' _ hb'' ih => @@ -123,11 +123,11 @@ theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠ b') : (a ⬝ b) ↠ (a ⬝ b -- instance RedMRedTrans : Trans Red Red MRed := -- ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ -lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠ a') (hb : b ↠ b') : +lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠a') (hb : b ↠b') : (a ⬝ b) ↠ (a' ⬝ b') := Trans.simple (MRed.head b ha) (MRed.tail a' hb) -lemma parallel_red {a a' b b' : SKI} (ha : a ⭢ a') (hb : b ⭢ b') : (a ⬝ b) ↠ (a' ⬝ b') := by +lemma parallel_red {a a' b b' : SKI} (ha : a ⭢a') (hb : b ⭢b') : (a ⬝ b) ↠ (a' ⬝ b') := by trans a' ⬝ b all_goals apply MRed.single · exact Red.red_head a a' b ha @@ -137,10 +137,10 @@ lemma parallel_red {a a' b b' : SKI} (ha : a ⭢ a') (hb : b ⭢ b') : (a ⬝ b) /-- Express that two terms have a reduce to a common term. -/ def CommonReduct : SKI → SKI → Prop := Relation.Join RedSKI.MRed -lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨b, h, by rfl⟩ +lemma commonReduct_of_single {a b : SKI} (h : a ↠b) : CommonReduct a b := ⟨b, h, by rfl⟩ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join -theorem reflexive_commonReduct : Reflexive CommonReduct := λ x => by +theorem reflexive_commonReduct : Reflexive CommonReduct := fun x => by refine ⟨x,?_,?_⟩ <;> rfl theorem commonReduct_head {x x' : SKI} (y : SKI) : CommonReduct x x' → CommonReduct (x ⬝ y) (x' ⬝ y) @@ -148,3 +148,5 @@ theorem commonReduct_head {x x' : SKI} (y : SKI) : CommonReduct x x' → CommonR theorem commonReduct_tail (x : SKI) {y y' : SKI} : CommonReduct y y' → CommonReduct (x ⬝ y) (x ⬝ y') | ⟨z, hz, hz'⟩ => ⟨x ⬝ z, MRed.tail x hz, MRed.tail x hz'⟩ + +end SKI diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 907b36b6..fef73975 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -186,7 +186,7 @@ theorem commonReduct_redexFree {x y : SKI} (hy : RedexFree y) (h : CommonReduct (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw /-- If `x` reduces to both `y` and `z`, and `z` is not reducible, then `y` reduces to `z`. -/ -lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : RedexFree z) : y ↠ z := +lemma confluent_redexFree {x y z : SKI} (hxy : x ↠y) (hxz : x ↠z) (hz : RedexFree z) : y ↠ z := let ⟨w, hyw, hzw⟩ := MRed.diamond x y z hxy hxz (redexFree_iff'.1 hz _ |>.1 hzw : z = w) ▸ hyw @@ -194,7 +194,7 @@ lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : Re If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. -/ lemma unique_normal_form {x y z : SKI} - (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := + (hxy : x ↠y) (hxz : x ↠z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) /-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ @@ -263,7 +263,7 @@ lemma churchK_size : (n : Nat) → (churchK n).size = n+1 lemma churchK_injective : Function.Injective churchK := fun n m h => by simpa using congrArg SKI.size h -/-- Injectivity for Church numerals-/ +/-- Injectivity for Church numerals -/ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsChurch m y) (hxy : CommonReduct x y) : n = m := by suffices CommonReduct (churchK n) (churchK m) by diff --git a/Cslib/Languages/CombinatoryLogic/Recursion.lean b/Cslib/Languages/CombinatoryLogic/Recursion.lean index 154cdf98..c6744103 100644 --- a/Cslib/Languages/CombinatoryLogic/Recursion.lean +++ b/Cslib/Languages/CombinatoryLogic/Recursion.lean @@ -21,8 +21,8 @@ correctness proofs `zero_correct` and `succ_correct`. - Predecessor : a term `Pred` so that (`pred_correct`) `IsChurch n a → IsChurch n.pred (Pred ⬝ a)`. - Primitive recursion : a term `Rec` so that (`rec_correct_succ`) `IsChurch (n+1) a` implies -`Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` implies -`Rec ⬝ x ⬝ g ⬝ a ↠ x`. +`Rec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))` and (`rec_correct_zero`) `IsChurch 0 a` +implies `Rec ⬝ x ⬝ g ⬝ a ↠ x`. - Unbounded root finding (μ-recursion) : given a term `f` representing a function `fℕ: Nat → Nat`, which takes on the value 0 a term `RFind` such that (`rFind_correct`) `RFind ⬝ f ↠ a` such that `IsChurch n a` for `n` the smallest root of `fℕ`. @@ -58,7 +58,7 @@ match n with | n+1 => f ⬝ (Church n f x) /-- `church` commutes with reduction. -/ -lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠ f') (hx : x ↠ x') : +lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠f') (hx : x ↠x') : Church n f x ↠ Church n f' x' := by induction n with | zero => exact hx @@ -69,7 +69,7 @@ def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x :SKI, (a ⬝ f ⬝ x) ↠ (Church n f x) /-- To show `IsChurch n a` it suffices to show the same for a reduct of `a`. -/ -theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠ a') : IsChurch n a' → IsChurch n a := by +theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠a') : IsChurch n a' → IsChurch n a := by simp_rw [IsChurch] intro ha' f x calc @@ -118,7 +118,7 @@ theorem predAux_def (p : SKI) : (PredAux ⬝ p) ↠ MkPair ⬝ (Snd ⬝ p) ⬝ def IsChurchPair (ns : Nat × Nat) (x : SKI) : Prop := IsChurch ns.1 (Fst ⬝ x) ∧ IsChurch ns.2 (Snd ⬝ x) -theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠ a') : +theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠a') : IsChurchPair ns a' → IsChurchPair ns a := by simp_rw [IsChurchPair] intro ⟨ha₁,ha₂⟩ @@ -227,7 +227,7 @@ theorem rec_zero (x g a : SKI) (ha : IsChurch 0 a) : (Rec ⬝ x ⬝ g ⬝ a) ↠ apply cond_correct exact isZero_correct 0 a ha -theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n+1) a) : +theorem rec_succ (n : Nat) (x g a : SKI) (ha : IsChurch (n + 1) a) : (Rec ⬝ x ⬝ g ⬝ a) ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a)) := by calc _ ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a) := rec_def _ _ _ @@ -257,7 +257,7 @@ theorem rfindAboveAux_base (R₀ f a : SKI) (hfa : IsChurch 0 (f ⬝ a)) : _ ↠ if (Nat.beq 0 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by apply cond_correct apply isZero_correct _ _ hfa -theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m+1) (f ⬝ a)) : +theorem rfindAboveAux_step (R₀ f a : SKI) {m : Nat} (hfa : IsChurch (m + 1) (f ⬝ a)) : (RFindAboveAux ⬝ R₀ ⬝ a ⬝ f) ↠ R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f := calc _ ↠ SKI.Cond ⬝ a ⬝ (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) ⬝ (IsZero ⬝ (f ⬝ a)) := rfindAboveAux_def _ _ _ _ ↠ if (Nat.beq (m+1) 0) then a else (R₀ ⬝ (SKI.Succ ⬝ a) ⬝ f) := by From faf31f840231486d3a1e76c8d746e8014a418f46 Mon Sep 17 00:00:00 2001 From: twwar Date: Wed, 10 Sep 2025 12:30:07 +0200 Subject: [PATCH 15/19] spaces after transitions --- Cslib/Foundations/Semantics/ReductionSystem/Basic.lean | 4 ++-- Cslib/Languages/CombinatoryLogic/Basic.lean | 2 +- Cslib/Languages/CombinatoryLogic/Confluence.lean | 4 ++-- Cslib/Languages/CombinatoryLogic/Defs.lean | 10 +++++----- Cslib/Languages/CombinatoryLogic/Evaluation.lean | 4 ++-- Cslib/Languages/CombinatoryLogic/Recursion.lean | 6 +++--- 6 files changed, 15 insertions(+), 15 deletions(-) diff --git a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean index dcf7e9ee..a589be12 100644 --- a/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean +++ b/Cslib/Foundations/Semantics/ReductionSystem/Basic.lean @@ -105,8 +105,8 @@ macro_rules ) | `(reduction_notation $rs) => `( - notation3 t:39 " ⭢" t':39 => (ReductionSystem.Red $rs) t t' - notation3 t:39 " ↠" t':39 => (ReductionSystem.MRed $rs) t t' + notation3 t:39 " ⭢ " t':39 => (ReductionSystem.Red $rs) t t' + notation3 t:39 " ↠ " t':39 => (ReductionSystem.MRed $rs) t t' ) diff --git a/Cslib/Languages/CombinatoryLogic/Basic.lean b/Cslib/Languages/CombinatoryLogic/Basic.lean index 837a6af4..2e7212f2 100644 --- a/Cslib/Languages/CombinatoryLogic/Basic.lean +++ b/Cslib/Languages/CombinatoryLogic/Basic.lean @@ -261,7 +261,7 @@ theorem Θ_correct (f : SKI) : (Θ ⬝ f) ↠ f ⬝ (Θ ⬝ f) := ΘAux_def ΘAu def IsBool (u : Bool) (a : SKI) : Prop := ∀ x y : SKI, (a ⬝ x ⬝ y) ↠ (if u then x else y) -theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠a') (ha' : IsBool u a') : +theorem isBool_trans (u : Bool) (a a' : SKI) (h : a ↠ a') (ha' : IsBool u a') : IsBool u a := by intro x y trans a' ⬝ x ⬝ y diff --git a/Cslib/Languages/CombinatoryLogic/Confluence.lean b/Cslib/Languages/CombinatoryLogic/Confluence.lean index e43c482d..00c8df1d 100644 --- a/Cslib/Languages/CombinatoryLogic/Confluence.lean +++ b/Cslib/Languages/CombinatoryLogic/Confluence.lean @@ -68,7 +68,7 @@ theorem mRed_of_parallelReduction {a a' : SKI} (h : a ⇒ₚ a') : a ↠ a' := b case red_S a b c => exact Relation.ReflTransGen.single (red_S a b c) /-- The inclusion `⇒ ⊆ ⇒ₚ` -/ -theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢a') : a ⇒ₚ a' := by +theorem parallelReduction_of_red {a a' : SKI} (h : a ⭢ a') : a ⇒ₚ a' := by cases h case red_S => apply ParallelReduction.red_S case red_K => apply ParallelReduction.red_K @@ -234,7 +234,7 @@ theorem commonReduct_equivalence : Equivalence CommonReduct := by exact join_parallelReduction_equivalence /-- The **Church-Rosser** theorem in the form it is usually stated. -/ -theorem MRed.diamond (a b c : SKI) (hab : a ↠b) (hac : a ↠c) : CommonReduct b c := by +theorem MRed.diamond (a b c : SKI) (hab : a ↠ b) (hac : a ↠ c) : CommonReduct b c := by apply commonReduct_equivalence.trans (y := a) · exact commonReduct_equivalence.symm (commonReduct_of_single hab) · exact commonReduct_of_single hac diff --git a/Cslib/Languages/CombinatoryLogic/Defs.lean b/Cslib/Languages/CombinatoryLogic/Defs.lean index f0417d0a..8c388e93 100644 --- a/Cslib/Languages/CombinatoryLogic/Defs.lean +++ b/Cslib/Languages/CombinatoryLogic/Defs.lean @@ -94,14 +94,14 @@ theorem MRed.S (x y z : SKI) : (S ⬝ x ⬝ y ⬝ z) ↠ (x ⬝ z ⬝ (y ⬝ z)) theorem MRed.K (x y : SKI) : (K ⬝ x ⬝ y) ↠ x := MRed.single RedSKI <| red_K .. theorem MRed.I (x : SKI) : (I ⬝ x) ↠ x := MRed.single RedSKI <| red_I .. -theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠a') : (a ⬝ b) ↠ (a' ⬝ b) := by +theorem MRed.head {a a' : SKI} (b : SKI) (h : a ↠ a') : (a ⬝ b) ↠ (a' ⬝ b) := by induction h with | refl => apply MRed.refl | @tail a' a'' _ ha'' ih => apply Relation.ReflTransGen.tail (b := a' ⬝ b) ih exact Red.red_head a' a'' b ha'' -theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠b') : (a ⬝ b) ↠ (a ⬝ b') := by +theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠ b') : (a ⬝ b) ↠ (a ⬝ b') := by induction h with | refl => apply MRed.refl | @tail b' b'' _ hb'' ih => @@ -123,11 +123,11 @@ theorem MRed.tail (a : SKI) {b b' : SKI} (h : b ↠b') : (a ⬝ b) ↠ (a ⬝ b' -- instance RedMRedTrans : Trans Red Red MRed := -- ⟨fun hab hbc => Relation.ReflTransGen.trans (MRed.single hab) (MRed.single hbc)⟩ -lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠a') (hb : b ↠b') : +lemma parallel_mRed {a a' b b' : SKI} (ha : a ↠ a') (hb : b ↠ b') : (a ⬝ b) ↠ (a' ⬝ b') := Trans.simple (MRed.head b ha) (MRed.tail a' hb) -lemma parallel_red {a a' b b' : SKI} (ha : a ⭢a') (hb : b ⭢b') : (a ⬝ b) ↠ (a' ⬝ b') := by +lemma parallel_red {a a' b b' : SKI} (ha : a ⭢ a') (hb : b ⭢ b') : (a ⬝ b) ↠ (a' ⬝ b') := by trans a' ⬝ b all_goals apply MRed.single · exact Red.red_head a a' b ha @@ -137,7 +137,7 @@ lemma parallel_red {a a' b b' : SKI} (ha : a ⭢a') (hb : b ⭢b') : (a ⬝ b) /-- Express that two terms have a reduce to a common term. -/ def CommonReduct : SKI → SKI → Prop := Relation.Join RedSKI.MRed -lemma commonReduct_of_single {a b : SKI} (h : a ↠b) : CommonReduct a b := ⟨b, h, by rfl⟩ +lemma commonReduct_of_single {a b : SKI} (h : a ↠ b) : CommonReduct a b := ⟨b, h, by rfl⟩ theorem symmetric_commonReduct : Symmetric CommonReduct := Relation.symmetric_join theorem reflexive_commonReduct : Reflexive CommonReduct := fun x => by diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index fef73975..ebfb441a 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -186,7 +186,7 @@ theorem commonReduct_redexFree {x y : SKI} (hy : RedexFree y) (h : CommonReduct (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw /-- If `x` reduces to both `y` and `z`, and `z` is not reducible, then `y` reduces to `z`. -/ -lemma confluent_redexFree {x y z : SKI} (hxy : x ↠y) (hxz : x ↠z) (hz : RedexFree z) : y ↠ z := +lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : RedexFree z) : y ↠ z := let ⟨w, hyw, hzw⟩ := MRed.diamond x y z hxy hxz (redexFree_iff'.1 hz _ |>.1 hzw : z = w) ▸ hyw @@ -194,7 +194,7 @@ lemma confluent_redexFree {x y z : SKI} (hxy : x ↠y) (hxz : x ↠z) (hz : Rede If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. -/ lemma unique_normal_form {x y z : SKI} - (hxy : x ↠y) (hxz : x ↠z) (hy : RedexFree y) (hz : RedexFree z) : y = z := + (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) /-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ diff --git a/Cslib/Languages/CombinatoryLogic/Recursion.lean b/Cslib/Languages/CombinatoryLogic/Recursion.lean index c6744103..2d16dc23 100644 --- a/Cslib/Languages/CombinatoryLogic/Recursion.lean +++ b/Cslib/Languages/CombinatoryLogic/Recursion.lean @@ -58,7 +58,7 @@ match n with | n+1 => f ⬝ (Church n f x) /-- `church` commutes with reduction. -/ -lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠f') (hx : x ↠x') : +lemma church_red (n : Nat) (f f' x x' : SKI) (hf : f ↠ f') (hx : x ↠ x') : Church n f x ↠ Church n f' x' := by induction n with | zero => exact hx @@ -69,7 +69,7 @@ def IsChurch (n : Nat) (a : SKI) : Prop := ∀ f x :SKI, (a ⬝ f ⬝ x) ↠ (Church n f x) /-- To show `IsChurch n a` it suffices to show the same for a reduct of `a`. -/ -theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠a') : IsChurch n a' → IsChurch n a := by +theorem isChurch_trans (n : Nat) {a a' : SKI} (h : a ↠ a') : IsChurch n a' → IsChurch n a := by simp_rw [IsChurch] intro ha' f x calc @@ -118,7 +118,7 @@ theorem predAux_def (p : SKI) : (PredAux ⬝ p) ↠ MkPair ⬝ (Snd ⬝ p) ⬝ def IsChurchPair (ns : Nat × Nat) (x : SKI) : Prop := IsChurch ns.1 (Fst ⬝ x) ∧ IsChurch ns.2 (Snd ⬝ x) -theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠a') : +theorem isChurchPair_trans (ns : Nat × Nat) (a a' : SKI) (h : a ↠ a') : IsChurchPair ns a' → IsChurchPair ns a := by simp_rw [IsChurchPair] intro ⟨ha₁,ha₂⟩ From d97644a208b1f53bad5c04ae59c887503d3861e6 Mon Sep 17 00:00:00 2001 From: twwar Date: Wed, 24 Sep 2025 15:03:30 +0200 Subject: [PATCH 16/19] documentation --- Cslib/Languages/CombinatoryLogic/Evaluation.lean | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index ebfb441a..78861eb7 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -277,7 +277,15 @@ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsC · simp_rw [churchK_church] exact commonReduct_of_single (hy K K) -/-- **Rice's theorem**: no SKI term is a non-trivial predicate. -/ +/-- +**Rice's theorem**: no SKI term is a non-trivial predicate. + +More specifically, say a term `P` is a *predicate* if, for every term `x`, `P · x` reduces to either +`TT` or `FF`. A predicate `P` is *trivial* if either it always reduces to true, or always to false. +This version of Rice's theorem derives a contradiction from the existence of a predicate `P` and the +existence of terms `x` for which `P · x` is true (`P · x ↠ TT`) and for which `P · x` is false +(`P · x ↠ FF`). +-/ theorem rice {P : SKI} (hP : ∀ x : SKI, ((P ⬝ x) ↠ TT) ∨ (P ⬝ x) ↠ FF) (hxt : ∃ x : SKI, (P ⬝ x) ↠ TT) (hxf : ∃ x : SKI, (P ⬝ x) ↠ FF) : False := by obtain ⟨a, ha⟩ := hxt @@ -304,7 +312,11 @@ theorem rice {P : SKI} (hP : ∀ x : SKI, ((P ⬝ x) ↠ TT) ∨ (P ⬝ x) ↠ F _ ↠ TT := ha exact TF_nequiv <| MRed.diamond _ _ _ this h -/-- **Rice's theorem**: any SKI predicate is trivial. -/ +/-- **Rice's theorem**: any SKI predicate is trivial. + +This version of Rice's theorem proves (classically) that any SKI predicate `P` either is constantly +true (ie `P · x ↠ TT` for every `x`) or is constantly false (`P · x ↠ FF` for every `x`). +-/ theorem rice' {P : SKI} (hP : ∀ x : SKI, ((P ⬝ x) ↠ TT) ∨ (P ⬝ x) ↠ FF) : (∀ x : SKI, (P ⬝ x) ↠ TT) ∨ (∀ x : SKI, (P ⬝ x) ↠ FF) := by by_contra! h From 37c800744088518dfd92ef9ef6c2acf5a2b8784c Mon Sep 17 00:00:00 2001 From: twwar Date: Wed, 24 Sep 2025 15:14:57 +0200 Subject: [PATCH 17/19] dot notation --- .../CombinatoryLogic/Evaluation.lean | 46 +++++++++---------- 1 file changed, 23 insertions(+), 23 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 78861eb7..d41b5840 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -36,14 +36,14 @@ This file draws heavily from True | K => True | I => True - | S ⬝ x => RedexFree x - | K ⬝ x => RedexFree x + | S ⬝ x => x.RedexFree + | K ⬝ x => x.RedexFree | I ⬝ _ => False - | S ⬝ x ⬝ y => RedexFree x ∧ RedexFree y + | S ⬝ x ⬝ y => x.RedexFree ∧ y.RedexFree | K ⬝ _ ⬝ _ => False | I ⬝ _ ⬝ _ => False | S ⬝ _ ⬝ _ ⬝ _ => False @@ -55,18 +55,18 @@ def RedexFree : SKI → Prop One-step evaluation as a function: either it returns a term that has been reduced by one step, or a proof that the term is redex free. Uses normal-order reduction. -/ -def evalStep : (x : SKI) → PLift (RedexFree x) ⊕ SKI +def SKI.evalStep : (x : SKI) → PLift (x.RedexFree) ⊕ SKI | S => Sum.inl (PLift.up trivial) | K => Sum.inl (PLift.up trivial) | I => Sum.inl (PLift.up trivial) - | S ⬝ x => match evalStep x with + | S ⬝ x => match x.evalStep with | Sum.inl h => Sum.inl h | Sum.inr x' => Sum.inr (S ⬝ x') - | K ⬝ x => match evalStep x with + | K ⬝ x => match x.evalStep with | Sum.inl h => Sum.inl h | Sum.inr x' => Sum.inr (K ⬝ x') | I ⬝ x => Sum.inr x - | S ⬝ x ⬝ y => match evalStep x, evalStep y with + | S ⬝ x ⬝ y => match x.evalStep, y.evalStep with | Sum.inl h1, Sum.inl h2 => Sum.inl (.up ⟨h1.down, h2.down⟩) | Sum.inl _, Sum.inr y' => Sum.inr (S ⬝ x ⬝ y') | Sum.inr x', _ => Sum.inr (S ⬝ x' ⬝ y) @@ -81,16 +81,16 @@ def evalStep : (x : SKI) → PLift (RedexFree x) ⊕ SKI | Sum.inr abcd', _ => Sum.inr (abcd' ⬝ e) /-- The normal-order reduction implemented by `evalStep` indeed computes a one-step reduction. -/ -theorem evalStep_right_correct : (x y : SKI) → (evalStep x = Sum.inr y) → x ⭢ y +theorem evalStep_right_correct : (x y : SKI) → (x.evalStep = Sum.inr y) → x ⭢ y | S ⬝ x, a, h => - match hx : evalStep x with + match hx : x.evalStep with | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h | Sum.inr x' => by simp only [evalStep, hx, Sum.inr.injEq] at h rw [←h] exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) | K ⬝ x, a, h => - match hx : evalStep x with + match hx : x.evalStep with | Sum.inl _ => by simp only [hx, evalStep, reduceCtorEq] at h | Sum.inr x' => by simp only [evalStep, hx, Sum.inr.injEq] at h @@ -98,7 +98,7 @@ theorem evalStep_right_correct : (x y : SKI) → (evalStep x = Sum.inr y) → x exact .red_tail _ _ _ (evalStep_right_correct _ _ hx) | I ⬝ x, a, h => Sum.inr.inj h ▸ red_I _ | S ⬝ x ⬝ y, a, h => - match hx : evalStep x, hy : evalStep y with + match hx : x.evalStep, hy : y.evalStep with | Sum.inl _, Sum.inl _ => by simp only [hx, hy, evalStep, reduceCtorEq] at h | Sum.inl _, Sum.inr y' => by simp only [hx, hy, evalStep, Sum.inr.injEq] at h @@ -125,12 +125,12 @@ theorem evalStep_right_correct : (x y : SKI) → (evalStep x = Sum.inr y) → x rw [←h] exact red_head _ _ _ <| evalStep_right_correct _ _ habcd -theorem redexFree_of_no_red {x : SKI} (h : ∀ y, ¬ (x ⭢ y)) : RedexFree x := by - match hx : evalStep x with +theorem redexFree_of_no_red {x : SKI} (h : ∀ y, ¬ (x ⭢ y)) : x.RedexFree := by + match hx : x.evalStep with | Sum.inl h' => exact h'.down | Sum.inr y => cases h _ (evalStep_right_correct x y hx) -theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) +theorem SKI.RedexFree.no_red : {x : SKI} → x.RedexFree → ∀ y, ¬ (x ⭢ y) | S ⬝ x, hx, S ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' | K ⬝ x, hx, K ⬝ y, red_tail _ _ _ hx' => by rw [RedexFree] at hx; exact hx.no_red y hx' | S ⬝ _ ⬝ _, ⟨hx, _⟩, S ⬝ _ ⬝ _, red_head _ _ _ (red_tail _ _ _ h3) => hx.no_red _ h3 @@ -139,26 +139,26 @@ theorem RedexFree.no_red : {x : SKI} → RedexFree x → ∀ y, ¬ (x ⭢ y) | _ ⬝ _ ⬝ _ ⬝ _ ⬝ _, ⟨_, hy⟩, _ ⬝ _, red_tail _ _ _ he => hy.no_red _ he /-- A term is redex free iff it has no one-step reductions. -/ -theorem redexFree_iff {x : SKI} : RedexFree x ↔ ∀ y, ¬ (x ⭢ y) := +theorem redexFree_iff {x : SKI} : x.RedexFree ↔ ∀ y, ¬ (x ⭢ y) := ⟨RedexFree.no_red, redexFree_of_no_red⟩ -theorem redexFree_iff_evalStep {x : SKI} : RedexFree x ↔ (evalStep x).isLeft = true := by +theorem redexFree_iff_evalStep {x : SKI} : x.RedexFree ↔ (x.evalStep).isLeft = true := by constructor case mp => intro h - match hx : evalStep x with + match hx : x.evalStep with | Sum.inl h' => exact rfl | Sum.inr y => cases h.no_red _ (evalStep_right_correct _ _ hx) case mpr => intro h - match hx : evalStep x with + match hx : x.evalStep with | Sum.inl h' => exact h'.down | Sum.inr y => rw [hx] at h; cases h instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_evalStep /-- A term is redex free iff its only many-step reduction is itself. -/ -theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := by +theorem redexFree_iff' {x : SKI} : x.RedexFree ↔ ∀ y, (x ↠ y) ↔ x = y := by constructor case mp => intro h y @@ -181,7 +181,7 @@ theorem redexFree_iff' {x : SKI} : RedexFree x ↔ ∀ y, (x ↠ y) ↔ x = y := exact Red.ne hy (h.1 (Relation.ReflTransGen.single hy)) /-- If a term has a common reduct with a normal term, it in fact reduces to that term. -/ -theorem commonReduct_redexFree {x y : SKI} (hy : RedexFree y) (h : CommonReduct x y) : x ↠ y := +theorem commonReduct_redexFree {x y : SKI} (hy : y.RedexFree) (h : CommonReduct x y) : x ↠ y := let ⟨w, hyw, hzw⟩ := h (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw @@ -194,12 +194,12 @@ lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : Re If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. -/ lemma unique_normal_form {x y z : SKI} - (hxy : x ↠ y) (hxz : x ↠ z) (hy : RedexFree y) (hz : RedexFree z) : y = z := + (hxy : x ↠ y) (hxz : x ↠ z) (hy : y.RedexFree) (hz : RedexFree z) : y = z := (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) /-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) - (hx : RedexFree x) (hy : RedexFree y) : x = y := + (hx : x.RedexFree) (hy : y.RedexFree) : x = y := (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) /-! ### Injectivity for datatypes -/ From 6b0b50c5c05a025554d84a6dac5c467909aab240 Mon Sep 17 00:00:00 2001 From: twwar Date: Sat, 27 Sep 2025 17:39:23 +0200 Subject: [PATCH 18/19] more descriptive names --- .../CombinatoryLogic/Evaluation.lean | 20 +++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index d41b5840..953ae069 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -22,7 +22,7 @@ This file formalises evaluation and normal forms of SKI terms. ## Main results - `evalStep_right_correct` : correctness for `evalStep` -- `redexFree_iff` and `redexFree_iff'` : a term is redex free if and only if it has (respectively) +- `redexFree_iff` and `redexFree_iff_mred_eq` : a term is redex free if and only if it has (respectively) no one-step reductions, or if its only many step reduction is itself. - `unique_normal_form` : if `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. @@ -158,7 +158,7 @@ theorem redexFree_iff_evalStep {x : SKI} : x.RedexFree ↔ (x.evalStep).isLeft = instance : DecidablePred RedexFree := fun _ => decidable_of_iff' _ redexFree_iff_evalStep /-- A term is redex free iff its only many-step reduction is itself. -/ -theorem redexFree_iff' {x : SKI} : x.RedexFree ↔ ∀ y, (x ↠ y) ↔ x = y := by +theorem redexFree_iff_mred_eq {x : SKI} : x.RedexFree ↔ ∀ y, (x ↠ y) ↔ x = y := by constructor case mp => intro h y @@ -183,24 +183,24 @@ theorem redexFree_iff' {x : SKI} : x.RedexFree ↔ ∀ y, (x ↠ y) ↔ x = y := /-- If a term has a common reduct with a normal term, it in fact reduces to that term. -/ theorem commonReduct_redexFree {x y : SKI} (hy : y.RedexFree) (h : CommonReduct x y) : x ↠ y := let ⟨w, hyw, hzw⟩ := h - (redexFree_iff'.1 hy _ |>.1 hzw : y = w) ▸ hyw + (redexFree_iff_mred_eq.1 hy _ |>.1 hzw : y = w) ▸ hyw /-- If `x` reduces to both `y` and `z`, and `z` is not reducible, then `y` reduces to `z`. -/ lemma confluent_redexFree {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hz : RedexFree z) : y ↠ z := let ⟨w, hyw, hzw⟩ := MRed.diamond x y z hxy hxz - (redexFree_iff'.1 hz _ |>.1 hzw : z = w) ▸ hyw + (redexFree_iff_mred_eq.1 hz _ |>.1 hzw : z = w) ▸ hyw /-- If `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. -/ lemma unique_normal_form {x y z : SKI} (hxy : x ↠ y) (hxz : x ↠ z) (hy : y.RedexFree) (hz : RedexFree z) : y = z := - (redexFree_iff'.1 hy _).1 (confluent_redexFree hxy hxz hz) + (redexFree_iff_mred_eq.1 hy _).1 (confluent_redexFree hxy hxz hz) /-- If `x` and `y` are normal and have a common reduct, then they are equal. -/ -lemma unique_normal_form' {x y : SKI} (h : CommonReduct x y) +lemma eq_of_commonReduct_redexFree {x y : SKI} (h : CommonReduct x y) (hx : x.RedexFree) (hy : y.RedexFree) : x = y := - (redexFree_iff'.1 hx _).1 (commonReduct_redexFree hy h) + (redexFree_iff_mred_eq.1 hx _).1 (commonReduct_redexFree hy h) /-! ### Injectivity for datatypes -/ @@ -208,8 +208,8 @@ lemma sk_nequiv : ¬ CommonReduct S K := by intro ⟨z, hsz, hkz⟩ have hS : RedexFree S := by simp [RedexFree] have hK : RedexFree K := by simp [RedexFree] - cases (redexFree_iff'.1 hS z).1 hsz - cases (redexFree_iff'.1 hK _).1 hkz + cases (redexFree_iff_mred_eq.1 hS z).1 hsz + cases (redexFree_iff_mred_eq.1 hK _).1 hkz /-- Injectivity for booleans. -/ theorem isBool_injective (x y : SKI) (u v : Bool) (hx : IsBool u x) (hy : IsBool v y) @@ -268,7 +268,7 @@ theorem isChurch_injective (x y : SKI) (n m : Nat) (hx : IsChurch n x) (hy : IsC (hxy : CommonReduct x y) : n = m := by suffices CommonReduct (churchK n) (churchK m) by apply churchK_injective - exact unique_normal_form' this (churchK_redexFree n) (churchK_redexFree m) + exact eq_of_commonReduct_redexFree this (churchK_redexFree n) (churchK_redexFree m) apply commonReduct_equivalence.trans (y := x ⬝ K ⬝ K) · simp_rw [churchK_church] exact commonReduct_equivalence.symm <| commonReduct_of_single (hx K K) From 74b7edb5f6cfa9ee24bf485baedc1b10701b75bf Mon Sep 17 00:00:00 2001 From: twwar Date: Sat, 27 Sep 2025 17:43:59 +0200 Subject: [PATCH 19/19] lint --- Cslib/Languages/CombinatoryLogic/Evaluation.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Cslib/Languages/CombinatoryLogic/Evaluation.lean b/Cslib/Languages/CombinatoryLogic/Evaluation.lean index 953ae069..02339a5f 100644 --- a/Cslib/Languages/CombinatoryLogic/Evaluation.lean +++ b/Cslib/Languages/CombinatoryLogic/Evaluation.lean @@ -22,8 +22,8 @@ This file formalises evaluation and normal forms of SKI terms. ## Main results - `evalStep_right_correct` : correctness for `evalStep` -- `redexFree_iff` and `redexFree_iff_mred_eq` : a term is redex free if and only if it has (respectively) -no one-step reductions, or if its only many step reduction is itself. +- `redexFree_iff` and `redexFree_iff_mred_eq` : a term is redex free if and only if it has +(respectively) no one-step reductions, or if its only many step reduction is itself. - `unique_normal_form` : if `x` reduces to both `y` and `z`, and both `y` and `z` are in normal form, then they are equal. - **Rice's theorem**: no SKI term is a non-trivial predicate.