From c034a7ea1335705ca1e1a7461fac257df6757d57 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 9 Jun 2023 16:07:39 +0200 Subject: Start working on extrinsic proofs of termination --- backends/lean/Base/Diverge.lean | 208 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 208 insertions(+) create mode 100644 backends/lean/Base/Diverge.lean (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean new file mode 100644 index 00000000..bd500c25 --- /dev/null +++ b/backends/lean/Base/Diverge.lean @@ -0,0 +1,208 @@ +import Lean +import Base.Primitives + +namespace Diverge + +open Primitives + +section Fix + +open Result + +variable {a b : Type} + +/-! # The least fixed point definition and its properties -/ + +def least_p (p : Nat → Prop) (n : Nat) : Prop := p n ∧ (∀ m, m < n → ¬ p m) +noncomputable def least (p : Nat → Prop) : Nat := + Classical.epsilon (least_p p) + +-- Auxiliary theorem for [least_spec]: if there exists an `n` satisfying `p`, +-- there there exists a least `m` satisfying `p`. +theorem least_spec_aux (p : Nat → Prop) : ∀ (n : Nat), (hn : p n) → ∃ m, least_p p m := by + apply Nat.strongRec' + intros n hi hn + -- Case disjunction on: is n the smallest n satisfying p? + match Classical.em (∀ m, m < n → ¬ p m) with + | .inl hlt => + -- Yes: trivial + exists n + | .inr hlt => + simp at * + let ⟨ m, ⟨ hmlt, hm ⟩ ⟩ := hlt + have hi := hi m hmlt hm + apply hi + +-- The specification of [least]: either `p` is never satisfied, or it is satisfied +-- by `least p` and no `n < least p` satisfies `p`. +theorem least_spec (p : Nat → Prop) : (∀ n, ¬ p n) ∨ (p (least p) ∧ ∀ n, n < least p → ¬ p n) := by + -- Case disjunction on the existence of an `n` which satisfies `p` + match Classical.em (∀ n, ¬ p n) with + | .inl h => + -- There doesn't exist: trivial + apply (Or.inl h) + | .inr h => + -- There exists: we simply use `least_spec_aux` in combination with the property + -- of the epsilon operator + simp at * + let ⟨ n, hn ⟩ := h + apply Or.inr + have hl := least_spec_aux p n hn + have he := Classical.epsilon_spec hl + apply he + +/-! # The fixed point definitions -/ + +def fix_fuel (n : Nat) (f : (a → Result b) → a → Result b) (x : a) : Result b := + match n with + | 0 => .div + | n + 1 => + f (fix_fuel n f) x + +@[simp] def fix_fuel_pred (f : (a → Result b) → a → Result b) (x : a) (n : Nat) := + not (div? (fix_fuel n f x)) + +def fix_fuel_P (f : (a → Result b) → a → Result b) (x : a) (n : Nat) : Prop := + fix_fuel_pred f x n + +noncomputable def fix (f : (a → Result b) → a → Result b) (x : a) : Result b := + fix_fuel (least (fix_fuel_P f x)) f x + +/-! # The proof of the fixed point equation -/ + +-- Monotonicity relation over results +-- TODO: generalize +def result_rel {a : Type u} (x1 x2 : Result a) : Prop := + match x1 with + | div => True + | fail _ => x2 = x1 + | ret _ => x2 = x1 -- TODO: generalize + +-- Monotonicity relation over monadic arrows +-- TODO: generalize +def marrow_rel (f g : a → Result b) : Prop := + ∀ x, result_rel (f x) (g x) + +-- Validity property for a body +def is_valid (f : (a → Result b) → a → Result b) : Prop := + ∀ {{g h}}, marrow_rel g h → marrow_rel (f g) (f h) + +/- + + -/ + +theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : + ∀ {{n m}}, n ≤ m → marrow_rel (fix_fuel n f) (fix_fuel m f) := by + intros n + induction n + case zero => simp [marrow_rel, fix_fuel, result_rel] + case succ n1 Hi => + intros m Hle x + simp [result_rel] + match m with + | 0 => + exfalso + -- TODO: annoying to do those conversions by hand - try zify? + have : n1 + 1 ≤ (0 : Int) := by simp [*] at * + have : 0 ≤ n1 := by simp [*] at * + linarith + | Nat.succ m1 => + simp_arith at Hle + simp [fix_fuel] + have Hi := Hi Hle + simp [is_valid] at Hvalid + have Hvalid := Hvalid Hi x + simp [result_rel] at Hvalid + apply Hvalid + +@[simp] theorem neg_fix_fuel_P {f : (a → Result b) → a → Result b} {x : a} {n : Nat} : + ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by + simp [fix_fuel_P, div?] + cases fix_fuel n f x <;> simp + +theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : + ∀ n, marrow_rel (fix_fuel n f) (fix f) := by + intros n x + simp [result_rel] + have Hl := least_spec (fix_fuel_P f x) + simp at Hl + match Hl with + | .inl Hl => simp [*] + | .inr ⟨ Hl, Hn ⟩ => + match Classical.em (fix_fuel n f x = div) with + | .inl Hd => + simp [*] + | .inr Hd => + have Hineq : least (fix_fuel_P f x) ≤ n := by + -- Proof by contradiction + cases Classical.em (least (fix_fuel_P f x) ≤ n) <;> simp [*] + simp at * + rename_i Hineq + have Hn := Hn n Hineq + contradiction + have Hfix : ¬ (fix f x = div) := by + simp [fix] + -- By property of the least upper bound + revert Hd Hl + -- TODO: there is no conversion to select the head of a function! + have : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := + by simp[fix_fuel_P] + simp [this, div?] + clear this + cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp + have Hmono := fix_fuel_mono Hvalid Hineq x + simp [result_rel] at Hmono + -- TODO: there is no conversion to select the head of a function! + revert Hmono Hfix Hd + simp [fix] + -- TODO: it would be good if cases actually introduces an equation: this + -- way we wouldn't have to do all the book-keeping + cases fix_fuel (least (fix_fuel_P f x)) f x <;> cases fix_fuel n f x <;> + intros <;> simp [*] at * + +theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : + ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by sorry + +theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_valid f) : + ∀ x, fix f x = f (fix f) x := by + intros x + -- conv => lhs; simp [fix] + -- Case disjunction: is there a fuel such that the execution successfully execute? + match Classical.em (∃ n, fix_fuel_P f x n) with + | .inr He => + -- No fuel: the fixed point evaluates to `div` + --simp [fix] at * + simp at * + simp [fix] + have He := He (Nat.succ (least (fix_fuel_P f x))) + simp [*, fix_fuel] at * + -- Use the monotonicity of `f` + have Hmono := fix_fuel_fix_mono Hvalid (least (fix_fuel_P f x)) x + simp [result_rel] at Hmono + simp [*] at * + -- TODO: we need a stronger validity predicate + sorry + | .inl ⟨ n, He ⟩ => + have Hl := fix_fuel_P_least Hvalid He + -- TODO: better control of simplification + have Heq : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := + by simp [fix_fuel_P] + simp [Heq] at Hl; clear Heq + -- The least upper bound is > 0 + have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by sorry + simp [Hsucc] at Hl + revert Hl + simp [*, div?, fix, fix_fuel] + -- Use the monotonicity + have Hineq : n ≤ Nat.succ n := by sorry + have Hmono := fix_fuel_fix_mono Hvalid n + have Hv := Hvalid Hmono x + -- Use functional extensionality + simp [result_rel, fix] at Hv + revert Hv + split <;> simp [*] <;> intros <;> simp [*] + + +end Fix + +end Diverge -- cgit v1.2.3 From ef6204e1e1b0a21975fcd9e3d0e5aa7ec3d9125f Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 13 Jun 2023 16:22:08 +0200 Subject: Find sufficient validity criteria for Diverge.lean --- backends/lean/Base/Diverge.lean | 350 +++++++++++++++++++++++++++++++++++----- 1 file changed, 309 insertions(+), 41 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index bd500c25..b5264d0d 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -5,7 +5,7 @@ namespace Diverge open Primitives -section Fix +namespace Fix open Result @@ -79,19 +79,32 @@ def result_rel {a : Type u} (x1 x2 : Result a) : Prop := | ret _ => x2 = x1 -- TODO: generalize -- Monotonicity relation over monadic arrows +-- TODO: Kleisli arrow -- TODO: generalize def marrow_rel (f g : a → Result b) : Prop := ∀ x, result_rel (f x) (g x) --- Validity property for a body -def is_valid (f : (a → Result b) → a → Result b) : Prop := +-- Monotonicity property +def is_mono (f : (a → Result b) → a → Result b) : Prop := ∀ {{g h}}, marrow_rel g h → marrow_rel (f g) (f h) +-- "Continuity" property. +-- We need this, and this looks a lot like continuity. Also see this paper: +-- https://inria.hal.science/file/index/docid/216187/filename/tarski.pdf +def is_cont (f : (a → Result b) → a → Result b) : Prop := + ∀ x, (Hdiv : ∀ n, fix_fuel (.succ n) f x = div) → f (fix f) x = div + +-- Validity property for a body +structure is_valid (f : (a → Result b) → a → Result b) := + intro:: + hmono : is_mono f + hcont : is_cont f + /- -/ -theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : +theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : ∀ {{n m}}, n ≤ m → marrow_rel (fix_fuel n f) (fix_fuel m f) := by intros n induction n @@ -110,17 +123,16 @@ theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hvalid : is_val simp_arith at Hle simp [fix_fuel] have Hi := Hi Hle - simp [is_valid] at Hvalid - have Hvalid := Hvalid Hi x - simp [result_rel] at Hvalid - apply Hvalid + have Hmono := Hmono Hi x + simp [result_rel] at Hmono + apply Hmono @[simp] theorem neg_fix_fuel_P {f : (a → Result b) → a → Result b} {x : a} {n : Nat} : ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by simp [fix_fuel_P, div?] cases fix_fuel n f x <;> simp -theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : +theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : ∀ n, marrow_rel (fix_fuel n f) (fix f) := by intros n x simp [result_rel] @@ -150,7 +162,7 @@ theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hvalid : is simp [this, div?] clear this cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp - have Hmono := fix_fuel_mono Hvalid Hineq x + have Hmono := fix_fuel_mono Hmono Hineq x simp [result_rel] at Hmono -- TODO: there is no conversion to select the head of a function! revert Hmono Hfix Hd @@ -160,9 +172,42 @@ theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hvalid : is cases fix_fuel (least (fix_fuel_P f x)) f x <;> cases fix_fuel n f x <;> intros <;> simp [*] at * -theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hvalid : is_valid f) : - ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by sorry +theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by + intros x n Hf + have Hfmono := fix_fuel_fix_mono Hmono n x + revert Hf Hfmono + -- TODO: would be good to be able to unfold fix_fuel_P only on the left + simp [fix_fuel_P, div?, result_rel, fix] + cases fix_fuel n f x <;> simp_all +-- Prove the fixed point equation in the case there exists some fuel for which +-- the execution terminates +theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono : is_mono f) + (x : a) (n : Nat) (He : fix_fuel_P f x n) : + fix f x = f (fix f) x := by + have Hl := fix_fuel_P_least Hmono He + -- TODO: better control of simplification + have Heq : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := + by simp [fix_fuel_P] + simp [Heq] at Hl; clear Heq + -- The least upper bound is > 0 + have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by + revert Hl + simp [div?] + cases least (fix_fuel_P f x) <;> simp [fix_fuel] + simp [Hsucc] at Hl + revert Hl + simp [*, div?, fix, fix_fuel] + -- Use the monotonicity + have Hfixmono := fix_fuel_fix_mono Hmono n + have Hvm := Hmono Hfixmono x + -- Use functional extensionality + simp [result_rel, fix] at Hvm + revert Hvm + split <;> simp [*] <;> intros <;> simp [*] + +-- The final fixed point equation theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_valid f) : ∀ x, fix f x = f (fix f) x := by intros x @@ -173,36 +218,259 @@ theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_vali -- No fuel: the fixed point evaluates to `div` --simp [fix] at * simp at * - simp [fix] - have He := He (Nat.succ (least (fix_fuel_P f x))) - simp [*, fix_fuel] at * - -- Use the monotonicity of `f` - have Hmono := fix_fuel_fix_mono Hvalid (least (fix_fuel_P f x)) x - simp [result_rel] at Hmono - simp [*] at * - -- TODO: we need a stronger validity predicate - sorry - | .inl ⟨ n, He ⟩ => - have Hl := fix_fuel_P_least Hvalid He - -- TODO: better control of simplification - have Heq : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := - by simp [fix_fuel_P] - simp [Heq] at Hl; clear Heq - -- The least upper bound is > 0 - have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by sorry - simp [Hsucc] at Hl - revert Hl - simp [*, div?, fix, fix_fuel] - -- Use the monotonicity - have Hineq : n ≤ Nat.succ n := by sorry - have Hmono := fix_fuel_fix_mono Hvalid n - have Hv := Hvalid Hmono x - -- Use functional extensionality - simp [result_rel, fix] at Hv - revert Hv - split <;> simp [*] <;> intros <;> simp [*] - + conv => lhs; simp [fix] + have Hel := He (Nat.succ (least (fix_fuel_P f x))); simp [*, fix_fuel] at *; clear Hel + -- Use the "continuity" of `f` + have He : ∀ n, fix_fuel (.succ n) f x = div := by intros; simp [*] + have Hcont := Hvalid.hcont x He + simp [Hcont] + | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hvalid.hmono x n He +/- +(∀ n, fix_fuel n f x = div) + +⊢ f (fun y => fix_fuel (least (fix_fuel_P f y)) f y) x = div + +(? x. p x) ==> p (epsilon p) + + +P (nf : a -> option Nat) := + match nf x with + | None => forall n, fix_fuel n f x = div + | Some n => fix_fuel n f x <> div + +TODO: theorem de Tarsky, +Gilles Dowek (Introduction à la théorie des langages de programmation) + +fix_f is f s.t.: f x = f (fix f) x ∧ ! g. g x = g (fix g) x ==> f <= g + +-/ + end Fix +namespace Ex1 + /- An example of use of the fixed-point -/ + open Fix + + variable {a : Type} (f : (List a × Int) → Result a) + + def list_nth_body (x : (List a × Int)) : Result a := + let (ls, i) := x + match ls with + | [] => .fail .panic + | hd :: tl => + if i = 0 then .ret hd + else f (tl, i - 1) + + theorem list_nth_body_mono : is_mono (@list_nth_body a) := by + simp [is_mono]; intro g h Hr (ls, i); simp [result_rel, list_nth_body] + cases ls <;> simp + rename_i hd tl + -- TODO: making a case disjunction over `i = 0` is annoying, we need a more + -- general tactic for this + cases (Classical.em (Eq i 0)) <;> simp [*] at * + apply Hr + + theorem list_nth_body_cont : is_cont (@list_nth_body a) := by + rw [is_cont]; intro (ls, i) Hdiv; simp [list_nth_body, fix_fuel] at * + cases ls <;> simp at * + -- TODO: automate this + cases (Classical.em (Eq i 0)) <;> simp [*] at * + -- Recursive call + apply Hdiv + + noncomputable + def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) + + theorem list_nth_eq (ls : List a) (i : Int) : + list_nth ls i = + match ls with + | [] => .fail .panic + | hd :: tl => + if i = 0 then .ret hd + else list_nth tl (i - 1) + := by + have Hvalid : is_valid (@list_nth_body a) := + is_valid.intro list_nth_body_mono list_nth_body_cont + have Heq := fix_fixed_eq (@list_nth_body a) Hvalid + simp [Heq, list_nth] + conv => lhs; rw [list_nth_body] + simp [Heq] + +end Ex1 + +namespace Ex2 + /- Higher-order example -/ + open Fix + + variable {a b : Type} + + /- An auxiliary function, which doesn't require the fixed-point -/ + def map (f : a → Result b) (ls : List a) : Result (List b) := + match ls with + | [] => .ret [] + | hd :: tl => + do + match f hd with + | .ret hd => + match map f tl with + | .ret tl => + .ret (hd :: tl) + | r => r + | .fail e => .fail e + | .div => .div + + theorem map_is_mono {{f g : a → Result b}} (Hr : marrow_rel f g) : + ∀ ls, result_rel (map f ls) (map g ls) := by + intro ls; induction ls <;> simp [result_rel, map] + case cons hd tl Hi => + have Hr1 := Hr hd; simp [result_rel] at Hr1 + -- TODO: reverting is annoying + revert Hr1 + cases f hd <;> intro Hr1 <;> simp [*] + -- ret case + simp [result_rel] at Hi + -- TODO: annoying + revert Hi + cases map f tl <;> intro Hi <;> simp [*] + + -- Auxiliary definition + def map_fix_fuel (n0 n1 : Nat) (f : (a → Result b) → a → Result b) (ls : List a) : Result (List b) := + match ls with + | [] => .ret [] + | hd :: tl => + do + match fix_fuel n0 f hd with + | .ret hd => + match map (fix_fuel n1 f) tl with + | .ret tl => + .ret (hd :: tl) + | r => r + | .fail e => .fail e + | .div => .div + + def exists_map_fix_fuel_not_div_imp {{f : (a → Result b) → a → Result b}} {{ls : List a}} + (Hmono : is_mono f) : + (∃ n0 n1, map_fix_fuel n0 n1 f ls ≠ .div) → + ∃ n2, map (fix_fuel n2 f) ls ≠ .div := by + intro ⟨ n0, n1, Hnd ⟩ + exists n0 + n1 + have Hineq0 : n0 ≤ n0 + n1 := by linarith + have Hineq1 : n1 ≤ n0 + n1 := by linarith + simp [map_fix_fuel] at Hnd + -- TODO: I would like a rewrite_once tactic + unfold map; simp + -- + revert Hnd + cases ls <;> simp + rename_i hd tl + -- Use the monotonicity of fix_fuel + have Hfmono := fix_fuel_mono Hmono Hineq0 hd + simp [result_rel] at Hfmono; revert Hfmono + cases fix_fuel n0 f hd <;> intro <;> simp [*] + -- Use the monotonicity of map + have Hfmono := fix_fuel_mono Hmono Hineq1 + have Hmmono := map_is_mono Hfmono tl + simp [result_rel] at Hmmono; revert Hmmono + cases map (fix_fuel n1 f) tl <;> intro <;> simp [*] + + -- TODO: it is simpler to prove the contrapositive of is_cont than is_cont itself. + -- The proof is still quite technical: think of a criteria whose proof is simpler + -- to automate. + theorem map_is_cont_contra_aux {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : + ∀ ls, map (fix f) ls ≠ .div → + ∃ n0 n1, map_fix_fuel n0 n1 f ls ≠ .div + := by + intro ls; induction ls <;> simp [result_rel, map_fix_fuel, map] + simp [fix] + case cons hd tl Hi => + -- Instantiate the first n and do a case disjunction + intro Hf; exists (least (fix_fuel_P f hd)); revert Hf + cases fix_fuel (least (fix_fuel_P f hd)) f hd <;> simp + -- Use the induction hyp + have Hr := Classical.em (map (fix f) tl = .div) + simp [fix] at * + cases Hr <;> simp_all + have Hj : ∃ n2, map (fix_fuel n2 f) tl ≠ .div := exists_map_fix_fuel_not_div_imp Hmono Hi + revert Hj; intro ⟨ n2, Hj ⟩ + intro Hf; exists n2; revert Hf + revert Hj; cases map (fix_fuel n2 f) tl <;> simp_all + + theorem map_is_cont_contra {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : + ∀ ls, map (fix f) ls ≠ .div → + ∃ n, map (fix_fuel n f) ls ≠ .div + := by + intro ls Hf + have Hc := map_is_cont_contra_aux Hmono ls Hf + apply exists_map_fix_fuel_not_div_imp <;> assumption + + theorem map_is_cont {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : + ∀ ls, (Hc : ∀ n, map (fix_fuel n f) ls = .div) → + map (fix f) ls = .div + := by + intro ls Hc + -- TODO: is there a tactic for proofs by contraposition? + apply Classical.byContradiction; intro Hndiv + let ⟨ n, Hcc ⟩ := map_is_cont_contra Hmono ls Hndiv + simp_all + + /- An example which uses map -/ + inductive Tree (a : Type) := + | leaf (x : a) + | node (tl : List (Tree a)) + + def id_body (f : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := + match t with + | .leaf x => .ret (.leaf x) + | .node tl => + match map f tl with + | .div => .div + | .fail e => .fail e + | .ret tl => .ret (.node tl) + + theorem id_body_mono : is_mono (@id_body a) := by + simp [is_mono]; intro g h Hr t; simp [result_rel, id_body] + cases t <;> simp + rename_i tl + have Hmmono := map_is_mono Hr tl + revert Hmmono; simp [result_rel] + cases map g tl <;> simp_all + + theorem id_body_cont : is_cont (@id_body a) := by + rw [is_cont]; intro t Hdiv + simp [fix_fuel] at * + -- TODO: weird things are happening with the rewriter and the simplifier here + rw [id_body] + simp [id_body] at Hdiv + -- + cases t <;> simp at * + rename_i tl + -- TODO: automate this + have Hmc := map_is_cont id_body_mono tl + have Hdiv : ∀ (n : ℕ), map (fix_fuel n id_body) tl = Result.div := by + intro n + have Hdiv := Hdiv n; revert Hdiv + cases map (fix_fuel n id_body) tl <;> simp_all + have Hmc := Hmc Hdiv; revert Hmc + cases map (fix id_body) tl <;> simp_all + + noncomputable def id (t : Tree a) := fix id_body t + + theorem id_eq (t : Tree a) : + id t = + match t with + | .leaf x => .ret (.leaf x) + | .node tl => + match map id tl with + | .div => .div + | .fail e => .fail e + | .ret tl => .ret (.node tl) + := by + have Hvalid : is_valid (@id_body a) := + is_valid.intro id_body_mono id_body_cont + have Heq := fix_fixed_eq (@id_body a) Hvalid + conv => lhs; rw [id, Heq, id_body] + +end Ex2 + end Diverge -- cgit v1.2.3 From ccc97b46c166a45255096d3fec2444c90f7c5aaa Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 14 Jun 2023 11:24:58 +0200 Subject: Make minor modifications --- backends/lean/Base/Diverge.lean | 31 +++++++++++++++++++++++++------ 1 file changed, 25 insertions(+), 6 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index b5264d0d..37d8eb27 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -11,6 +11,27 @@ open Result variable {a b : Type} +/- +TODO: +- we want an easier to use cases: + - keeps in the goal an equation of the shape: `t = case` + - if called on Prop terms, uses Classical.em + Actually, the cases from mathlib seems already quite powerful + (https://leanprover-community.github.io/mathlib_docs/tactics.html#cases) + For instance: cases h : e + Also: cases_matching +- better split tactic +- we need conversions to operate on the head of applications. + Actually, something like this works: + ``` + conv at Hl => + apply congr_fun + simp [fix_fuel_P] + ``` + Maybe we need a rpt ... ; focus? +- simplifier/rewriter have a strange behavior sometimes +-/ + /-! # The least fixed point definition and its properties -/ def least_p (p : Nat → Prop) (n : Nat) : Prop := p n ∧ (∀ m, m < n → ¬ p m) @@ -115,9 +136,7 @@ theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono match m with | 0 => exfalso - -- TODO: annoying to do those conversions by hand - try zify? - have : n1 + 1 ≤ (0 : Int) := by simp [*] at * - have : 0 ≤ n1 := by simp [*] at * + zify at * linarith | Nat.succ m1 => simp_arith at Hle @@ -188,9 +207,9 @@ theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono fix f x = f (fix f) x := by have Hl := fix_fuel_P_least Hmono He -- TODO: better control of simplification - have Heq : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := - by simp [fix_fuel_P] - simp [Heq] at Hl; clear Heq + conv at Hl => + apply congr_fun + simp [fix_fuel_P] -- The least upper bound is > 0 have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by revert Hl -- cgit v1.2.3 From 75f5f8a68b0ce028689c1d880ec99448e6d8dc3a Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 15:03:00 +0200 Subject: Make progress on making the proofs in Diverge more systematic --- backends/lean/Base/Diverge.lean | 260 ++++++++++++++++++++++++++++++++++++++-- 1 file changed, 252 insertions(+), 8 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 37d8eb27..0eff17e3 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -1,12 +1,76 @@ import Lean -import Base.Primitives +import Lean.Meta.Tactic.Simp +import Init.Data.List.Basic +import Mathlib.Tactic.RunCmd +import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Tauto namespace Diverge -open Primitives +namespace Primitives + +inductive Error where + | assertionFailure: Error + | integerOverflow: Error + | divisionByZero: Error + | arrayOutOfBounds: Error + | maximumSizeExceeded: Error + | panic: Error +deriving Repr, BEq + +open Error + +inductive Result (α : Type u) where + | ret (v: α): Result α + | fail (e: Error): Result α + | div +deriving Repr, BEq + +open Result + +-- instance Result_Inhabited (α : Type u) : Inhabited (Result α) := +-- Inhabited.mk (fail panic) + +-- instance Result_Nonempty (α : Type u) : Nonempty (Result α) := +-- Nonempty.intro div + +def bind (x: Result α) (f: α -> Result β) : Result β := + match x with + | ret v => f v + | fail v => fail v + | div => div + +@[simp] theorem bind_ret (x : α) (f : α → Result β) : bind (.ret x) f = f x := by simp [bind] +@[simp] theorem bind_fail (x : Error) (f : α → Result β) : bind (.fail x) f = .fail x := by simp [bind] +@[simp] theorem bind_div (f : α → Result β) : bind .div f = .div := by simp [bind] + +-- Allows using Result in do-blocks +instance : Bind Result where + bind := bind + +-- Allows using return x in do-blocks +instance : Pure Result where + pure := fun x => ret x + +@[simp] theorem bind_tc_ret (x : α) (f : α → Result β) : + (do let y ← .ret x; f y) = f x := by simp [Bind.bind, bind] + +@[simp] theorem bind_tc_fail (x : Error) (f : α → Result β) : + (do let y ← fail x; f y) = fail x := by simp [Bind.bind, bind] + +@[simp] theorem bind_tc_div (f : α → Result β) : + (do let y ← div; f y) = div := by simp [Bind.bind, bind] + +def div? {α: Type} (r: Result α): Bool := + match r with + | div => true + | ret _ | fail _ => false + +end Primitives namespace Fix +open Primitives open Result variable {a b : Type} @@ -123,7 +187,7 @@ structure is_valid (f : (a → Result b) → a → Result b) := /- - -/ + -/ theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : ∀ {{n m}}, n ≤ m → marrow_rel (fix_fuel n f) (fix_fuel m f) := by @@ -244,7 +308,38 @@ theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_vali have Hcont := Hvalid.hcont x He simp [Hcont] | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hvalid.hmono x n He + + /- Making the proofs more systematic -/ + -- TODO: rewrite is_mono in terms of is_mono_p + def is_mono_p (body : (a → Result b) → Result b) : Prop := + ∀ {{g h}}, marrow_rel g h → result_rel (body g) (body h) + + @[simp] theorem is_mono_p_same (x : Result b) : + @is_mono_p a b (λ _ => x) := by + simp [is_mono_p, marrow_rel, result_rel] + split <;> simp + + -- TODO: generalize + @[simp] theorem is_mono_p_tail_rec (x : a) : + @is_mono_p a b (λ f => f x) := by + simp_all [is_mono_p, marrow_rel, result_rel] + + -- TODO: rewrite is_cont in terms of is_cont_p + def is_cont_p (f : (a → Result b) → a → Result b) + (body : (a → Result b) → Result b) : Prop := + (Hc : ∀ n, body (fix_fuel n f) = .div) → + body (fix f) = .div + + @[simp] theorem is_cont_p_same (f : (a → Result b) → a → Result b) (x : Result b) : + is_cont_p f (λ _ => x) := by + simp [is_cont_p] + + -- TODO: generalize + @[simp] theorem is_cont_p_tail_rec (f : (a → Result b) → a → Result b) (x : a) : + is_cont_p f (λ f => f x) := by + simp_all [is_cont_p, fix] + /- (∀ n, fix_fuel n f x = div) @@ -269,7 +364,7 @@ end Fix namespace Ex1 /- An example of use of the fixed-point -/ - open Fix + open Primitives Fix variable {a : Type} (f : (List a × Int) → Result a) @@ -298,6 +393,20 @@ namespace Ex1 -- Recursive call apply Hdiv + /- Making the monotonicity/continuity proofs more systematic -/ + + theorem list_nth_body_mono2 : ∀ x, is_mono_p (λ f => @list_nth_body a f x) := by + intro x + simp [list_nth_body] + split <;> simp + split <;> simp + + theorem list_nth_body_cont2: ∀ f x, is_cont_p f (λ f => @list_nth_body a f x) := by + intro f x + simp [list_nth_body] + split <;> simp + split <;> simp + noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) @@ -319,8 +428,142 @@ namespace Ex1 end Ex1 namespace Ex2 + /- Same as Ex1, but we make the body of nth non tail-rec -/ + open Primitives Fix + + variable {a : Type} (f : (List a × Int) → Result a) + + def list_nth_body (x : (List a × Int)) : Result a := + let (ls, i) := x + match ls with + | [] => .fail .panic + | hd :: tl => + if i = 0 then .ret hd + else + do + let y ← f (tl, i - 1) + .ret y + + -- Lean is good at applying lemmas: we can write a very general version + theorem is_mono_p_bind + (g : (a → Result b) → Result b) + (h : b → (a → Result b) → Result b) : + is_mono_p g → + (∀ y, is_mono_p (h y)) → + is_mono_p (λ f => do let y ← g f; h y f) := by + intro hg hh + simp [is_mono_p] + intro fg fh Hrgh + simp [marrow_rel, result_rel] + have hg := hg Hrgh; simp [result_rel] at hg + cases heq0: g fg <;> simp_all + rename_i y _ + have hh := hh y Hrgh; simp [result_rel] at hh + simp_all + + -- Lean is good at applying lemmas: we can write a very general version + theorem is_cont_p_bind + (f : (a → Result b) → a → Result b) + (g : (a → Result b) → Result b) + (h : b → (a → Result b) → Result b) : + is_cont_p f (λ f => g f) → + (∀ y, is_cont_p f (h y)) → + is_cont_p f (λ f => do let y ← g f; h y f) := by + intro Hg Hh + simp [is_cont_p] + intro Hdiv + -- Case on `g (fix... f)`: is there an n s.t. it terminates? + cases Classical.em (∀ n, g (fix_fuel n f) = .div) <;> rename_i Hn + . -- Case 1: g diverges + have Hg := Hg Hn + simp_all + . -- Case 2: g doesn't diverge + simp at Hn + let ⟨ n, Hn ⟩ := Hn + have Hdivn := Hdiv n + -- TODO: we need monotonicity of g and f + have Hgmono : is_mono_p g := by sorry + have Hfmono : is_mono f := by sorry + have Hffmono := fix_fuel_fix_mono Hfmono n + have Hgeq := Hgmono Hffmono + simp [result_rel] at Hgeq + cases Heq: g (fix_fuel n f) <;> rename_i y <;> simp_all + -- Remains the .ret case + -- TODO: we need monotonicity of h? + have Hhmono : is_mono_p (h y) := by sorry + -- Use Hdiv to prove that: ∀ n, h y (fix_fuel n f) = div + -- We do this in two steps: first we prove it for m ≥ n + have Hhdiv: ∀ m, h y (fix_fuel m f) = .div := by + have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m f) = .div := by + -- We use the fact that `g (fix_fuel n f) = .div`, combined with Hdiv + intro m Hle + have Hdivm := Hdiv m + -- Monotonicity of g + have Hffmono := fix_fuel_mono Hfmono Hle + have Hgmono := Hgmono Hffmono + -- We need to clear Hdiv because otherwise simp_all rewrites Hdivm with Hdiv + clear Hdiv + simp_all [result_rel] + intro m + -- TODO: we shouldn't need the excluded middle here because it is decidable + cases Classical.em (n ≤ m) <;> rename_i Hl + . apply Hhdiv; assumption + . simp at Hl + -- Make a case disjunction on `h y (fix_fuel m f)`: if it is not equal + -- to div, use the monotonicity of `h y` + have Hle : m ≤ n := by linarith + have Hffmono := fix_fuel_mono Hfmono Hle + have Hmono := Hhmono Hffmono + simp [result_rel] at Hmono + cases Heq: h y (fix_fuel m f) <;> simp_all + -- We can now use the continuity hypothesis for h + apply Hh; assumption + + + -- TODO: what is the name of this theorem? + -- theorem eta_app_eq (x : a) (f : a → b) : f x = (λ x => f x) x := by simp + -- theorem eta_eq (x : a) (f : a → b) : (λ x => f x) = f := by simp + + --set_option pp.funBinderTypes true + --set_option pp.explicit true + --set_option pp.notation false + + theorem list_nth_body_mono : ∀ x, is_mono_p (λ f => @list_nth_body a f x) := by + intro x + simp [list_nth_body] + split <;> simp + split <;> simp + apply is_mono_p_bind <;> intros <;> simp + + theorem list_nth_body_cont2: ∀ f x, is_cont_p f (λ f => @list_nth_body a f x) := by + intro f x + simp [list_nth_body] + split <;> simp + split <;> simp + + noncomputable + def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) + + theorem list_nth_eq (ls : List a) (i : Int) : + list_nth ls i = + match ls with + | [] => .fail .panic + | hd :: tl => + if i = 0 then .ret hd + else list_nth tl (i - 1) + := by + have Hvalid : is_valid (@list_nth_body a) := + is_valid.intro list_nth_body_mono list_nth_body_cont + have Heq := fix_fixed_eq (@list_nth_body a) Hvalid + simp [Heq, list_nth] + conv => lhs; rw [list_nth_body] + simp [Heq] + +end Ex2 + +namespace Ex3 /- Higher-order example -/ - open Fix + open Primitives Fix variable {a b : Type} @@ -330,6 +573,7 @@ namespace Ex2 | [] => .ret [] | hd :: tl => do + -- TODO: monadic syntax match f hd with | .ret hd => match map f tl with @@ -423,7 +667,7 @@ namespace Ex2 have Hc := map_is_cont_contra_aux Hmono ls Hf apply exists_map_fix_fuel_not_div_imp <;> assumption - theorem map_is_cont {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : + theorem map_is_cont {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : ∀ ls, (Hc : ∀ n, map (fix_fuel n f) ls = .div) → map (fix f) ls = .div := by @@ -431,7 +675,7 @@ namespace Ex2 -- TODO: is there a tactic for proofs by contraposition? apply Classical.byContradiction; intro Hndiv let ⟨ n, Hcc ⟩ := map_is_cont_contra Hmono ls Hndiv - simp_all + simp_all /- An example which uses map -/ inductive Tree (a : Type) := @@ -490,6 +734,6 @@ namespace Ex2 have Heq := fix_fixed_eq (@id_body a) Hvalid conv => lhs; rw [id, Heq, id_body] -end Ex2 +end Ex3 end Diverge -- cgit v1.2.3 From 6297cdd89299452f8043f7aed75cf2eb01d31e24 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 16:20:35 +0200 Subject: Further simplify the proofs in Diverge.lean --- backends/lean/Base/Diverge.lean | 273 ++++++++++++++++++++++------------------ 1 file changed, 151 insertions(+), 122 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 0eff17e3..759773c9 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -4,6 +4,7 @@ import Init.Data.List.Basic import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith import Mathlib.Tactic.Tauto +--import Mathlib.Logic namespace Diverge @@ -291,7 +292,7 @@ theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono split <;> simp [*] <;> intros <;> simp [*] -- The final fixed point equation -theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_valid f) : +theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_valid f) : ∀ x, fix f x = f (fix f) x := by intros x -- conv => lhs; simp [fix] @@ -320,7 +321,7 @@ theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_vali simp [is_mono_p, marrow_rel, result_rel] split <;> simp - -- TODO: generalize + -- TODO: remove @[simp] theorem is_mono_p_tail_rec (x : a) : @is_mono_p a b (λ f => f x) := by simp_all [is_mono_p, marrow_rel, result_rel] @@ -335,30 +336,147 @@ theorem fix_fixed_eq (f : (a → Result b) → a → Result b) (Hvalid : is_vali is_cont_p f (λ _ => x) := by simp [is_cont_p] - -- TODO: generalize + -- TODO: remove @[simp] theorem is_cont_p_tail_rec (f : (a → Result b) → a → Result b) (x : a) : is_cont_p f (λ f => f x) := by simp_all [is_cont_p, fix] -/- -(∀ n, fix_fuel n f x = div) - -⊢ f (fun y => fix_fuel (least (fix_fuel_P f y)) f y) x = div - -(? x. p x) ==> p (epsilon p) - - -P (nf : a -> option Nat) := - match nf x with - | None => forall n, fix_fuel n f x = div - | Some n => fix_fuel n f x <> div + -- Lean is good at unification: we can write a very general version + theorem is_mono_p_bind + (g : (a → Result b) → Result b) + (h : b → (a → Result b) → Result b) : + is_mono_p g → + (∀ y, is_mono_p (h y)) → + is_mono_p (λ f => do let y ← g f; h y f) := by + intro hg hh + simp [is_mono_p] + intro fg fh Hrgh + simp [marrow_rel, result_rel] + have hg := hg Hrgh; simp [result_rel] at hg + cases heq0: g fg <;> simp_all + rename_i y _ + have hh := hh y Hrgh; simp [result_rel] at hh + simp_all -TODO: theorem de Tarsky, -Gilles Dowek (Introduction à la théorie des langages de programmation) + -- Lean is good at unification: we can write a very general version + -- (in particular, it will manage to figure out `g` and `h` when we + -- apply the lemma) + theorem is_cont_p_bind + (f : (a → Result b) → a → Result b) + (Hfmono : is_mono f) + (g : (a → Result b) → Result b) + (h : b → (a → Result b) → Result b) : + is_mono_p g → + is_cont_p f g → + (∀ y, is_mono_p (h y)) → + (∀ y, is_cont_p f (h y)) → + is_cont_p f (λ f => do let y ← g f; h y f) := by + intro Hgmono Hgcont Hhmono Hhcont + simp [is_cont_p] + intro Hdiv + -- Case on `g (fix... f)`: is there an n s.t. it terminates? + cases Classical.em (∀ n, g (fix_fuel n f) = .div) <;> rename_i Hn + . -- Case 1: g diverges + have Hgcont := Hgcont Hn + simp_all + . -- Case 2: g doesn't diverge + simp at Hn + let ⟨ n, Hn ⟩ := Hn + have Hdivn := Hdiv n + have Hffmono := fix_fuel_fix_mono Hfmono n + have Hgeq := Hgmono Hffmono + simp [result_rel] at Hgeq + cases Heq: g (fix_fuel n f) <;> rename_i y <;> simp_all + -- Remains the .ret case + -- Use Hdiv to prove that: ∀ n, h y (fix_fuel n f) = div + -- We do this in two steps: first we prove it for m ≥ n + have Hhdiv: ∀ m, h y (fix_fuel m f) = .div := by + have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m f) = .div := by + -- We use the fact that `g (fix_fuel n f) = .div`, combined with Hdiv + intro m Hle + have Hdivm := Hdiv m + -- Monotonicity of g + have Hffmono := fix_fuel_mono Hfmono Hle + have Hgmono := Hgmono Hffmono + -- We need to clear Hdiv because otherwise simp_all rewrites Hdivm with Hdiv + clear Hdiv + simp_all [result_rel] + intro m + -- TODO: we shouldn't need the excluded middle here because it is decidable + cases Classical.em (n ≤ m) <;> rename_i Hl + . apply Hhdiv; assumption + . simp at Hl + -- Make a case disjunction on `h y (fix_fuel m f)`: if it is not equal + -- to div, use the monotonicity of `h y` + have Hle : m ≤ n := by linarith + have Hffmono := fix_fuel_mono Hfmono Hle + have Hmono := Hhmono y Hffmono + simp [result_rel] at Hmono + cases Heq: h y (fix_fuel m f) <;> simp_all + -- We can now use the continuity hypothesis for h + apply Hhcont; assumption -fix_f is f s.t.: f x = f (fix f) x ∧ ! g. g x = g (fix g) x ==> f <= g + -- TODO: move + def is_valid_p (f : (a → Result b) → a → Result b) + (body : (a → Result b) → Result b) : Prop := + is_mono_p body ∧ + (is_mono f → is_cont_p f body) + + @[simp] theorem is_valid_p_same (f : (a → Result b) → a → Result b) (x : Result b) : + is_valid_p f (λ _ => x) := by + simp [is_valid_p] + + @[simp] theorem is_valid_p_rec (f : (a → Result b) → a → Result b) (x : a) : + is_valid_p f (λ f => f x) := by + simp [is_valid_p] + + -- Lean is good at unification: we can write a very general version + -- (in particular, it will manage to figure out `g` and `h` when we + -- apply the lemma) + theorem is_valid_p_bind + {{f : (a → Result b) → a → Result b}} + {{g : (a → Result b) → Result b}} + {{h : b → (a → Result b) → Result b}} + (Hgvalid : is_valid_p f g) + (Hhvalid : ∀ y, is_valid_p f (h y)) : + is_valid_p f (λ f => do let y ← g f; h y f) := by + let ⟨ Hgmono, Hgcont ⟩ := Hgvalid + -- TODO: conversion to move forall below and conjunction? + simp [is_valid_p, forall_and] at Hhvalid + have ⟨ Hhmono, Hhcont ⟩ := Hhvalid + simp [← imp_forall_iff] at Hhcont + simp [is_valid_p]; constructor + . -- Monotonicity + apply is_mono_p_bind <;> assumption + . -- Continuity + intro Hfmono + have Hgcont := Hgcont Hfmono + have Hhcont := Hhcont Hfmono + apply is_cont_p_bind <;> assumption + + theorem is_valid_p_imp_is_valid {{body : (a → Result b) → a → Result b}} + (Hvalid : ∀ f x, is_valid_p f (λ f => body f x)) : + is_valid body := by + have Hmono : is_mono body := by + intro f h Hr x + have Hmono := Hvalid (λ _ _ => .div) x + have Hmono := Hmono.left + apply Hmono; assumption + have Hcont : is_cont body := by + intro x Hdiv + have Hcont := (Hvalid body x).right Hmono + simp [is_cont_p] at Hcont + apply Hcont + intro n + have Hdiv := Hdiv n + simp [fix_fuel] at Hdiv + simp [*] + apply is_valid.intro Hmono Hcont --/ + theorem is_valid_p_fix_fixed_eq {{body : (a → Result b) → a → Result b}} + (Hvalid : ∀ f x, is_valid_p f (λ f => body f x)) : + ∀ x, fix body x = body (fix body) x := + fix_fixed_eq (is_valid_p_imp_is_valid Hvalid) end Fix @@ -420,7 +538,7 @@ namespace Ex1 := by have Hvalid : is_valid (@list_nth_body a) := is_valid.intro list_nth_body_mono list_nth_body_cont - have Heq := fix_fixed_eq (@list_nth_body a) Hvalid + have Heq := fix_fixed_eq Hvalid simp [Heq, list_nth] conv => lhs; rw [list_nth_body] simp [Heq] @@ -444,117 +562,28 @@ namespace Ex2 let y ← f (tl, i - 1) .ret y - -- Lean is good at applying lemmas: we can write a very general version - theorem is_mono_p_bind - (g : (a → Result b) → Result b) - (h : b → (a → Result b) → Result b) : - is_mono_p g → - (∀ y, is_mono_p (h y)) → - is_mono_p (λ f => do let y ← g f; h y f) := by - intro hg hh - simp [is_mono_p] - intro fg fh Hrgh - simp [marrow_rel, result_rel] - have hg := hg Hrgh; simp [result_rel] at hg - cases heq0: g fg <;> simp_all - rename_i y _ - have hh := hh y Hrgh; simp [result_rel] at hh - simp_all - - -- Lean is good at applying lemmas: we can write a very general version - theorem is_cont_p_bind - (f : (a → Result b) → a → Result b) - (g : (a → Result b) → Result b) - (h : b → (a → Result b) → Result b) : - is_cont_p f (λ f => g f) → - (∀ y, is_cont_p f (h y)) → - is_cont_p f (λ f => do let y ← g f; h y f) := by - intro Hg Hh - simp [is_cont_p] - intro Hdiv - -- Case on `g (fix... f)`: is there an n s.t. it terminates? - cases Classical.em (∀ n, g (fix_fuel n f) = .div) <;> rename_i Hn - . -- Case 1: g diverges - have Hg := Hg Hn - simp_all - . -- Case 2: g doesn't diverge - simp at Hn - let ⟨ n, Hn ⟩ := Hn - have Hdivn := Hdiv n - -- TODO: we need monotonicity of g and f - have Hgmono : is_mono_p g := by sorry - have Hfmono : is_mono f := by sorry - have Hffmono := fix_fuel_fix_mono Hfmono n - have Hgeq := Hgmono Hffmono - simp [result_rel] at Hgeq - cases Heq: g (fix_fuel n f) <;> rename_i y <;> simp_all - -- Remains the .ret case - -- TODO: we need monotonicity of h? - have Hhmono : is_mono_p (h y) := by sorry - -- Use Hdiv to prove that: ∀ n, h y (fix_fuel n f) = div - -- We do this in two steps: first we prove it for m ≥ n - have Hhdiv: ∀ m, h y (fix_fuel m f) = .div := by - have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m f) = .div := by - -- We use the fact that `g (fix_fuel n f) = .div`, combined with Hdiv - intro m Hle - have Hdivm := Hdiv m - -- Monotonicity of g - have Hffmono := fix_fuel_mono Hfmono Hle - have Hgmono := Hgmono Hffmono - -- We need to clear Hdiv because otherwise simp_all rewrites Hdivm with Hdiv - clear Hdiv - simp_all [result_rel] - intro m - -- TODO: we shouldn't need the excluded middle here because it is decidable - cases Classical.em (n ≤ m) <;> rename_i Hl - . apply Hhdiv; assumption - . simp at Hl - -- Make a case disjunction on `h y (fix_fuel m f)`: if it is not equal - -- to div, use the monotonicity of `h y` - have Hle : m ≤ n := by linarith - have Hffmono := fix_fuel_mono Hfmono Hle - have Hmono := Hhmono Hffmono - simp [result_rel] at Hmono - cases Heq: h y (fix_fuel m f) <;> simp_all - -- We can now use the continuity hypothesis for h - apply Hh; assumption - - - -- TODO: what is the name of this theorem? - -- theorem eta_app_eq (x : a) (f : a → b) : f x = (λ x => f x) x := by simp - -- theorem eta_eq (x : a) (f : a → b) : (λ x => f x) = f := by simp - - --set_option pp.funBinderTypes true - --set_option pp.explicit true - --set_option pp.notation false - - theorem list_nth_body_mono : ∀ x, is_mono_p (λ f => @list_nth_body a f x) := by - intro x - simp [list_nth_body] - split <;> simp - split <;> simp - apply is_mono_p_bind <;> intros <;> simp - - theorem list_nth_body_cont2: ∀ f x, is_cont_p f (λ f => @list_nth_body a f x) := by + theorem list_nth_body_valid: ∀ f x, is_valid_p f (λ f => @list_nth_body a f x) := by intro f x simp [list_nth_body] split <;> simp split <;> simp + apply is_valid_p_bind <;> intros <;> simp_all noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) theorem list_nth_eq (ls : List a) (i : Int) : - list_nth ls i = - match ls with - | [] => .fail .panic - | hd :: tl => - if i = 0 then .ret hd - else list_nth tl (i - 1) + (list_nth ls i = + match ls with + | [] => .fail .panic + | hd :: tl => + if i = 0 then .ret hd + else + do + let y ← list_nth tl (i - 1) + .ret y) := by - have Hvalid : is_valid (@list_nth_body a) := - is_valid.intro list_nth_body_mono list_nth_body_cont - have Heq := fix_fixed_eq (@list_nth_body a) Hvalid + have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_valid a) simp [Heq, list_nth] conv => lhs; rw [list_nth_body] simp [Heq] @@ -731,7 +760,7 @@ namespace Ex3 := by have Hvalid : is_valid (@id_body a) := is_valid.intro id_body_mono id_body_cont - have Heq := fix_fixed_eq (@id_body a) Hvalid + have Heq := fix_fixed_eq Hvalid conv => lhs; rw [id, Heq, id_body] end Ex3 -- cgit v1.2.3 From 34a471c02d6c49aa34b7f353b28b90b09a69864a Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 17:10:24 +0200 Subject: Simplify the id example in Diverge.lean --- backends/lean/Base/Diverge.lean | 119 ++++++++++++++++++++++++++++++++-------- 1 file changed, 95 insertions(+), 24 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 759773c9..2c764c5e 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -74,7 +74,7 @@ namespace Fix open Primitives open Result -variable {a b : Type} +variable {a b c d : Type} /- TODO: @@ -292,6 +292,7 @@ theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono split <;> simp [*] <;> intros <;> simp [*] -- The final fixed point equation +-- TODO: remove the `forall x` theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_valid f) : ∀ x, fix f x = f (fix f) x := by intros x @@ -313,26 +314,26 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va /- Making the proofs more systematic -/ -- TODO: rewrite is_mono in terms of is_mono_p - def is_mono_p (body : (a → Result b) → Result b) : Prop := + def is_mono_p (body : (a → Result b) → Result c) : Prop := ∀ {{g h}}, marrow_rel g h → result_rel (body g) (body h) - @[simp] theorem is_mono_p_same (x : Result b) : - @is_mono_p a b (λ _ => x) := by + @[simp] theorem is_mono_p_same (x : Result c) : + @is_mono_p a b c (λ _ => x) := by simp [is_mono_p, marrow_rel, result_rel] split <;> simp -- TODO: remove @[simp] theorem is_mono_p_tail_rec (x : a) : - @is_mono_p a b (λ f => f x) := by + @is_mono_p a b b (λ f => f x) := by simp_all [is_mono_p, marrow_rel, result_rel] -- TODO: rewrite is_cont in terms of is_cont_p def is_cont_p (f : (a → Result b) → a → Result b) - (body : (a → Result b) → Result b) : Prop := + (body : (a → Result b) → Result c) : Prop := (Hc : ∀ n, body (fix_fuel n f) = .div) → body (fix f) = .div - @[simp] theorem is_cont_p_same (f : (a → Result b) → a → Result b) (x : Result b) : + @[simp] theorem is_cont_p_same (f : (a → Result b) → a → Result b) (x : Result c) : is_cont_p f (λ _ => x) := by simp [is_cont_p] @@ -343,8 +344,8 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va -- Lean is good at unification: we can write a very general version theorem is_mono_p_bind - (g : (a → Result b) → Result b) - (h : b → (a → Result b) → Result b) : + (g : (a → Result b) → Result c) + (h : c → (a → Result b) → Result d) : is_mono_p g → (∀ y, is_mono_p (h y)) → is_mono_p (λ f => do let y ← g f; h y f) := by @@ -364,8 +365,8 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va theorem is_cont_p_bind (f : (a → Result b) → a → Result b) (Hfmono : is_mono f) - (g : (a → Result b) → Result b) - (h : b → (a → Result b) → Result b) : + (g : (a → Result b) → Result c) + (h : c → (a → Result b) → Result d) : is_mono_p g → is_cont_p f g → (∀ y, is_mono_p (h y)) → @@ -417,12 +418,12 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va apply Hhcont; assumption -- TODO: move - def is_valid_p (f : (a → Result b) → a → Result b) - (body : (a → Result b) → Result b) : Prop := + def is_valid_p (k : (a → Result b) → a → Result b) + (body : (a → Result b) → Result c) : Prop := is_mono_p body ∧ - (is_mono f → is_cont_p f body) + (is_mono k → is_cont_p k body) - @[simp] theorem is_valid_p_same (f : (a → Result b) → a → Result b) (x : Result b) : + @[simp] theorem is_valid_p_same (f : (a → Result b) → a → Result b) (x : Result c) : is_valid_p f (λ _ => x) := by simp [is_valid_p] @@ -435,8 +436,8 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va -- apply the lemma) theorem is_valid_p_bind {{f : (a → Result b) → a → Result b}} - {{g : (a → Result b) → Result b}} - {{h : b → (a → Result b) → Result b}} + {{g : (a → Result b) → Result c}} + {{h : c → (a → Result b) → Result d}} (Hgvalid : is_valid_p f g) (Hhvalid : ∀ y, is_valid_p f (h y)) : is_valid_p f (λ f => do let y ← g f; h y f) := by @@ -473,10 +474,12 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va simp [*] apply is_valid.intro Hmono Hcont + -- TODO: functional extensionality theorem is_valid_p_fix_fixed_eq {{body : (a → Result b) → a → Result b}} (Hvalid : ∀ f x, is_valid_p f (λ f => body f x)) : - ∀ x, fix body x = body (fix body) x := - fix_fixed_eq (is_valid_p_imp_is_valid Hvalid) + fix body = body (fix body) := by + apply funext + exact fix_fixed_eq (is_valid_p_imp_is_valid Hvalid) end Fix @@ -562,8 +565,8 @@ namespace Ex2 let y ← f (tl, i - 1) .ret y - theorem list_nth_body_valid: ∀ f x, is_valid_p f (λ f => @list_nth_body a f x) := by - intro f x + theorem list_nth_body_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by + intro k x simp [list_nth_body] split <;> simp split <;> simp @@ -584,9 +587,8 @@ namespace Ex2 .ret y) := by have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_valid a) - simp [Heq, list_nth] - conv => lhs; rw [list_nth_body] - simp [Heq] + simp [list_nth] + conv => lhs; rw [Heq] end Ex2 @@ -765,4 +767,73 @@ namespace Ex3 end Ex3 +namespace Ex4 + /- Higher-order example -/ + open Primitives Fix + + variable {a b : Type} + + /- An auxiliary function, which doesn't require the fixed-point -/ + def map (f : a → Result b) (ls : List a) : Result (List b) := + match ls with + | [] => .ret [] + | hd :: tl => + do + let hd ← f hd + let tl ← map f tl + .ret (hd :: tl) + + /- The validity theorem for `map`, generic in `f` -/ + /- TODO: rename the condition to k in all the lemma statements -/ + theorem map_is_valid + {{f : (a → Result b) → a → Result c}} + (Hfvalid : ∀ k x, is_valid_p k (λ k => f k x)) + (k : (a → Result b) → a → Result b) + (ls : List a) : + is_valid_p k (λ k => map (f k) ls) := by + induction ls <;> simp [map] + apply is_valid_p_bind <;> simp_all + intros + apply is_valid_p_bind <;> simp_all + + /- An example which uses map -/ + inductive Tree (a : Type) := + | leaf (x : a) + | node (tl : List (Tree a)) + + def id_body (f : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := + match t with + | .leaf x => .ret (.leaf x) + | .node tl => + do + let tl ← map f tl + .ret (.node tl) + + /- TODO: make the naming consistent (suffix with "_is") -/ + theorem id_body_is_valid : + ∀ k x, is_valid_p k (λ k => @id_body a k x) := by + intro k x + simp [id_body] + split <;> simp + apply is_valid_p_bind <;> simp_all + -- We have to show that `map k tl` is valid + apply map_is_valid; simp + + noncomputable def id (t : Tree a) := fix id_body t + + theorem id_eq (t : Tree a) : + (id t = + match t with + | .leaf x => .ret (.leaf x) + | .node tl => + do + let tl ← map id tl + .ret (.node tl)) + := by + have Heq := is_valid_p_fix_fixed_eq (@id_body_is_valid a) + simp [id] + conv => lhs; rw [Heq]; simp; rw [id_body] + +end Ex4 + end Diverge -- cgit v1.2.3 From 5d8eea6504d9dcfa43844d5ba51c7abf6c589701 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 17:26:51 +0200 Subject: Remove the obsolete examples from Diverge --- backends/lean/Base/Diverge.lean | 213 +--------------------------------------- 1 file changed, 5 insertions(+), 208 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 2c764c5e..0e3e96c3 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -497,33 +497,8 @@ namespace Ex1 if i = 0 then .ret hd else f (tl, i - 1) - theorem list_nth_body_mono : is_mono (@list_nth_body a) := by - simp [is_mono]; intro g h Hr (ls, i); simp [result_rel, list_nth_body] - cases ls <;> simp - rename_i hd tl - -- TODO: making a case disjunction over `i = 0` is annoying, we need a more - -- general tactic for this - cases (Classical.em (Eq i 0)) <;> simp [*] at * - apply Hr - - theorem list_nth_body_cont : is_cont (@list_nth_body a) := by - rw [is_cont]; intro (ls, i) Hdiv; simp [list_nth_body, fix_fuel] at * - cases ls <;> simp at * - -- TODO: automate this - cases (Classical.em (Eq i 0)) <;> simp [*] at * - -- Recursive call - apply Hdiv - - /- Making the monotonicity/continuity proofs more systematic -/ - - theorem list_nth_body_mono2 : ∀ x, is_mono_p (λ f => @list_nth_body a f x) := by - intro x - simp [list_nth_body] - split <;> simp - split <;> simp - - theorem list_nth_body_cont2: ∀ f x, is_cont_p f (λ f => @list_nth_body a f x) := by - intro f x + theorem list_nth_body_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by + intro k x simp [list_nth_body] split <;> simp split <;> simp @@ -539,12 +514,9 @@ namespace Ex1 if i = 0 then .ret hd else list_nth tl (i - 1) := by - have Hvalid : is_valid (@list_nth_body a) := - is_valid.intro list_nth_body_mono list_nth_body_cont - have Heq := fix_fixed_eq Hvalid - simp [Heq, list_nth] - conv => lhs; rw [list_nth_body] - simp [Heq] + have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_valid a) + simp [list_nth] + conv => lhs; rw [Heq] end Ex1 @@ -592,181 +564,6 @@ namespace Ex2 end Ex2 -namespace Ex3 - /- Higher-order example -/ - open Primitives Fix - - variable {a b : Type} - - /- An auxiliary function, which doesn't require the fixed-point -/ - def map (f : a → Result b) (ls : List a) : Result (List b) := - match ls with - | [] => .ret [] - | hd :: tl => - do - -- TODO: monadic syntax - match f hd with - | .ret hd => - match map f tl with - | .ret tl => - .ret (hd :: tl) - | r => r - | .fail e => .fail e - | .div => .div - - theorem map_is_mono {{f g : a → Result b}} (Hr : marrow_rel f g) : - ∀ ls, result_rel (map f ls) (map g ls) := by - intro ls; induction ls <;> simp [result_rel, map] - case cons hd tl Hi => - have Hr1 := Hr hd; simp [result_rel] at Hr1 - -- TODO: reverting is annoying - revert Hr1 - cases f hd <;> intro Hr1 <;> simp [*] - -- ret case - simp [result_rel] at Hi - -- TODO: annoying - revert Hi - cases map f tl <;> intro Hi <;> simp [*] - - -- Auxiliary definition - def map_fix_fuel (n0 n1 : Nat) (f : (a → Result b) → a → Result b) (ls : List a) : Result (List b) := - match ls with - | [] => .ret [] - | hd :: tl => - do - match fix_fuel n0 f hd with - | .ret hd => - match map (fix_fuel n1 f) tl with - | .ret tl => - .ret (hd :: tl) - | r => r - | .fail e => .fail e - | .div => .div - - def exists_map_fix_fuel_not_div_imp {{f : (a → Result b) → a → Result b}} {{ls : List a}} - (Hmono : is_mono f) : - (∃ n0 n1, map_fix_fuel n0 n1 f ls ≠ .div) → - ∃ n2, map (fix_fuel n2 f) ls ≠ .div := by - intro ⟨ n0, n1, Hnd ⟩ - exists n0 + n1 - have Hineq0 : n0 ≤ n0 + n1 := by linarith - have Hineq1 : n1 ≤ n0 + n1 := by linarith - simp [map_fix_fuel] at Hnd - -- TODO: I would like a rewrite_once tactic - unfold map; simp - -- - revert Hnd - cases ls <;> simp - rename_i hd tl - -- Use the monotonicity of fix_fuel - have Hfmono := fix_fuel_mono Hmono Hineq0 hd - simp [result_rel] at Hfmono; revert Hfmono - cases fix_fuel n0 f hd <;> intro <;> simp [*] - -- Use the monotonicity of map - have Hfmono := fix_fuel_mono Hmono Hineq1 - have Hmmono := map_is_mono Hfmono tl - simp [result_rel] at Hmmono; revert Hmmono - cases map (fix_fuel n1 f) tl <;> intro <;> simp [*] - - -- TODO: it is simpler to prove the contrapositive of is_cont than is_cont itself. - -- The proof is still quite technical: think of a criteria whose proof is simpler - -- to automate. - theorem map_is_cont_contra_aux {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : - ∀ ls, map (fix f) ls ≠ .div → - ∃ n0 n1, map_fix_fuel n0 n1 f ls ≠ .div - := by - intro ls; induction ls <;> simp [result_rel, map_fix_fuel, map] - simp [fix] - case cons hd tl Hi => - -- Instantiate the first n and do a case disjunction - intro Hf; exists (least (fix_fuel_P f hd)); revert Hf - cases fix_fuel (least (fix_fuel_P f hd)) f hd <;> simp - -- Use the induction hyp - have Hr := Classical.em (map (fix f) tl = .div) - simp [fix] at * - cases Hr <;> simp_all - have Hj : ∃ n2, map (fix_fuel n2 f) tl ≠ .div := exists_map_fix_fuel_not_div_imp Hmono Hi - revert Hj; intro ⟨ n2, Hj ⟩ - intro Hf; exists n2; revert Hf - revert Hj; cases map (fix_fuel n2 f) tl <;> simp_all - - theorem map_is_cont_contra {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : - ∀ ls, map (fix f) ls ≠ .div → - ∃ n, map (fix_fuel n f) ls ≠ .div - := by - intro ls Hf - have Hc := map_is_cont_contra_aux Hmono ls Hf - apply exists_map_fix_fuel_not_div_imp <;> assumption - - theorem map_is_cont {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) : - ∀ ls, (Hc : ∀ n, map (fix_fuel n f) ls = .div) → - map (fix f) ls = .div - := by - intro ls Hc - -- TODO: is there a tactic for proofs by contraposition? - apply Classical.byContradiction; intro Hndiv - let ⟨ n, Hcc ⟩ := map_is_cont_contra Hmono ls Hndiv - simp_all - - /- An example which uses map -/ - inductive Tree (a : Type) := - | leaf (x : a) - | node (tl : List (Tree a)) - - def id_body (f : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := - match t with - | .leaf x => .ret (.leaf x) - | .node tl => - match map f tl with - | .div => .div - | .fail e => .fail e - | .ret tl => .ret (.node tl) - - theorem id_body_mono : is_mono (@id_body a) := by - simp [is_mono]; intro g h Hr t; simp [result_rel, id_body] - cases t <;> simp - rename_i tl - have Hmmono := map_is_mono Hr tl - revert Hmmono; simp [result_rel] - cases map g tl <;> simp_all - - theorem id_body_cont : is_cont (@id_body a) := by - rw [is_cont]; intro t Hdiv - simp [fix_fuel] at * - -- TODO: weird things are happening with the rewriter and the simplifier here - rw [id_body] - simp [id_body] at Hdiv - -- - cases t <;> simp at * - rename_i tl - -- TODO: automate this - have Hmc := map_is_cont id_body_mono tl - have Hdiv : ∀ (n : ℕ), map (fix_fuel n id_body) tl = Result.div := by - intro n - have Hdiv := Hdiv n; revert Hdiv - cases map (fix_fuel n id_body) tl <;> simp_all - have Hmc := Hmc Hdiv; revert Hmc - cases map (fix id_body) tl <;> simp_all - - noncomputable def id (t : Tree a) := fix id_body t - - theorem id_eq (t : Tree a) : - id t = - match t with - | .leaf x => .ret (.leaf x) - | .node tl => - match map id tl with - | .div => .div - | .fail e => .fail e - | .ret tl => .ret (.node tl) - := by - have Hvalid : is_valid (@id_body a) := - is_valid.intro id_body_mono id_body_cont - have Heq := fix_fixed_eq Hvalid - conv => lhs; rw [id, Heq, id_body] - -end Ex3 - namespace Ex4 /- Higher-order example -/ open Primitives Fix -- cgit v1.2.3 From 8db6718d06023ffa77035b29ec92cec03ee838bc Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 18:13:29 +0200 Subject: Add an example with even/odd in Diverge.lean --- backends/lean/Base/Diverge.lean | 120 +++++++++++++++++++++++++++++++++++++++- 1 file changed, 119 insertions(+), 1 deletion(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 0e3e96c3..2e77c5e0 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -564,6 +564,124 @@ namespace Ex2 end Ex2 +namespace Ex3 + /- Mutually recursive functions -/ + open Primitives Fix + + /- Because we have mutually recursive functions, we use a sum for the inputs + and the output types: + - inputs: the sum allows to select the function to call in the recursive + calls (and the functions may not have the same input types) + - outpus: this case is degenerate because `even` and `odd` both have the + return type `Bool`, but generally speaking we need a sum type because + the functions in the mutually recursive group may not have the same + return type. + -/ + variable (f : (Int ⊕ Int) → Result (Bool ⊕ Bool)) + + def is_even_is_odd_body (x : (Int ⊕ Int)) : Result (Bool ⊕ Bool) := + match x with + | .inl i => + -- Body of `is_even` + if i = 0 + then .ret (.inl true) -- We return .inl because this is `is_even` + else + do + let b ← + do + -- Call `odd`: we need to wrap the input value in `.inr`, then + -- extract the output value + let r ← f (.inr (i- 1)) + match r with + | .inl _ => .fail .panic -- Invalid output + | .inr b => .ret b + -- Wrap the return value + .ret (.inl b) + | .inr i => + -- Body of `is_odd` + if i = 0 + then .ret (.inr false) -- We return .inr because this is `is_odd` + else + do + let b ← + do + -- Call `is_even`: we need to wrap the input value in .inr, then + -- extract the output value + let r ← f (.inl (i- 1)) + match r with + | .inl b => .ret b + | .inr _ => .fail .panic -- Invalid output + -- Wrap the return value + .ret (.inr b) + + theorem is_even_is_odd_body_is_valid: + ∀ k x, is_valid_p k (λ k => is_even_is_odd_body k x) := by + intro k x + simp [is_even_is_odd_body] + split <;> simp <;> split <;> simp + apply is_valid_p_bind; simp + intros; split <;> simp + apply is_valid_p_bind; simp + intros; split <;> simp + + noncomputable + def is_even (i : Int): Result Bool := + do + let r ← fix is_even_is_odd_body (.inl i) + match r with + | .inl b => .ret b + | .inr _ => .fail .panic + + noncomputable + def is_odd (i : Int): Result Bool := + do + let r ← fix is_even_is_odd_body (.inr i) + match r with + | .inl _ => .fail .panic + | .inr b => .ret b + + -- TODO: move + -- TODO: this is not enough + theorem swap_if_bind {a b : Type} (e : Prop) [Decidable e] (x y : Result a) (f : a → Result b) : + (do + let z ← (if e then x else y) + f z) + = + (if e then do let z ← x; f z + else do let z ← y; f z) := by + split <;> simp + + theorem is_even_eq (i : Int) : + is_even i = (if i = 0 then .ret true else is_odd (i - 1)) + := by + have Heq := is_valid_p_fix_fixed_eq is_even_is_odd_body_is_valid + simp [is_even, is_odd] + conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp + -- Very annoying: we need to swap the matches + -- Doing this with rewriting lemmas is hard generally speaking + -- (especially as we may have to generate lemmas for user-defined + -- inductives on the fly). + -- The simplest is to repeatedly split then simplify (we identify + -- the outer match or monadic let-binding, and split on its scrutinee) + split <;> simp + cases H0 : fix is_even_is_odd_body (Sum.inr (i - 1)) <;> simp + rename_i v + split <;> simp + + theorem is_odd_eq (i : Int) : + is_odd i = (if i = 0 then .ret false else is_even (i - 1)) + := by + have Heq := is_valid_p_fix_fixed_eq is_even_is_odd_body_is_valid + simp [is_even, is_odd] + conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp + -- Same remark as for `even` + split <;> simp + cases H0 : fix is_even_is_odd_body (Sum.inl (i - 1)) <;> simp + rename_i v + split <;> simp + +end Ex3 + namespace Ex4 /- Higher-order example -/ open Primitives Fix @@ -581,7 +699,7 @@ namespace Ex4 .ret (hd :: tl) /- The validity theorem for `map`, generic in `f` -/ - /- TODO: rename the condition to k in all the lemma statements -/ + /- TODO: rename the continuation to k in all the lemma statements -/ theorem map_is_valid {{f : (a → Result b) → a → Result c}} (Hfvalid : ∀ k x, is_valid_p k (λ k => f k x)) -- cgit v1.2.3 From a2670f4d097075c23b9affceb8ed8498b73c4b8c Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 19 Jun 2023 18:52:29 +0200 Subject: Cleanup Diverge.lean --- backends/lean/Base/Diverge.lean | 657 +++++++++++++++++++--------------------- 1 file changed, 320 insertions(+), 337 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 2e77c5e0..65c061bd 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -3,12 +3,32 @@ import Lean.Meta.Tactic.Simp import Init.Data.List.Basic import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith -import Mathlib.Tactic.Tauto ---import Mathlib.Logic + +/- +TODO: +- we want an easier to use cases: + - keeps in the goal an equation of the shape: `t = case` + - if called on Prop terms, uses Classical.em + Actually, the cases from mathlib seems already quite powerful + (https://leanprover-community.github.io/mathlib_docs/tactics.html#cases) + For instance: cases h : e + Also: cases_matching +- better split tactic +- we need conversions to operate on the head of applications. + Actually, something like this works: + ``` + conv at Hl => + apply congr_fun + simp [fix_fuel_P] + ``` + Maybe we need a rpt ... ; focus? +- simplifier/rewriter have a strange behavior sometimes +-/ namespace Diverge namespace Primitives +/-! # Copy-pasting from Primitives to make the file self-contained -/ inductive Error where | assertionFailure: Error @@ -29,12 +49,6 @@ deriving Repr, BEq open Result --- instance Result_Inhabited (α : Type u) : Inhabited (Result α) := --- Inhabited.mk (fail panic) - --- instance Result_Nonempty (α : Type u) : Nonempty (Result α) := --- Nonempty.intro div - def bind (x: Result α) (f: α -> Result β) : Result β := match x with | ret v => f v @@ -71,312 +85,276 @@ end Primitives namespace Fix -open Primitives -open Result - -variable {a b c d : Type} - -/- -TODO: -- we want an easier to use cases: - - keeps in the goal an equation of the shape: `t = case` - - if called on Prop terms, uses Classical.em - Actually, the cases from mathlib seems already quite powerful - (https://leanprover-community.github.io/mathlib_docs/tactics.html#cases) - For instance: cases h : e - Also: cases_matching -- better split tactic -- we need conversions to operate on the head of applications. - Actually, something like this works: - ``` - conv at Hl => - apply congr_fun - simp [fix_fuel_P] - ``` - Maybe we need a rpt ... ; focus? -- simplifier/rewriter have a strange behavior sometimes --/ - -/-! # The least fixed point definition and its properties -/ - -def least_p (p : Nat → Prop) (n : Nat) : Prop := p n ∧ (∀ m, m < n → ¬ p m) -noncomputable def least (p : Nat → Prop) : Nat := - Classical.epsilon (least_p p) - --- Auxiliary theorem for [least_spec]: if there exists an `n` satisfying `p`, --- there there exists a least `m` satisfying `p`. -theorem least_spec_aux (p : Nat → Prop) : ∀ (n : Nat), (hn : p n) → ∃ m, least_p p m := by - apply Nat.strongRec' - intros n hi hn - -- Case disjunction on: is n the smallest n satisfying p? - match Classical.em (∀ m, m < n → ¬ p m) with - | .inl hlt => - -- Yes: trivial - exists n - | .inr hlt => - simp at * - let ⟨ m, ⟨ hmlt, hm ⟩ ⟩ := hlt - have hi := hi m hmlt hm - apply hi - --- The specification of [least]: either `p` is never satisfied, or it is satisfied --- by `least p` and no `n < least p` satisfies `p`. -theorem least_spec (p : Nat → Prop) : (∀ n, ¬ p n) ∨ (p (least p) ∧ ∀ n, n < least p → ¬ p n) := by - -- Case disjunction on the existence of an `n` which satisfies `p` - match Classical.em (∀ n, ¬ p n) with - | .inl h => - -- There doesn't exist: trivial - apply (Or.inl h) - | .inr h => - -- There exists: we simply use `least_spec_aux` in combination with the property - -- of the epsilon operator - simp at * - let ⟨ n, hn ⟩ := h - apply Or.inr - have hl := least_spec_aux p n hn - have he := Classical.epsilon_spec hl - apply he - -/-! # The fixed point definitions -/ - -def fix_fuel (n : Nat) (f : (a → Result b) → a → Result b) (x : a) : Result b := - match n with - | 0 => .div - | n + 1 => - f (fix_fuel n f) x - -@[simp] def fix_fuel_pred (f : (a → Result b) → a → Result b) (x : a) (n : Nat) := - not (div? (fix_fuel n f x)) - -def fix_fuel_P (f : (a → Result b) → a → Result b) (x : a) (n : Nat) : Prop := - fix_fuel_pred f x n - -noncomputable def fix (f : (a → Result b) → a → Result b) (x : a) : Result b := - fix_fuel (least (fix_fuel_P f x)) f x - -/-! # The proof of the fixed point equation -/ - --- Monotonicity relation over results --- TODO: generalize -def result_rel {a : Type u} (x1 x2 : Result a) : Prop := - match x1 with - | div => True - | fail _ => x2 = x1 - | ret _ => x2 = x1 -- TODO: generalize - --- Monotonicity relation over monadic arrows --- TODO: Kleisli arrow --- TODO: generalize -def marrow_rel (f g : a → Result b) : Prop := - ∀ x, result_rel (f x) (g x) - --- Monotonicity property -def is_mono (f : (a → Result b) → a → Result b) : Prop := - ∀ {{g h}}, marrow_rel g h → marrow_rel (f g) (f h) - --- "Continuity" property. --- We need this, and this looks a lot like continuity. Also see this paper: --- https://inria.hal.science/file/index/docid/216187/filename/tarski.pdf -def is_cont (f : (a → Result b) → a → Result b) : Prop := - ∀ x, (Hdiv : ∀ n, fix_fuel (.succ n) f x = div) → f (fix f) x = div - --- Validity property for a body -structure is_valid (f : (a → Result b) → a → Result b) := - intro:: - hmono : is_mono f - hcont : is_cont f - -/- - - -/ - -theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : - ∀ {{n m}}, n ≤ m → marrow_rel (fix_fuel n f) (fix_fuel m f) := by - intros n - induction n - case zero => simp [marrow_rel, fix_fuel, result_rel] - case succ n1 Hi => - intros m Hle x + open Primitives + open Result + + variable {a b c d : Type} + + /-! # The least fixed point definition and its properties -/ + + def least_p (p : Nat → Prop) (n : Nat) : Prop := p n ∧ (∀ m, m < n → ¬ p m) + noncomputable def least (p : Nat → Prop) : Nat := + Classical.epsilon (least_p p) + + -- Auxiliary theorem for [least_spec]: if there exists an `n` satisfying `p`, + -- there there exists a least `m` satisfying `p`. + theorem least_spec_aux (p : Nat → Prop) : ∀ (n : Nat), (hn : p n) → ∃ m, least_p p m := by + apply Nat.strongRec' + intros n hi hn + -- Case disjunction on: is n the smallest n satisfying p? + match Classical.em (∀ m, m < n → ¬ p m) with + | .inl hlt => + -- Yes: trivial + exists n + | .inr hlt => + simp at * + let ⟨ m, ⟨ hmlt, hm ⟩ ⟩ := hlt + have hi := hi m hmlt hm + apply hi + + -- The specification of [least]: either `p` is never satisfied, or it is satisfied + -- by `least p` and no `n < least p` satisfies `p`. + theorem least_spec (p : Nat → Prop) : (∀ n, ¬ p n) ∨ (p (least p) ∧ ∀ n, n < least p → ¬ p n) := by + -- Case disjunction on the existence of an `n` which satisfies `p` + match Classical.em (∀ n, ¬ p n) with + | .inl h => + -- There doesn't exist: trivial + apply (Or.inl h) + | .inr h => + -- There exists: we simply use `least_spec_aux` in combination with the property + -- of the epsilon operator + simp at * + let ⟨ n, hn ⟩ := h + apply Or.inr + have hl := least_spec_aux p n hn + have he := Classical.epsilon_spec hl + apply he + + /-! # The fixed point definitions -/ + + def fix_fuel (n : Nat) (f : (a → Result b) → a → Result b) (x : a) : Result b := + match n with + | 0 => .div + | n + 1 => + f (fix_fuel n f) x + + @[simp] def fix_fuel_pred (f : (a → Result b) → a → Result b) (x : a) (n : Nat) := + not (div? (fix_fuel n f x)) + + def fix_fuel_P (f : (a → Result b) → a → Result b) (x : a) (n : Nat) : Prop := + fix_fuel_pred f x n + + noncomputable def fix (f : (a → Result b) → a → Result b) (x : a) : Result b := + fix_fuel (least (fix_fuel_P f x)) f x + + /-! # The validity property -/ + + -- Monotonicity relation over results + -- TODO: generalize (we should parameterize the definition by a relation over `a`) + def result_rel {a : Type u} (x1 x2 : Result a) : Prop := + match x1 with + | div => True + | fail _ => x2 = x1 + | ret _ => x2 = x1 -- TODO: generalize + + -- Monotonicity relation over monadic arrows (i.e., Kleisli arrows) + def karrow_rel (k1 k2 : a → Result b) : Prop := + ∀ x, result_rel (k1 x) (k2 x) + + -- Monotonicity property for function bodies + def is_mono (f : (a → Result b) → a → Result b) : Prop := + ∀ {{k1 k2}}, karrow_rel k1 k2 → karrow_rel (f k1) (f k2) + + -- "Continuity" property. + -- We need this, and this looks a lot like continuity. Also see this paper: + -- https://inria.hal.science/file/index/docid/216187/filename/tarski.pdf + -- We define our "continuity" criteria so that it gives us what we need to + -- prove the fixed-point equation, and we can also easily manipulate it. + def is_cont (f : (a → Result b) → a → Result b) : Prop := + ∀ x, (Hdiv : ∀ n, fix_fuel (.succ n) f x = div) → f (fix f) x = div + + /-! # The proof of the fixed-point equation -/ + theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + ∀ {{n m}}, n ≤ m → karrow_rel (fix_fuel n f) (fix_fuel m f) := by + intros n + induction n + case zero => simp [karrow_rel, fix_fuel, result_rel] + case succ n1 Hi => + intros m Hle x + simp [result_rel] + match m with + | 0 => + exfalso + zify at * + linarith + | Nat.succ m1 => + simp_arith at Hle + simp [fix_fuel] + have Hi := Hi Hle + have Hmono := Hmono Hi x + simp [result_rel] at Hmono + apply Hmono + + @[simp] theorem neg_fix_fuel_P {f : (a → Result b) → a → Result b} {x : a} {n : Nat} : + ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by + simp [fix_fuel_P, div?] + cases fix_fuel n f x <;> simp + + theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + ∀ n, karrow_rel (fix_fuel n f) (fix f) := by + intros n x simp [result_rel] - match m with - | 0 => - exfalso - zify at * - linarith - | Nat.succ m1 => - simp_arith at Hle - simp [fix_fuel] - have Hi := Hi Hle - have Hmono := Hmono Hi x - simp [result_rel] at Hmono - apply Hmono - -@[simp] theorem neg_fix_fuel_P {f : (a → Result b) → a → Result b} {x : a} {n : Nat} : - ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by - simp [fix_fuel_P, div?] - cases fix_fuel n f x <;> simp - -theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : - ∀ n, marrow_rel (fix_fuel n f) (fix f) := by - intros n x - simp [result_rel] - have Hl := least_spec (fix_fuel_P f x) - simp at Hl - match Hl with - | .inl Hl => simp [*] - | .inr ⟨ Hl, Hn ⟩ => - match Classical.em (fix_fuel n f x = div) with - | .inl Hd => - simp [*] - | .inr Hd => - have Hineq : least (fix_fuel_P f x) ≤ n := by - -- Proof by contradiction - cases Classical.em (least (fix_fuel_P f x) ≤ n) <;> simp [*] - simp at * - rename_i Hineq - have Hn := Hn n Hineq - contradiction - have Hfix : ¬ (fix f x = div) := by - simp [fix] - -- By property of the least upper bound - revert Hd Hl - -- TODO: there is no conversion to select the head of a function! - have : fix_fuel_P f x (least (fix_fuel_P f x)) = fix_fuel_pred f x (least (fix_fuel_P f x)) := - by simp[fix_fuel_P] - simp [this, div?] - clear this - cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp - have Hmono := fix_fuel_mono Hmono Hineq x - simp [result_rel] at Hmono - -- TODO: there is no conversion to select the head of a function! - revert Hmono Hfix Hd - simp [fix] - -- TODO: it would be good if cases actually introduces an equation: this - -- way we wouldn't have to do all the book-keeping - cases fix_fuel (least (fix_fuel_P f x)) f x <;> cases fix_fuel n f x <;> - intros <;> simp [*] at * - -theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : - ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by - intros x n Hf - have Hfmono := fix_fuel_fix_mono Hmono n x - revert Hf Hfmono - -- TODO: would be good to be able to unfold fix_fuel_P only on the left - simp [fix_fuel_P, div?, result_rel, fix] - cases fix_fuel n f x <;> simp_all - --- Prove the fixed point equation in the case there exists some fuel for which --- the execution terminates -theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono : is_mono f) - (x : a) (n : Nat) (He : fix_fuel_P f x n) : - fix f x = f (fix f) x := by - have Hl := fix_fuel_P_least Hmono He - -- TODO: better control of simplification - conv at Hl => - apply congr_fun - simp [fix_fuel_P] - -- The least upper bound is > 0 - have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by + have Hl := least_spec (fix_fuel_P f x) + simp at Hl + match Hl with + | .inl Hl => simp [*] + | .inr ⟨ Hl, Hn ⟩ => + match Classical.em (fix_fuel n f x = div) with + | .inl Hd => + simp [*] + | .inr Hd => + have Hineq : least (fix_fuel_P f x) ≤ n := by + -- Proof by contradiction + cases Classical.em (least (fix_fuel_P f x) ≤ n) <;> simp [*] + simp at * + rename_i Hineq + have Hn := Hn n Hineq + contradiction + have Hfix : ¬ (fix f x = div) := by + simp [fix] + -- By property of the least upper bound + revert Hd Hl + -- TODO: there is no conversion to select the head of a function! + conv => lhs; apply congr_fun; apply congr_fun; apply congr_fun; simp [fix_fuel_P, div?] + cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp + have Hmono := fix_fuel_mono Hmono Hineq x + simp [result_rel] at Hmono + simp [fix] at * + cases Heq: fix_fuel (least (fix_fuel_P f x)) f x <;> + cases Heq':fix_fuel n f x <;> + simp_all + + theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by + intros x n Hf + have Hfmono := fix_fuel_fix_mono Hmono n x + -- TODO: there is no conversion to select the head of a function! + conv => apply congr_fun; simp [fix_fuel_P] + simp [fix_fuel_P] at Hf + revert Hf Hfmono + simp [div?, result_rel, fix] + cases fix_fuel n f x <;> simp_all + + -- Prove the fixed point equation in the case there exists some fuel for which + -- the execution terminates + theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono : is_mono f) + (x : a) (n : Nat) (He : fix_fuel_P f x n) : + fix f x = f (fix f) x := by + have Hl := fix_fuel_P_least Hmono He + -- TODO: better control of simplification + conv at Hl => + apply congr_fun + simp [fix_fuel_P] + -- The least upper bound is > 0 + have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by + revert Hl + simp [div?] + cases least (fix_fuel_P f x) <;> simp [fix_fuel] + simp [Hsucc] at Hl revert Hl - simp [div?] - cases least (fix_fuel_P f x) <;> simp [fix_fuel] - simp [Hsucc] at Hl - revert Hl - simp [*, div?, fix, fix_fuel] - -- Use the monotonicity - have Hfixmono := fix_fuel_fix_mono Hmono n - have Hvm := Hmono Hfixmono x - -- Use functional extensionality - simp [result_rel, fix] at Hvm - revert Hvm - split <;> simp [*] <;> intros <;> simp [*] - --- The final fixed point equation --- TODO: remove the `forall x` -theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_valid f) : - ∀ x, fix f x = f (fix f) x := by - intros x - -- conv => lhs; simp [fix] - -- Case disjunction: is there a fuel such that the execution successfully execute? - match Classical.em (∃ n, fix_fuel_P f x n) with - | .inr He => - -- No fuel: the fixed point evaluates to `div` - --simp [fix] at * - simp at * - conv => lhs; simp [fix] - have Hel := He (Nat.succ (least (fix_fuel_P f x))); simp [*, fix_fuel] at *; clear Hel - -- Use the "continuity" of `f` - have He : ∀ n, fix_fuel (.succ n) f x = div := by intros; simp [*] - have Hcont := Hvalid.hcont x He - simp [Hcont] - | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hvalid.hmono x n He - - /- Making the proofs more systematic -/ - - -- TODO: rewrite is_mono in terms of is_mono_p - def is_mono_p (body : (a → Result b) → Result c) : Prop := - ∀ {{g h}}, marrow_rel g h → result_rel (body g) (body h) - - @[simp] theorem is_mono_p_same (x : Result c) : + simp [*, div?, fix, fix_fuel] + -- Use the monotonicity + have Hfixmono := fix_fuel_fix_mono Hmono n + have Hvm := Hmono Hfixmono x + -- Use functional extensionality + simp [result_rel, fix] at Hvm + revert Hvm + split <;> simp [*] <;> intros <;> simp [*] + + -- The final fixed point equation + -- TODO: remove the `forall x` + theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} + (Hmono : is_mono f) (Hcont : is_cont f) : + ∀ x, fix f x = f (fix f) x := by + intros x + -- Case disjunction: is there a fuel such that the execution successfully execute? + match Classical.em (∃ n, fix_fuel_P f x n) with + | .inr He => + -- No fuel: the fixed point evaluates to `div` + --simp [fix] at * + simp at * + conv => lhs; simp [fix] + have Hel := He (Nat.succ (least (fix_fuel_P f x))); simp [*, fix_fuel] at *; clear Hel + -- Use the "continuity" of `f` + have He : ∀ n, fix_fuel (.succ n) f x = div := by intros; simp [*] + have Hcont := Hcont x He + simp [Hcont] + | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hmono x n He + + + /-! # Making the proofs of validity manageable (and automatable) -/ + + -- Monotonicity property for expressions + def is_mono_p (e : (a → Result b) → Result c) : Prop := + ∀ {{k1 k2}}, karrow_rel k1 k2 → result_rel (e k1) (e k2) + + theorem is_mono_p_same (x : Result c) : @is_mono_p a b c (λ _ => x) := by - simp [is_mono_p, marrow_rel, result_rel] + simp [is_mono_p, karrow_rel, result_rel] split <;> simp - -- TODO: remove - @[simp] theorem is_mono_p_tail_rec (x : a) : + theorem is_mono_p_rec (x : a) : @is_mono_p a b b (λ f => f x) := by - simp_all [is_mono_p, marrow_rel, result_rel] + simp_all [is_mono_p, karrow_rel, result_rel] - -- TODO: rewrite is_cont in terms of is_cont_p - def is_cont_p (f : (a → Result b) → a → Result b) - (body : (a → Result b) → Result c) : Prop := - (Hc : ∀ n, body (fix_fuel n f) = .div) → - body (fix f) = .div - - @[simp] theorem is_cont_p_same (f : (a → Result b) → a → Result b) (x : Result c) : - is_cont_p f (λ _ => x) := by - simp [is_cont_p] - - -- TODO: remove - @[simp] theorem is_cont_p_tail_rec (f : (a → Result b) → a → Result b) (x : a) : - is_cont_p f (λ f => f x) := by - simp_all [is_cont_p, fix] - - -- Lean is good at unification: we can write a very general version + -- The important lemma about `is_mono_p` theorem is_mono_p_bind (g : (a → Result b) → Result c) (h : c → (a → Result b) → Result d) : is_mono_p g → (∀ y, is_mono_p (h y)) → - is_mono_p (λ f => do let y ← g f; h y f) := by + is_mono_p (λ k => do let y ← g k; h y k) := by intro hg hh simp [is_mono_p] intro fg fh Hrgh - simp [marrow_rel, result_rel] + simp [karrow_rel, result_rel] have hg := hg Hrgh; simp [result_rel] at hg cases heq0: g fg <;> simp_all rename_i y _ have hh := hh y Hrgh; simp [result_rel] at hh simp_all - -- Lean is good at unification: we can write a very general version - -- (in particular, it will manage to figure out `g` and `h` when we - -- apply the lemma) + -- Continuity property for expressions - note that we take the continuation + -- as parameter + def is_cont_p (k : (a → Result b) → a → Result b) + (e : (a → Result b) → Result c) : Prop := + (Hc : ∀ n, e (fix_fuel n k) = .div) → + e (fix k) = .div + + theorem is_cont_p_same (k : (a → Result b) → a → Result b) (x : Result c) : + is_cont_p k (λ _ => x) := by + simp [is_cont_p] + + theorem is_cont_p_rec (f : (a → Result b) → a → Result b) (x : a) : + is_cont_p f (λ f => f x) := by + simp_all [is_cont_p, fix] + + -- The important lemma about `is_cont_p` theorem is_cont_p_bind - (f : (a → Result b) → a → Result b) - (Hfmono : is_mono f) + (k : (a → Result b) → a → Result b) + (Hkmono : is_mono k) (g : (a → Result b) → Result c) (h : c → (a → Result b) → Result d) : is_mono_p g → - is_cont_p f g → + is_cont_p k g → (∀ y, is_mono_p (h y)) → - (∀ y, is_cont_p f (h y)) → - is_cont_p f (λ f => do let y ← g f; h y f) := by + (∀ y, is_cont_p k (h y)) → + is_cont_p k (λ k => do let y ← g k; h y k) := by intro Hgmono Hgcont Hhmono Hhcont simp [is_cont_p] intro Hdiv - -- Case on `g (fix... f)`: is there an n s.t. it terminates? - cases Classical.em (∀ n, g (fix_fuel n f) = .div) <;> rename_i Hn + -- Case on `g (fix... k)`: is there an n s.t. it terminates? + cases Classical.em (∀ n, g (fix_fuel n k) = .div) <;> rename_i Hn . -- Case 1: g diverges have Hgcont := Hgcont Hn simp_all @@ -384,20 +362,20 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va simp at Hn let ⟨ n, Hn ⟩ := Hn have Hdivn := Hdiv n - have Hffmono := fix_fuel_fix_mono Hfmono n + have Hffmono := fix_fuel_fix_mono Hkmono n have Hgeq := Hgmono Hffmono simp [result_rel] at Hgeq - cases Heq: g (fix_fuel n f) <;> rename_i y <;> simp_all + cases Heq: g (fix_fuel n k) <;> rename_i y <;> simp_all -- Remains the .ret case -- Use Hdiv to prove that: ∀ n, h y (fix_fuel n f) = div -- We do this in two steps: first we prove it for m ≥ n - have Hhdiv: ∀ m, h y (fix_fuel m f) = .div := by - have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m f) = .div := by + have Hhdiv: ∀ m, h y (fix_fuel m k) = .div := by + have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m k) = .div := by -- We use the fact that `g (fix_fuel n f) = .div`, combined with Hdiv intro m Hle have Hdivm := Hdiv m -- Monotonicity of g - have Hffmono := fix_fuel_mono Hfmono Hle + have Hffmono := fix_fuel_mono Hkmono Hle have Hgmono := Hgmono Hffmono -- We need to clear Hdiv because otherwise simp_all rewrites Hdivm with Hdiv clear Hdiv @@ -407,42 +385,41 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va cases Classical.em (n ≤ m) <;> rename_i Hl . apply Hhdiv; assumption . simp at Hl - -- Make a case disjunction on `h y (fix_fuel m f)`: if it is not equal + -- Make a case disjunction on `h y (fix_fuel m k)`: if it is not equal -- to div, use the monotonicity of `h y` have Hle : m ≤ n := by linarith - have Hffmono := fix_fuel_mono Hfmono Hle + have Hffmono := fix_fuel_mono Hkmono Hle have Hmono := Hhmono y Hffmono simp [result_rel] at Hmono - cases Heq: h y (fix_fuel m f) <;> simp_all + cases Heq: h y (fix_fuel m k) <;> simp_all -- We can now use the continuity hypothesis for h apply Hhcont; assumption - -- TODO: move + -- The validity property for an expression def is_valid_p (k : (a → Result b) → a → Result b) - (body : (a → Result b) → Result c) : Prop := - is_mono_p body ∧ - (is_mono k → is_cont_p k body) + (e : (a → Result b) → Result c) : Prop := + is_mono_p e ∧ + (is_mono k → is_cont_p k e) - @[simp] theorem is_valid_p_same (f : (a → Result b) → a → Result b) (x : Result c) : - is_valid_p f (λ _ => x) := by - simp [is_valid_p] + @[simp] theorem is_valid_p_same (k : (a → Result b) → a → Result b) (x : Result c) : + is_valid_p k (λ _ => x) := by + simp [is_valid_p, is_mono_p_same, is_cont_p_same] - @[simp] theorem is_valid_p_rec (f : (a → Result b) → a → Result b) (x : a) : - is_valid_p f (λ f => f x) := by - simp [is_valid_p] + @[simp] theorem is_valid_p_rec (k : (a → Result b) → a → Result b) (x : a) : + is_valid_p k (λ k => k x) := by + simp_all [is_valid_p, is_mono_p_rec, is_cont_p_rec] -- Lean is good at unification: we can write a very general version -- (in particular, it will manage to figure out `g` and `h` when we -- apply the lemma) theorem is_valid_p_bind - {{f : (a → Result b) → a → Result b}} + {{k : (a → Result b) → a → Result b}} {{g : (a → Result b) → Result c}} {{h : c → (a → Result b) → Result d}} - (Hgvalid : is_valid_p f g) - (Hhvalid : ∀ y, is_valid_p f (h y)) : - is_valid_p f (λ f => do let y ← g f; h y f) := by + (Hgvalid : is_valid_p k g) + (Hhvalid : ∀ y, is_valid_p k (h y)) : + is_valid_p k (λ k => do let y ← g k; h y k) := by let ⟨ Hgmono, Hgcont ⟩ := Hgvalid - -- TODO: conversion to move forall below and conjunction? simp [is_valid_p, forall_and] at Hhvalid have ⟨ Hhmono, Hhcont ⟩ := Hhvalid simp [← imp_forall_iff] at Hhcont @@ -450,36 +427,37 @@ theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} (Hvalid : is_va . -- Monotonicity apply is_mono_p_bind <;> assumption . -- Continuity - intro Hfmono - have Hgcont := Hgcont Hfmono - have Hhcont := Hhcont Hfmono + intro Hkmono + have Hgcont := Hgcont Hkmono + have Hhcont := Hhcont Hkmono apply is_cont_p_bind <;> assumption - theorem is_valid_p_imp_is_valid {{body : (a → Result b) → a → Result b}} - (Hvalid : ∀ f x, is_valid_p f (λ f => body f x)) : - is_valid body := by - have Hmono : is_mono body := by + theorem is_valid_p_imp_is_valid {{e : (a → Result b) → a → Result b}} + (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : + is_mono e ∧ is_cont e := by + have Hmono : is_mono e := by intro f h Hr x have Hmono := Hvalid (λ _ _ => .div) x have Hmono := Hmono.left apply Hmono; assumption - have Hcont : is_cont body := by + have Hcont : is_cont e := by intro x Hdiv - have Hcont := (Hvalid body x).right Hmono + have Hcont := (Hvalid e x).right Hmono simp [is_cont_p] at Hcont apply Hcont intro n have Hdiv := Hdiv n simp [fix_fuel] at Hdiv simp [*] - apply is_valid.intro Hmono Hcont + simp [*] -- TODO: functional extensionality - theorem is_valid_p_fix_fixed_eq {{body : (a → Result b) → a → Result b}} - (Hvalid : ∀ f x, is_valid_p f (λ f => body f x)) : - fix body = body (fix body) := by + theorem is_valid_p_fix_fixed_eq {{e : (a → Result b) → a → Result b}} + (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : + fix e = e (fix e) := by apply funext - exact fix_fixed_eq (is_valid_p_imp_is_valid Hvalid) + have ⟨ Hmono, Hcont ⟩ := is_valid_p_imp_is_valid Hvalid + exact fix_fixed_eq Hmono Hcont end Fix @@ -487,7 +465,7 @@ namespace Ex1 /- An example of use of the fixed-point -/ open Primitives Fix - variable {a : Type} (f : (List a × Int) → Result a) + variable {a : Type} (k : (List a × Int) → Result a) def list_nth_body (x : (List a × Int)) : Result a := let (ls, i) := x @@ -495,9 +473,9 @@ namespace Ex1 | [] => .fail .panic | hd :: tl => if i = 0 then .ret hd - else f (tl, i - 1) + else k (tl, i - 1) - theorem list_nth_body_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by + theorem list_nth_body_is_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by intro k x simp [list_nth_body] split <;> simp @@ -506,6 +484,7 @@ namespace Ex1 noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) + -- The unfolding equation theorem list_nth_eq (ls : List a) (i : Int) : list_nth ls i = match ls with @@ -514,17 +493,18 @@ namespace Ex1 if i = 0 then .ret hd else list_nth tl (i - 1) := by - have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_valid a) + have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_is_valid a) simp [list_nth] conv => lhs; rw [Heq] end Ex1 namespace Ex2 - /- Same as Ex1, but we make the body of nth non tail-rec -/ + /- Same as Ex1, but we make the body of nth non tail-rec (this is mostly + to see what happens when there are let-bindings) -/ open Primitives Fix - variable {a : Type} (f : (List a × Int) → Result a) + variable {a : Type} (k : (List a × Int) → Result a) def list_nth_body (x : (List a × Int)) : Result a := let (ls, i) := x @@ -534,10 +514,10 @@ namespace Ex2 if i = 0 then .ret hd else do - let y ← f (tl, i - 1) + let y ← k (tl, i - 1) .ret y - theorem list_nth_body_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by + theorem list_nth_body_is_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by intro k x simp [list_nth_body] split <;> simp @@ -547,6 +527,7 @@ namespace Ex2 noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) + -- The unfolding equation theorem list_nth_eq (ls : List a) (i : Int) : (list_nth ls i = match ls with @@ -558,7 +539,7 @@ namespace Ex2 let y ← list_nth tl (i - 1) .ret y) := by - have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_valid a) + have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_is_valid a) simp [list_nth] conv => lhs; rw [Heq] @@ -577,7 +558,7 @@ namespace Ex3 the functions in the mutually recursive group may not have the same return type. -/ - variable (f : (Int ⊕ Int) → Result (Bool ⊕ Bool)) + variable (k : (Int ⊕ Int) → Result (Bool ⊕ Bool)) def is_even_is_odd_body (x : (Int ⊕ Int)) : Result (Bool ⊕ Bool) := match x with @@ -591,7 +572,7 @@ namespace Ex3 do -- Call `odd`: we need to wrap the input value in `.inr`, then -- extract the output value - let r ← f (.inr (i- 1)) + let r ← k (.inr (i- 1)) match r with | .inl _ => .fail .panic -- Invalid output | .inr b => .ret b @@ -607,7 +588,7 @@ namespace Ex3 do -- Call `is_even`: we need to wrap the input value in .inr, then -- extract the output value - let r ← f (.inl (i- 1)) + let r ← k (.inl (i- 1)) match r with | .inl b => .ret b | .inr _ => .fail .panic -- Invalid output @@ -642,7 +623,8 @@ namespace Ex3 -- TODO: move -- TODO: this is not enough - theorem swap_if_bind {a b : Type} (e : Prop) [Decidable e] (x y : Result a) (f : a → Result b) : + theorem swap_if_bind {a b : Type} (e : Prop) [Decidable e] + (x y : Result a) (f : a → Result b) : (do let z ← (if e then x else y) f z) @@ -651,6 +633,7 @@ namespace Ex3 else do let z ← y; f z) := by split <;> simp + -- The unfolding equation for `is_even` theorem is_even_eq (i : Int) : is_even i = (if i = 0 then .ret true else is_odd (i - 1)) := by @@ -668,6 +651,7 @@ namespace Ex3 rename_i v split <;> simp + -- The unfolding equation for `is_odd` theorem is_odd_eq (i : Int) : is_odd i = (if i = 0 then .ret false else is_even (i - 1)) := by @@ -699,7 +683,6 @@ namespace Ex4 .ret (hd :: tl) /- The validity theorem for `map`, generic in `f` -/ - /- TODO: rename the continuation to k in all the lemma statements -/ theorem map_is_valid {{f : (a → Result b) → a → Result c}} (Hfvalid : ∀ k x, is_valid_p k (λ k => f k x)) @@ -724,7 +707,6 @@ namespace Ex4 let tl ← map f tl .ret (.node tl) - /- TODO: make the naming consistent (suffix with "_is") -/ theorem id_body_is_valid : ∀ k x, is_valid_p k (λ k => @id_body a k x) := by intro k x @@ -736,6 +718,7 @@ namespace Ex4 noncomputable def id (t : Tree a) := fix id_body t + -- The unfolding equation theorem id_eq (t : Tree a) : (id t = match t with -- cgit v1.2.3 From 393748cc3dd0f43a79d2342379008bbf445f116d Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 20 Jun 2023 12:30:39 +0200 Subject: Remove the use of fun. ext. in Diverge.lean --- backends/lean/Base/Diverge.lean | 34 +++++++++++++++++++--------------- 1 file changed, 19 insertions(+), 15 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 65c061bd..97ffa214 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -271,9 +271,7 @@ namespace Fix revert Hvm split <;> simp [*] <;> intros <;> simp [*] - -- The final fixed point equation - -- TODO: remove the `forall x` - theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} + theorem fix_fixed_eq_forall {{f : (a → Result b) → a → Result b}} (Hmono : is_mono f) (Hcont : is_cont f) : ∀ x, fix f x = f (fix f) x := by intros x @@ -291,6 +289,14 @@ namespace Fix simp [Hcont] | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hmono x n He + -- The final fixed point equation + theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} + (Hmono : is_mono f) (Hcont : is_cont f) : + fix f = f (fix f) := by + have Heq := fix_fixed_eq_forall Hmono Hcont + have Heq1 : fix f = (λ x => fix f x) := by simp + rw [Heq1] + conv => lhs; ext; simp [Heq] /-! # Making the proofs of validity manageable (and automatable) -/ @@ -451,11 +457,9 @@ namespace Fix simp [*] simp [*] - -- TODO: functional extensionality theorem is_valid_p_fix_fixed_eq {{e : (a → Result b) → a → Result b}} (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : fix e = e (fix e) := by - apply funext have ⟨ Hmono, Hcont ⟩ := is_valid_p_imp_is_valid Hvalid exact fix_fixed_eq Hmono Hcont @@ -484,7 +488,7 @@ namespace Ex1 noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) - -- The unfolding equation + -- The unfolding equation - diverges if `i < 0` theorem list_nth_eq (ls : List a) (i : Int) : list_nth ls i = match ls with @@ -527,7 +531,7 @@ namespace Ex2 noncomputable def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) - -- The unfolding equation + -- The unfolding equation - diverges if `i < 0` theorem list_nth_eq (ls : List a) (i : Int) : (list_nth ls i = match ls with @@ -553,10 +557,10 @@ namespace Ex3 and the output types: - inputs: the sum allows to select the function to call in the recursive calls (and the functions may not have the same input types) - - outpus: this case is degenerate because `even` and `odd` both have the + - outputs: this case is degenerate because `even` and `odd` have the same return type `Bool`, but generally speaking we need a sum type because - the functions in the mutually recursive group may not have the same - return type. + the functions in the mutually recursive group may have different + return types. -/ variable (k : (Int ⊕ Int) → Result (Bool ⊕ Bool)) @@ -565,7 +569,7 @@ namespace Ex3 | .inl i => -- Body of `is_even` if i = 0 - then .ret (.inl true) -- We return .inl because this is `is_even` + then .ret (.inl true) -- We use .inl because this is `is_even` else do let b ← @@ -581,7 +585,7 @@ namespace Ex3 | .inr i => -- Body of `is_odd` if i = 0 - then .ret (.inr false) -- We return .inr because this is `is_odd` + then .ret (.inr false) -- We use .inr because this is `is_odd` else do let b ← @@ -633,7 +637,7 @@ namespace Ex3 else do let z ← y; f z) := by split <;> simp - -- The unfolding equation for `is_even` + -- The unfolding equation for `is_even` - diverges if `i < 0` theorem is_even_eq (i : Int) : is_even i = (if i = 0 then .ret true else is_odd (i - 1)) := by @@ -651,9 +655,9 @@ namespace Ex3 rename_i v split <;> simp - -- The unfolding equation for `is_odd` + -- The unfolding equation for `is_odd` - diverges if `i < 0` theorem is_odd_eq (i : Int) : - is_odd i = (if i = 0 then .ret false else is_even (i - 1)) + is_odd i = (if i = 0 then .ret false else is_even (i - 1)) := by have Heq := is_valid_p_fix_fixed_eq is_even_is_odd_body_is_valid simp [is_even, is_odd] -- cgit v1.2.3 From 3971da603ee54d373b4c73d6a20b3d83dea7b5b9 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 21 Jun 2023 16:20:25 +0200 Subject: Start working on Arith.lean --- backends/lean/Base/Diverge.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 97ffa214..1ff34516 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -703,12 +703,12 @@ namespace Ex4 | leaf (x : a) | node (tl : List (Tree a)) - def id_body (f : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := + def id_body (k : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := match t with | .leaf x => .ret (.leaf x) | .node tl => do - let tl ← map f tl + let tl ← map k tl .ret (.node tl) theorem id_body_is_valid : -- cgit v1.2.3 From ffdc2f47bc4b21df491e1a2efb6cd19637fb399b Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 26 Jun 2023 17:38:49 +0200 Subject: Start working on a better encoding of mut rec defs for Diverge --- backends/lean/Base/Diverge.lean | 102 +++++++++++++++++++++++++++++++++++++++- 1 file changed, 100 insertions(+), 2 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 1ff34516..a5cf3459 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -550,7 +550,7 @@ namespace Ex2 end Ex2 namespace Ex3 - /- Mutually recursive functions -/ + /- Mutually recursive functions - first encoding (see Ex4 for a better encoding) -/ open Primitives Fix /- Because we have mutually recursive functions, we use a sum for the inputs @@ -671,6 +671,104 @@ namespace Ex3 end Ex3 namespace Ex4 + /- Mutually recursive functions - 2nd encoding -/ + open Primitives Fix + + attribute [local simp] List.get + + /- We make the input type and output types dependent on a parameter -/ + @[simp] def input_ty (i : Fin 2) : Type := + [Int, Int].get i + + @[simp] def output_ty (i : Fin 2) : Type := + [Bool, Bool].get i + + /- The continuation -/ + variable (k : (i : Fin 2) → input_ty i → Result (output_ty i)) + + /- The bodies are more natural -/ + def is_even_body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i : Int) : Result Bool := + if i = 0 + then .ret true + else do + let b ← k 1 (i - 1) + .ret b + + def is_odd_body (i : Int) : Result Bool := + if i = 0 + then .ret false + else do + let b ← k 0 (i - 1) + .ret b + + inductive Funs : List (Type 0) → List (Type 0) → Type 1 := + | Nil : Funs [] [] + | Cons {ity oty : Type 0} {itys otys : List (Type 0)} (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) + + theorem Funs.length_eq {itys otys : List (Type 0)} (fl : Funs itys otys) : itys.length = otys.length := + match fl with + | .Nil => by simp + | .Cons f tl => + have h:= Funs.length_eq tl + by simp [h] + + @[simp] def Funs.cast_fin {itys otys : List (Type 0)} (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := + ⟨ i.val, by have h:= fl.length_eq; have h1:= i.isLt; simp_all ⟩ + + @[simp] def bodies (k : (i : Fin 2) → input_ty i → Result (output_ty i)) : Funs [Int, Int] [Bool, Bool] := + Funs.Cons (is_even_body k) (Funs.Cons (is_odd_body k) Funs.Nil) + + @[simp] def get_fun {itys otys : List (Type 0)} (fl : Funs itys otys) : + (i : Fin itys.length) → itys.get i → Result (otys.get (fl.cast_fin i)) := + match fl with + | .Nil => λ i => by have h:= i.isLt; simp at h + | @Funs.Cons ity oty itys1 otys1 f tl => + λ i => + if h: i.val = 0 then + Eq.mp (by cases i; simp_all [List.get]) f + else + let j := i.val - 1 + have Hj: j < itys1.length := by + have Hi := i.isLt + simp at Hi + revert Hi + cases Heq: i.val <;> simp_all + simp_arith + let j: Fin itys1.length := ⟨ j, Hj ⟩ + Eq.mp (by cases Heq: i; rename_i val isLt; cases Heq': j; rename_i val' isLt; cases val <;> simp_all [List.get]) (get_fun tl j) + + def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : input_ty i → Result (output_ty i) := get_fun (bodies k) i + + def fix_ {n : Nat} {ity oty : Fin n → Type 0} (f : ((i:Fin n) → ity i → Result (oty i)) → (i:Fin n) → ity i → Result (oty i)) : + (i:Fin n) → ity i → Result (oty i) := + sorry + + theorem body_fix_eq : fix_ body = body (fix_ body) := sorry + + def is_even (i : Int) : Result Bool := fix_ body 0 i + def is_odd (i : Int) : Result Bool := fix_ body 1 i + + theorem is_even_eq (i : Int) : is_even i = + (if i = 0 + then .ret true + else do + let b ← is_odd (i - 1) + .ret b) := by + simp [is_even, is_odd]; + conv => lhs; rw [body_fix_eq] + + theorem is_odd_eq (i : Int) : is_odd i = + (if i = 0 + then .ret false + else do + let b ← is_even (i - 1) + .ret b) := by + simp [is_even, is_odd]; + conv => lhs; rw [body_fix_eq] + +end Ex4 + +namespace Ex5 /- Higher-order example -/ open Primitives Fix @@ -736,6 +834,6 @@ namespace Ex4 simp [id] conv => lhs; rw [Heq]; simp; rw [id_body] -end Ex4 +end Ex5 end Diverge -- cgit v1.2.3 From 87fd14e74fe00752df7759372093543ae77a51ae Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 26 Jun 2023 18:33:26 +0200 Subject: Make the definitions in Diverge.Fix dependently typed --- backends/lean/Base/Diverge.lean | 95 ++++++++++++++++++++--------------------- 1 file changed, 47 insertions(+), 48 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index a5cf3459..d65e77a1 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -88,7 +88,8 @@ namespace Fix open Primitives open Result - variable {a b c d : Type} + variable {a : Type} {b : a → Type} + variable {c d : Type} /-! # The least fixed point definition and its properties -/ @@ -132,19 +133,23 @@ namespace Fix /-! # The fixed point definitions -/ - def fix_fuel (n : Nat) (f : (a → Result b) → a → Result b) (x : a) : Result b := + def fix_fuel (n : Nat) (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : + Result (b x) := match n with | 0 => .div | n + 1 => f (fix_fuel n f) x - @[simp] def fix_fuel_pred (f : (a → Result b) → a → Result b) (x : a) (n : Nat) := + @[simp] def fix_fuel_pred (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) + (x : a) (n : Nat) := not (div? (fix_fuel n f x)) - def fix_fuel_P (f : (a → Result b) → a → Result b) (x : a) (n : Nat) : Prop := + def fix_fuel_P (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) + (x : a) (n : Nat) : Prop := fix_fuel_pred f x n - noncomputable def fix (f : (a → Result b) → a → Result b) (x : a) : Result b := + noncomputable + def fix (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : Result (b x) := fix_fuel (least (fix_fuel_P f x)) f x /-! # The validity property -/ @@ -158,11 +163,11 @@ namespace Fix | ret _ => x2 = x1 -- TODO: generalize -- Monotonicity relation over monadic arrows (i.e., Kleisli arrows) - def karrow_rel (k1 k2 : a → Result b) : Prop := + def karrow_rel (k1 k2 : (x:a) → Result (b x)) : Prop := ∀ x, result_rel (k1 x) (k2 x) -- Monotonicity property for function bodies - def is_mono (f : (a → Result b) → a → Result b) : Prop := + def is_mono (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := ∀ {{k1 k2}}, karrow_rel k1 k2 → karrow_rel (f k1) (f k2) -- "Continuity" property. @@ -170,11 +175,12 @@ namespace Fix -- https://inria.hal.science/file/index/docid/216187/filename/tarski.pdf -- We define our "continuity" criteria so that it gives us what we need to -- prove the fixed-point equation, and we can also easily manipulate it. - def is_cont (f : (a → Result b) → a → Result b) : Prop := + def is_cont (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := ∀ x, (Hdiv : ∀ n, fix_fuel (.succ n) f x = div) → f (fix f) x = div /-! # The proof of the fixed-point equation -/ - theorem fix_fuel_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + theorem fix_fuel_mono {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} + (Hmono : is_mono f) : ∀ {{n m}}, n ≤ m → karrow_rel (fix_fuel n f) (fix_fuel m f) := by intros n induction n @@ -195,12 +201,13 @@ namespace Fix simp [result_rel] at Hmono apply Hmono - @[simp] theorem neg_fix_fuel_P {f : (a → Result b) → a → Result b} {x : a} {n : Nat} : + @[simp] theorem neg_fix_fuel_P + {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} {x : a} {n : Nat} : ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by simp [fix_fuel_P, div?] cases fix_fuel n f x <;> simp - theorem fix_fuel_fix_mono {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + theorem fix_fuel_fix_mono {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} (Hmono : is_mono f) : ∀ n, karrow_rel (fix_fuel n f) (fix f) := by intros n x simp [result_rel] @@ -234,7 +241,7 @@ namespace Fix cases Heq':fix_fuel n f x <;> simp_all - theorem fix_fuel_P_least {f : (a → Result b) → a → Result b} (Hmono : is_mono f) : + theorem fix_fuel_P_least {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} (Hmono : is_mono f) : ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by intros x n Hf have Hfmono := fix_fuel_fix_mono Hmono n x @@ -247,7 +254,7 @@ namespace Fix -- Prove the fixed point equation in the case there exists some fuel for which -- the execution terminates - theorem fix_fixed_eq_terminates (f : (a → Result b) → a → Result b) (Hmono : is_mono f) + theorem fix_fixed_eq_terminates (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (Hmono : is_mono f) (x : a) (n : Nat) (He : fix_fuel_P f x n) : fix f x = f (fix f) x := by have Hl := fix_fuel_P_least Hmono He @@ -271,7 +278,7 @@ namespace Fix revert Hvm split <;> simp [*] <;> intros <;> simp [*] - theorem fix_fixed_eq_forall {{f : (a → Result b) → a → Result b}} + theorem fix_fixed_eq_forall {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} (Hmono : is_mono f) (Hcont : is_cont f) : ∀ x, fix f x = f (fix f) x := by intros x @@ -290,7 +297,7 @@ namespace Fix | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hmono x n He -- The final fixed point equation - theorem fix_fixed_eq {{f : (a → Result b) → a → Result b}} + theorem fix_fixed_eq {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} (Hmono : is_mono f) (Hcont : is_cont f) : fix f = f (fix f) := by have Heq := fix_fixed_eq_forall Hmono Hcont @@ -301,7 +308,7 @@ namespace Fix /-! # Making the proofs of validity manageable (and automatable) -/ -- Monotonicity property for expressions - def is_mono_p (e : (a → Result b) → Result c) : Prop := + def is_mono_p (e : ((x:a) → Result (b x)) → Result c) : Prop := ∀ {{k1 k2}}, karrow_rel k1 k2 → result_rel (e k1) (e k2) theorem is_mono_p_same (x : Result c) : @@ -310,16 +317,17 @@ namespace Fix split <;> simp theorem is_mono_p_rec (x : a) : - @is_mono_p a b b (λ f => f x) := by + @is_mono_p a b (b x) (λ f => f x) := by simp_all [is_mono_p, karrow_rel, result_rel] -- The important lemma about `is_mono_p` + -- TODO: generalize d? theorem is_mono_p_bind - (g : (a → Result b) → Result c) - (h : c → (a → Result b) → Result d) : + (g : ((x:a) → Result (b x)) → Result c) + (h : c → ((x:a) → Result (b x)) → Result d) : is_mono_p g → (∀ y, is_mono_p (h y)) → - is_mono_p (λ k => do let y ← g k; h y k) := by + @is_mono_p a b d (λ k => do let y ← g k; h y k) := by intro hg hh simp [is_mono_p] intro fg fh Hrgh @@ -332,25 +340,26 @@ namespace Fix -- Continuity property for expressions - note that we take the continuation -- as parameter - def is_cont_p (k : (a → Result b) → a → Result b) - (e : (a → Result b) → Result c) : Prop := + def is_cont_p (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) + (e : ((x:a) → Result (b x)) → Result c) : Prop := (Hc : ∀ n, e (fix_fuel n k) = .div) → e (fix k) = .div - theorem is_cont_p_same (k : (a → Result b) → a → Result b) (x : Result c) : + theorem is_cont_p_same (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) + (x : Result c) : is_cont_p k (λ _ => x) := by simp [is_cont_p] - theorem is_cont_p_rec (f : (a → Result b) → a → Result b) (x : a) : + theorem is_cont_p_rec (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : is_cont_p f (λ f => f x) := by simp_all [is_cont_p, fix] -- The important lemma about `is_cont_p` theorem is_cont_p_bind - (k : (a → Result b) → a → Result b) + (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (Hkmono : is_mono k) - (g : (a → Result b) → Result c) - (h : c → (a → Result b) → Result d) : + (g : ((x:a) → Result (b x)) → Result c) + (h : c → ((x:a) → Result (b x)) → Result d) : is_mono_p g → is_cont_p k g → (∀ y, is_mono_p (h y)) → @@ -402,16 +411,18 @@ namespace Fix apply Hhcont; assumption -- The validity property for an expression - def is_valid_p (k : (a → Result b) → a → Result b) - (e : (a → Result b) → Result c) : Prop := + def is_valid_p (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) + (e : ((x:a) → Result (b x)) → Result c) : Prop := is_mono_p e ∧ (is_mono k → is_cont_p k e) - @[simp] theorem is_valid_p_same (k : (a → Result b) → a → Result b) (x : Result c) : + @[simp] theorem is_valid_p_same + (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : Result c) : is_valid_p k (λ _ => x) := by simp [is_valid_p, is_mono_p_same, is_cont_p_same] - @[simp] theorem is_valid_p_rec (k : (a → Result b) → a → Result b) (x : a) : + @[simp] theorem is_valid_p_rec + (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : is_valid_p k (λ k => k x) := by simp_all [is_valid_p, is_mono_p_rec, is_cont_p_rec] @@ -419,9 +430,9 @@ namespace Fix -- (in particular, it will manage to figure out `g` and `h` when we -- apply the lemma) theorem is_valid_p_bind - {{k : (a → Result b) → a → Result b}} - {{g : (a → Result b) → Result c}} - {{h : c → (a → Result b) → Result d}} + {{k : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} + {{g : ((x:a) → Result (b x)) → Result c}} + {{h : c → ((x:a) → Result (b x)) → Result d}} (Hgvalid : is_valid_p k g) (Hhvalid : ∀ y, is_valid_p k (h y)) : is_valid_p k (λ k => do let y ← g k; h y k) := by @@ -438,7 +449,7 @@ namespace Fix have Hhcont := Hhcont Hkmono apply is_cont_p_bind <;> assumption - theorem is_valid_p_imp_is_valid {{e : (a → Result b) → a → Result b}} + theorem is_valid_p_imp_is_valid {{e : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : is_mono e ∧ is_cont e := by have Hmono : is_mono e := by @@ -457,7 +468,7 @@ namespace Fix simp [*] simp [*] - theorem is_valid_p_fix_fixed_eq {{e : (a → Result b) → a → Result b}} + theorem is_valid_p_fix_fixed_eq {{e : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : fix e = e (fix e) := by have ⟨ Hmono, Hcont ⟩ := is_valid_p_imp_is_valid Hvalid @@ -625,18 +636,6 @@ namespace Ex3 | .inl _ => .fail .panic | .inr b => .ret b - -- TODO: move - -- TODO: this is not enough - theorem swap_if_bind {a b : Type} (e : Prop) [Decidable e] - (x y : Result a) (f : a → Result b) : - (do - let z ← (if e then x else y) - f z) - = - (if e then do let z ← x; f z - else do let z ← y; f z) := by - split <;> simp - -- The unfolding equation for `is_even` - diverges if `i < 0` theorem is_even_eq (i : Int) : is_even i = (if i = 0 then .ret true else is_odd (i - 1)) -- cgit v1.2.3 From 4cc411a30b19f5c5eea67b2e4da232337af8f12b Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 26 Jun 2023 18:40:47 +0200 Subject: Generalize some definitions --- backends/lean/Base/Diverge.lean | 34 ++++++++++++++++++++++------------ 1 file changed, 22 insertions(+), 12 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index d65e77a1..0c1028bd 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -321,7 +321,6 @@ namespace Fix simp_all [is_mono_p, karrow_rel, result_rel] -- The important lemma about `is_mono_p` - -- TODO: generalize d? theorem is_mono_p_bind (g : ((x:a) → Result (b x)) → Result c) (h : c → ((x:a) → Result (b x)) → Result d) : @@ -700,24 +699,28 @@ namespace Ex4 let b ← k 0 (i - 1) .ret b - inductive Funs : List (Type 0) → List (Type 0) → Type 1 := + inductive Funs : List (Type u) → List (Type u) → Type (u + 1) := | Nil : Funs [] [] - | Cons {ity oty : Type 0} {itys otys : List (Type 0)} (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) + | Cons {ity oty : Type u} {itys otys : List (Type u)} + (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) - theorem Funs.length_eq {itys otys : List (Type 0)} (fl : Funs itys otys) : itys.length = otys.length := + theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs itys otys) : + itys.length = otys.length := match fl with | .Nil => by simp | .Cons f tl => have h:= Funs.length_eq tl by simp [h] - @[simp] def Funs.cast_fin {itys otys : List (Type 0)} (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := + @[simp] def Funs.cast_fin {itys otys : List (Type)} + (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := ⟨ i.val, by have h:= fl.length_eq; have h1:= i.isLt; simp_all ⟩ - @[simp] def bodies (k : (i : Fin 2) → input_ty i → Result (output_ty i)) : Funs [Int, Int] [Bool, Bool] := + @[simp] def bodies (k : (i : Fin 2) → input_ty i → Result (output_ty i)) : + Funs [Int, Int] [Bool, Bool] := Funs.Cons (is_even_body k) (Funs.Cons (is_odd_body k) Funs.Nil) - @[simp] def get_fun {itys otys : List (Type 0)} (fl : Funs itys otys) : + @[simp] def get_fun {itys otys : List (Type)} (fl : Funs itys otys) : (i : Fin itys.length) → itys.get i → Result (otys.get (fl.cast_fin i)) := match fl with | .Nil => λ i => by have h:= i.isLt; simp at h @@ -734,11 +737,18 @@ namespace Ex4 cases Heq: i.val <;> simp_all simp_arith let j: Fin itys1.length := ⟨ j, Hj ⟩ - Eq.mp (by cases Heq: i; rename_i val isLt; cases Heq': j; rename_i val' isLt; cases val <;> simp_all [List.get]) (get_fun tl j) - - def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : input_ty i → Result (output_ty i) := get_fun (bodies k) i - - def fix_ {n : Nat} {ity oty : Fin n → Type 0} (f : ((i:Fin n) → ity i → Result (oty i)) → (i:Fin n) → ity i → Result (oty i)) : + Eq.mp + (by + cases Heq: i; rename_i val isLt; + cases Heq': j; rename_i val' isLt; + cases val <;> simp_all [List.get]) + (get_fun tl j) + + def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : + input_ty i → Result (output_ty i) := get_fun (bodies k) i + + def fix_ {n : Nat} {ity oty : Fin n → Type} + (f : ((i:Fin n) → ity i → Result (oty i)) → (i:Fin n) → ity i → Result (oty i)) : (i:Fin n) → ity i → Result (oty i) := sorry -- cgit v1.2.3 From f4ee75da0959ff06ce4cfaab817de540fcd0433f Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 26 Jun 2023 19:28:03 +0200 Subject: Add FixI in Diverge --- backends/lean/Base/Diverge.lean | 127 +++++++++++++++++++++++++++++++++------- 1 file changed, 105 insertions(+), 22 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 0c1028bd..907075d7 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -448,17 +448,20 @@ namespace Fix have Hhcont := Hhcont Hkmono apply is_cont_p_bind <;> assumption - theorem is_valid_p_imp_is_valid {{e : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : - is_mono e ∧ is_cont e := by - have Hmono : is_mono e := by + def is_valid (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := + ∀ k x, is_valid_p k (λ k => f k x) + + theorem is_valid_p_imp_is_valid {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} + (Hvalid : is_valid f) : + is_mono f ∧ is_cont f := by + have Hmono : is_mono f := by intro f h Hr x have Hmono := Hvalid (λ _ _ => .div) x have Hmono := Hmono.left apply Hmono; assumption - have Hcont : is_cont e := by + have Hcont : is_cont f := by intro x Hdiv - have Hcont := (Hvalid e x).right Hmono + have Hcont := (Hvalid f x).right Hmono simp [is_cont_p] at Hcont apply Hcont intro n @@ -467,14 +470,96 @@ namespace Fix simp [*] simp [*] - theorem is_valid_p_fix_fixed_eq {{e : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hvalid : ∀ k x, is_valid_p k (λ k => e k x)) : - fix e = e (fix e) := by + theorem is_valid_fix_fixed_eq {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} + (Hvalid : is_valid f) : + fix f = f (fix f) := by have ⟨ Hmono, Hcont ⟩ := is_valid_p_imp_is_valid Hvalid exact fix_fixed_eq Hmono Hcont end Fix +namespace FixI + /- Indexed fixed-point: definitions with indexed types, convenient to use for mutually + recursive definitions. We simply port the definitions and proofs from Fix to a more + specific case. + -/ + open Primitives Fix + + -- The index type + variable {id : Type} + + -- The input/output types + variable {a b : id → Type} + + -- Monotonicity relation over monadic arrows (i.e., Kleisli arrows) + def karrow_rel (k1 k2 : (i:id) → a i → Result (b i)) : Prop := + ∀ i x, result_rel (k1 i x) (k2 i x) + + def kk_to_gen (k : (i:id) → a i → Result (b i)) : + (x: (i:id) × a i) → Result (b x.fst) := + λ ⟨ i, x ⟩ => k i x + + def kk_of_gen (k : (x: (i:id) × a i) → Result (b x.fst)) : + (i:id) → a i → Result (b i) := + λ i x => k ⟨ i, x ⟩ + + def k_to_gen (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : + ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst) := + λ kk => kk_to_gen (k (kk_of_gen kk)) + + def k_of_gen (k : ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst)) : + ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i) := + λ kk => kk_of_gen (k (kk_to_gen kk)) + + def e_to_gen (e : ((i:id) → a i → Result (b i)) → Result c) : + ((x: (i:id) × a i) → Result (b x.fst)) → Result c := + λ k => e (kk_of_gen k) + + def is_valid_p (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) + (e : ((i:id) → a i → Result (b i)) → Result c) : Prop := + Fix.is_valid_p (k_to_gen k) (e_to_gen e) + + def is_valid (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : Prop := + ∀ k i x, is_valid_p k (λ k => f k i x) + + @[simp] theorem kk_to_gen_kk_of_gen + (k : (x: (i:id) × a i) → Result (b x.fst)) : + kk_to_gen (kk_of_gen kk) = kk := by + simp [kk_to_gen, kk_of_gen] + + @[simp] theorem k_to_gen_k_of_gen + (k : ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst)) : + k_to_gen (k_of_gen kk) = kk := by + simp [k_to_gen, k_of_gen] + apply funext + intro kk_1 + -- TODO: some simplifications don't work + simp [kk_to_gen, kk_of_gen] + + noncomputable def fix + (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : + (i:id) → a i → Result (b i) := + kk_of_gen (Fix.fix (k_to_gen f)) + + theorem is_valid_fix_fixed_eq + {{f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)}} + (Hvalid : is_valid f) : + fix f = f (fix f) := by + have Hvalid' : Fix.is_valid (k_to_gen f) := by + intro k x + simp [is_valid, is_valid_p] at Hvalid + --simp [Fix.is_valid_p] + let ⟨ i, x ⟩ := x + have Hvalid := Hvalid (k_of_gen k) i x + -- TODO: some simplifications don't work + simp [k_to_gen, k_of_gen, kk_to_gen, kk_of_gen] at Hvalid + refine Hvalid + have Heq := Fix.is_valid_fix_fixed_eq Hvalid' + simp [fix] + conv => lhs; rw [Heq] + +end FixI + namespace Ex1 /- An example of use of the fixed-point -/ open Primitives Fix @@ -507,7 +592,7 @@ namespace Ex1 if i = 0 then .ret hd else list_nth tl (i - 1) := by - have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_is_valid a) + have Heq := is_valid_fix_fixed_eq (@list_nth_body_is_valid a) simp [list_nth] conv => lhs; rw [Heq] @@ -553,7 +638,7 @@ namespace Ex2 let y ← list_nth tl (i - 1) .ret y) := by - have Heq := is_valid_p_fix_fixed_eq (@list_nth_body_is_valid a) + have Heq := is_valid_fix_fixed_eq (@list_nth_body_is_valid a) simp [list_nth] conv => lhs; rw [Heq] @@ -639,7 +724,7 @@ namespace Ex3 theorem is_even_eq (i : Int) : is_even i = (if i = 0 then .ret true else is_odd (i - 1)) := by - have Heq := is_valid_p_fix_fixed_eq is_even_is_odd_body_is_valid + have Heq := is_valid_fix_fixed_eq is_even_is_odd_body_is_valid simp [is_even, is_odd] conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp -- Very annoying: we need to swap the matches @@ -657,7 +742,7 @@ namespace Ex3 theorem is_odd_eq (i : Int) : is_odd i = (if i = 0 then .ret false else is_even (i - 1)) := by - have Heq := is_valid_p_fix_fixed_eq is_even_is_odd_body_is_valid + have Heq := is_valid_fix_fixed_eq is_even_is_odd_body_is_valid simp [is_even, is_odd] conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp -- Same remark as for `even` @@ -670,7 +755,7 @@ end Ex3 namespace Ex4 /- Mutually recursive functions - 2nd encoding -/ - open Primitives Fix + open Primitives FixI attribute [local simp] List.get @@ -747,15 +832,13 @@ namespace Ex4 def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : input_ty i → Result (output_ty i) := get_fun (bodies k) i - def fix_ {n : Nat} {ity oty : Fin n → Type} - (f : ((i:Fin n) → ity i → Result (oty i)) → (i:Fin n) → ity i → Result (oty i)) : - (i:Fin n) → ity i → Result (oty i) := - sorry + theorem body_is_valid : is_valid body := by sorry - theorem body_fix_eq : fix_ body = body (fix_ body) := sorry + theorem body_fix_eq : fix body = body (fix body) := + is_valid_fix_fixed_eq body_is_valid - def is_even (i : Int) : Result Bool := fix_ body 0 i - def is_odd (i : Int) : Result Bool := fix_ body 1 i + noncomputable def is_even (i : Int) : Result Bool := fix body 0 i + noncomputable def is_odd (i : Int) : Result Bool := fix body 1 i theorem is_even_eq (i : Int) : is_even i = (if i = 0 @@ -839,7 +922,7 @@ namespace Ex5 let tl ← map id tl .ret (.node tl)) := by - have Heq := is_valid_p_fix_fixed_eq (@id_body_is_valid a) + have Heq := is_valid_fix_fixed_eq (@id_body_is_valid a) simp [id] conv => lhs; rw [Heq]; simp; rw [id_body] -- cgit v1.2.3 From 7f3604c21bb9f923aecb98917b5c7a33bafd1bcb Mon Sep 17 00:00:00 2001 From: Son Ho Date: Mon, 26 Jun 2023 23:41:01 +0200 Subject: Make minor modifications --- backends/lean/Base/Diverge.lean | 16 ---------------- 1 file changed, 16 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 907075d7..f3fa4815 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -522,20 +522,6 @@ namespace FixI def is_valid (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : Prop := ∀ k i x, is_valid_p k (λ k => f k i x) - @[simp] theorem kk_to_gen_kk_of_gen - (k : (x: (i:id) × a i) → Result (b x.fst)) : - kk_to_gen (kk_of_gen kk) = kk := by - simp [kk_to_gen, kk_of_gen] - - @[simp] theorem k_to_gen_k_of_gen - (k : ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst)) : - k_to_gen (k_of_gen kk) = kk := by - simp [k_to_gen, k_of_gen] - apply funext - intro kk_1 - -- TODO: some simplifications don't work - simp [kk_to_gen, kk_of_gen] - noncomputable def fix (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : (i:id) → a i → Result (b i) := @@ -548,10 +534,8 @@ namespace FixI have Hvalid' : Fix.is_valid (k_to_gen f) := by intro k x simp [is_valid, is_valid_p] at Hvalid - --simp [Fix.is_valid_p] let ⟨ i, x ⟩ := x have Hvalid := Hvalid (k_of_gen k) i x - -- TODO: some simplifications don't work simp [k_to_gen, k_of_gen, kk_to_gen, kk_of_gen] at Hvalid refine Hvalid have Heq := Fix.is_valid_fix_fixed_eq Hvalid' -- cgit v1.2.3 From 0a62cf3f7d58b31c75344172bad1032e14a4082f Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 27 Jun 2023 10:52:07 +0200 Subject: Finish the proofs which use FixI --- backends/lean/Base/Diverge.lean | 199 +++++++++++++++++++++++++++++++--------- 1 file changed, 157 insertions(+), 42 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index f3fa4815..76f0543a 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -542,6 +542,147 @@ namespace FixI simp [fix] conv => lhs; rw [Heq] + /- Some utilities to define the mutually recursive functions -/ + + inductive Funs : List (Type u) → List (Type u) → Type (u + 1) := + | Nil : Funs [] [] + | Cons {ity oty : Type u} {itys otys : List (Type u)} + (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) + + theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs itys otys) : + itys.length = otys.length := + match fl with + | .Nil => by simp + | .Cons f tl => + have h:= Funs.length_eq tl + by simp [h] + + @[simp] def Funs.cast_fin {itys otys : List (Type)} + (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := + ⟨ i.val, by have h:= fl.length_eq; have h1:= i.isLt; simp_all ⟩ + + def get_fun {itys otys : List (Type)} (fl : Funs itys otys) : + (i : Fin itys.length) → itys.get i → Result (otys.get (fl.cast_fin i)) := + match fl with + | .Nil => λ i => by have h:= i.isLt; simp at h + | @Funs.Cons ity oty itys1 otys1 f tl => + λ i => + if h: i.val = 0 then + Eq.mp (by cases i; simp_all [List.get]) f + else + let j := i.val - 1 + have Hj: j < itys1.length := by + have Hi := i.isLt + simp at Hi + revert Hi + cases Heq: i.val <;> simp_all + simp_arith + let j: Fin itys1.length := ⟨ j, Hj ⟩ + Eq.mp + (by + cases Heq: i; rename_i val isLt; + cases Heq': j; rename_i val' isLt; + cases val <;> simp_all [List.get]) + (get_fun tl j) + + + -- TODO: move + theorem add_one_le_iff_le_ne (n m : Nat) (h1 : m ≤ n) (h2 : m ≠ n) : m + 1 ≤ n := by + -- Damn, those proofs on natural numbers are hard - I wish Omega was in mathlib4... + simp [Nat.add_one_le_iff] + simp [Nat.lt_iff_le_and_ne] + simp_all + + def for_all_fin_aux {n : Nat} (f : Fin n → Prop) (m : Nat) (h : m ≤ n) : Prop := + if heq: m = n then True + else + f ⟨ m, by simp_all [Nat.lt_iff_le_and_ne] ⟩ ∧ + for_all_fin_aux f (m + 1) (by simp_all [add_one_le_iff_le_ne]) + termination_by for_all_fin_aux n _ m h => n - m + decreasing_by + simp_wf + apply Nat.sub_add_lt_sub <;> simp_all + simp_all [add_one_le_iff_le_ne] + + def for_all_fin {n : Nat} (f : Fin n → Prop) := for_all_fin_aux f 0 (by simp) + + theorem for_all_fin_aux_imp_forall {n : Nat} (f : Fin n → Prop) (m : Nat) : + (h : m ≤ n) → + for_all_fin_aux f m h → ∀ i, m ≤ i.val → f i + := by + generalize h: (n - m) = k + revert m + induction k + case zero => + simp_all + intro m h1 h2 + have h: n = m := by + linarith + unfold for_all_fin_aux; simp_all + simp_all + -- There is no i s.t. m ≤ i + intro i h3; cases i; simp_all + linarith + case succ k hi => + simp_all + intro m hk hmn + intro hf i hmi + have hne: m ≠ n := by + have hineq := Nat.lt_of_sub_eq_succ hk + linarith + -- m = i? + if heq: m = i then + -- Yes: simply use the `for_all_fin_aux` hyp + unfold for_all_fin_aux at hf + simp_all + tauto + else + -- No: use the induction hypothesis + have hlt: m < i := by simp_all [Nat.lt_iff_le_and_ne] + have hineq: m + 1 ≤ n := by + have hineq := Nat.lt_of_sub_eq_succ hk + simp [*, Nat.add_one_le_iff] + have heq1: n - (m + 1) = k := by + -- TODO: very annoying arithmetic proof + simp [Nat.sub_eq_iff_eq_add hineq] + have hineq1: m ≤ n := by linarith + simp [Nat.sub_eq_iff_eq_add hineq1] at hk + simp_arith [hk] + have hi := hi (m + 1) heq1 hineq + apply hi <;> simp_all + . unfold for_all_fin_aux at hf + simp_all + . simp_all [add_one_le_iff_le_ne] + + theorem for_all_fin_imp_forall (n : Nat) (f : Fin n → Prop) : + for_all_fin f → ∀ i, f i + := by + intro Hf i + apply for_all_fin_aux_imp_forall <;> try assumption + simp + + /- Automating the proofs -/ + @[simp] theorem is_valid_p_same + (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) (x : Result c) : + is_valid_p k (λ _ => x) := by + simp [is_valid_p, k_to_gen, e_to_gen] + + @[simp] theorem is_valid_p_rec + (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) (i : id) (x : a i) : + is_valid_p k (λ k => k i x) := by + simp [is_valid_p, k_to_gen, e_to_gen, kk_to_gen, kk_of_gen] + + theorem is_valid_p_bind + {{k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)}} + {{g : ((i:id) → a i → Result (b i)) → Result c}} + {{h : c → ((i:id) → a i → Result (b i)) → Result d}} + (Hgvalid : is_valid_p k g) + (Hhvalid : ∀ y, is_valid_p k (h y)) : + is_valid_p k (λ k => do let y ← g k; h y k) := by + apply Fix.is_valid_p_bind + . apply Hgvalid + . apply Hhvalid + end FixI namespace Ex1 @@ -768,55 +909,29 @@ namespace Ex4 let b ← k 0 (i - 1) .ret b - inductive Funs : List (Type u) → List (Type u) → Type (u + 1) := - | Nil : Funs [] [] - | Cons {ity oty : Type u} {itys otys : List (Type u)} - (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) - - theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs itys otys) : - itys.length = otys.length := - match fl with - | .Nil => by simp - | .Cons f tl => - have h:= Funs.length_eq tl - by simp [h] - - @[simp] def Funs.cast_fin {itys otys : List (Type)} - (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := - ⟨ i.val, by have h:= fl.length_eq; have h1:= i.isLt; simp_all ⟩ - @[simp] def bodies (k : (i : Fin 2) → input_ty i → Result (output_ty i)) : Funs [Int, Int] [Bool, Bool] := Funs.Cons (is_even_body k) (Funs.Cons (is_odd_body k) Funs.Nil) - @[simp] def get_fun {itys otys : List (Type)} (fl : Funs itys otys) : - (i : Fin itys.length) → itys.get i → Result (otys.get (fl.cast_fin i)) := - match fl with - | .Nil => λ i => by have h:= i.isLt; simp at h - | @Funs.Cons ity oty itys1 otys1 f tl => - λ i => - if h: i.val = 0 then - Eq.mp (by cases i; simp_all [List.get]) f - else - let j := i.val - 1 - have Hj: j < itys1.length := by - have Hi := i.isLt - simp at Hi - revert Hi - cases Heq: i.val <;> simp_all - simp_arith - let j: Fin itys1.length := ⟨ j, Hj ⟩ - Eq.mp - (by - cases Heq: i; rename_i val isLt; - cases Heq': j; rename_i val' isLt; - cases val <;> simp_all [List.get]) - (get_fun tl j) - def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : input_ty i → Result (output_ty i) := get_fun (bodies k) i - theorem body_is_valid : is_valid body := by sorry + theorem body_is_valid : is_valid body := by + -- Split the proof into proofs of validity of the individual bodies + rw [is_valid] + simp [body] + intro k + apply for_all_fin_imp_forall + simp [for_all_fin] + repeat (unfold for_all_fin_aux; simp) + simp [get_fun] + (repeat (apply And.intro)) <;> intro x <;> simp at x <;> + simp [is_even_body, is_odd_body] + -- Prove the validity of the individual bodies + . split <;> simp + apply is_valid_p_bind <;> simp + . split <;> simp + apply is_valid_p_bind <;> simp theorem body_fix_eq : fix body = body (fix body) := is_valid_fix_fixed_eq body_is_valid -- cgit v1.2.3 From 8db1af5afcb414b502a58a87f6bdcc1c08cbe3d2 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 27 Jun 2023 18:22:22 +0200 Subject: Finish some proofs in Diverge --- backends/lean/Base/Diverge.lean | 119 ++++++++++++++++++++++++++++++++++------ 1 file changed, 102 insertions(+), 17 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index 76f0543a..c97674dd 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -544,28 +544,40 @@ namespace FixI /- Some utilities to define the mutually recursive functions -/ - inductive Funs : List (Type u) → List (Type u) → Type (u + 1) := - | Nil : Funs [] [] + -- TODO: use more + @[simp] def kk_ty (id : Type) (a b : id → Type) := (i:id) → a i → Result (b i) + @[simp] def k_ty (id : Type) (a b : id → Type) := kk_ty id a b → kk_ty id a b + + -- Initially, we had left out the parameters id, a and b. + -- However, by parameterizing Funs with those parameters, we can state + -- and prove lemmas like Funs.is_valid_p_is_valid_p + inductive Funs (id : Type) (a b : id → Type) : + List (Type u) → List (Type u) → Type (u + 1) := + | Nil : Funs id a b [] [] | Cons {ity oty : Type u} {itys otys : List (Type u)} - (f : ity → Result oty) (tl : Funs itys otys) : Funs (ity :: itys) (oty :: otys) + (f : kk_ty id a b → ity → Result oty) (tl : Funs id a b itys otys) : + Funs id a b (ity :: itys) (oty :: otys) - theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs itys otys) : - itys.length = otys.length := + theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs id a b itys otys) : + otys.length = itys.length := match fl with | .Nil => by simp | .Cons f tl => have h:= Funs.length_eq tl by simp [h] + def fin_cast {n m : Nat} (h : m = n) (i : Fin n) : Fin m := + ⟨ i.val, by have h1:= i.isLt; simp_all ⟩ + @[simp] def Funs.cast_fin {itys otys : List (Type)} - (fl : Funs itys otys) (i : Fin itys.length) : Fin otys.length := - ⟨ i.val, by have h:= fl.length_eq; have h1:= i.isLt; simp_all ⟩ + (fl : Funs id a b itys otys) (i : Fin itys.length) : Fin otys.length := + fin_cast (fl.length_eq) i - def get_fun {itys otys : List (Type)} (fl : Funs itys otys) : - (i : Fin itys.length) → itys.get i → Result (otys.get (fl.cast_fin i)) := + def get_fun {itys otys : List (Type)} (fl : Funs id a b itys otys) : + (i : Fin itys.length) → kk_ty id a b → itys.get i → Result (otys.get (fl.cast_fin i)) := match fl with | .Nil => λ i => by have h:= i.isLt; simp at h - | @Funs.Cons ity oty itys1 otys1 f tl => + | @Funs.Cons id a b ity oty itys1 otys1 f tl => λ i => if h: i.val = 0 then Eq.mp (by cases i; simp_all [List.get]) f @@ -582,10 +594,9 @@ namespace FixI (by cases Heq: i; rename_i val isLt; cases Heq': j; rename_i val' isLt; - cases val <;> simp_all [List.get]) + cases val <;> simp_all [List.get, fin_cast]) (get_fun tl j) - -- TODO: move theorem add_one_le_iff_le_ne (n m : Nat) (h1 : m ≤ n) (h2 : m ≠ n) : m + 1 ≤ n := by -- Damn, those proofs on natural numbers are hard - I wish Omega was in mathlib4... @@ -612,7 +623,7 @@ namespace FixI := by generalize h: (n - m) = k revert m - induction k + induction k -- TODO: induction h rather? case zero => simp_all intro m h1 h2 @@ -683,6 +694,65 @@ namespace FixI . apply Hgvalid . apply Hhvalid + def Funs.is_valid_p + (k : k_ty id a b) + (fl : Funs id a b itys otys) : + Prop := + match fl with + | .Nil => True + | .Cons f fl => (∀ x, FixI.is_valid_p k (λ k => f k x)) ∧ fl.is_valid_p k + + #check Subtype + def Funs.is_valid_p_is_valid_p_aux + {k : k_ty id a b} + {itys otys : List Type} + (Heq : List.length otys = List.length itys) + (fl : Funs id a b itys otys) (Hvalid : is_valid_p k fl) : + ∀ (i : Fin itys.length) (x : itys.get i), FixI.is_valid_p k (fun k => get_fun fl i k x) := by + -- Prepare the induction + have ⟨ n, Hn ⟩ : { n : Nat // itys.length = n } := ⟨ itys.length, by rfl ⟩ + revert itys otys Heq fl Hvalid + induction n + -- + case zero => + intro itys otys Heq fl Hvalid Hlen; + have Heq: itys = [] := by cases itys <;> simp_all + have Heq: otys = [] := by cases otys <;> simp_all + intro i x + simp_all + have Hi := i.isLt + simp_all + case succ n Hn => + intro itys otys Heq fl Hvalid Hlen i x; + cases fl <;> simp at * + rename_i ity oty itys otys f fl + have ⟨ Hvf, Hvalid ⟩ := Hvalid + have Hvf1: is_valid_p k fl := by + simp_all [Funs.is_valid_p] + have Hn := @Hn itys otys (by simp[*]) fl Hvf1 (by simp [*]) + -- Case disjunction on i + match i with + | ⟨ 0, _ ⟩ => + simp at x + simp [get_fun] + apply (Hvf x) + | ⟨ .succ j, HiLt ⟩ => + simp_arith at HiLt + simp at x + let j : Fin (List.length itys) := ⟨ j, by simp_arith [HiLt] ⟩ + have Hn := Hn j x + apply Hn + + def Funs.is_valid_p_is_valid_p + (itys otys : List (Type)) (Heq: otys.length = itys.length := by decide) + (k : k_ty (Fin (List.length itys)) (List.get itys) fun i => List.get otys (fin_cast Heq i)) + (fl : Funs (Fin itys.length) itys.get (λ i => otys.get (fin_cast Heq i)) itys otys) : + fl.is_valid_p k → + ∀ (i : Fin itys.length) (x : itys.get i), FixI.is_valid_p k (fun k => get_fun fl i k x) + := by + intro Hvalid + apply is_valid_p_is_valid_p_aux <;> simp [*] + end FixI namespace Ex1 @@ -909,12 +979,12 @@ namespace Ex4 let b ← k 0 (i - 1) .ret b - @[simp] def bodies (k : (i : Fin 2) → input_ty i → Result (output_ty i)) : - Funs [Int, Int] [Bool, Bool] := - Funs.Cons (is_even_body k) (Funs.Cons (is_odd_body k) Funs.Nil) + @[simp] def bodies : + Funs (Fin 2) input_ty output_ty [Int, Int] [Bool, Bool] := + Funs.Cons (is_even_body) (Funs.Cons (is_odd_body) Funs.Nil) def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : - input_ty i → Result (output_ty i) := get_fun (bodies k) i + input_ty i → Result (output_ty i) := get_fun bodies i k theorem body_is_valid : is_valid body := by -- Split the proof into proofs of validity of the individual bodies @@ -933,6 +1003,21 @@ namespace Ex4 . split <;> simp apply is_valid_p_bind <;> simp + theorem body_is_valid' : is_valid body := by + -- Split the proof into proofs of validity of the individual bodies + rw [is_valid] + simp [body] + intro k + apply (Funs.is_valid_p_is_valid_p [Int, Int] [Bool, Bool]) + simp [Funs.is_valid_p] + (repeat (apply And.intro)) <;> intro x <;> simp at x <;> + simp [is_even_body, is_odd_body] + -- Prove the validity of the individual bodies + . split <;> simp + apply is_valid_p_bind <;> simp + . split <;> simp + apply is_valid_p_bind <;> simp + theorem body_fix_eq : fix body = body (fix body) := is_valid_fix_fixed_eq body_is_valid -- cgit v1.2.3 From 2554a0a64d761a82789b7eacbfa3ca2c88eec7df Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 27 Jun 2023 19:06:14 +0200 Subject: Reduce the time spent on some proofs --- backends/lean/Base/Diverge.lean | 48 +++++++++++++++++------------------------ 1 file changed, 20 insertions(+), 28 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index c97674dd..c62e6dd5 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -25,6 +25,11 @@ TODO: - simplifier/rewriter have a strange behavior sometimes -/ + +/- TODO: this is very useful, but is there more? -/ +set_option profiler true +set_option profiler.threshold 100 + namespace Diverge namespace Primitives @@ -533,10 +538,10 @@ namespace FixI fix f = f (fix f) := by have Hvalid' : Fix.is_valid (k_to_gen f) := by intro k x - simp [is_valid, is_valid_p] at Hvalid + simp only [is_valid, is_valid_p] at Hvalid let ⟨ i, x ⟩ := x have Hvalid := Hvalid (k_of_gen k) i x - simp [k_to_gen, k_of_gen, kk_to_gen, kk_of_gen] at Hvalid + simp only [k_to_gen, k_of_gen, kk_to_gen, kk_of_gen] at Hvalid refine Hvalid have Heq := Fix.is_valid_fix_fixed_eq Hvalid' simp [fix] @@ -612,7 +617,7 @@ namespace FixI termination_by for_all_fin_aux n _ m h => n - m decreasing_by simp_wf - apply Nat.sub_add_lt_sub <;> simp_all + apply Nat.sub_add_lt_sub <;> simp simp_all [add_one_le_iff_le_ne] def for_all_fin {n : Nat} (f : Fin n → Prop) := for_all_fin_aux f 0 (by simp) @@ -665,6 +670,7 @@ namespace FixI simp_all . simp_all [add_one_le_iff_le_ne] + -- TODO: this is not necessary anymore theorem for_all_fin_imp_forall (n : Nat) (f : Fin n → Prop) : for_all_fin f → ∀ i, f i := by @@ -702,7 +708,6 @@ namespace FixI | .Nil => True | .Cons f fl => (∀ x, FixI.is_valid_p k (λ k => f k x)) ∧ fl.is_valid_p k - #check Subtype def Funs.is_valid_p_is_valid_p_aux {k : k_ty id a b} {itys otys : List Type} @@ -724,11 +729,11 @@ namespace FixI simp_all case succ n Hn => intro itys otys Heq fl Hvalid Hlen i x; - cases fl <;> simp at * + cases fl <;> simp at Hlen i x Heq Hvalid rename_i ity oty itys otys f fl have ⟨ Hvf, Hvalid ⟩ := Hvalid have Hvf1: is_valid_p k fl := by - simp_all [Funs.is_valid_p] + simp [Hvalid, Funs.is_valid_p] have Hn := @Hn itys otys (by simp[*]) fl Hvf1 (by simp [*]) -- Case disjunction on i match i with @@ -989,29 +994,12 @@ namespace Ex4 theorem body_is_valid : is_valid body := by -- Split the proof into proofs of validity of the individual bodies rw [is_valid] - simp [body] - intro k - apply for_all_fin_imp_forall - simp [for_all_fin] - repeat (unfold for_all_fin_aux; simp) - simp [get_fun] - (repeat (apply And.intro)) <;> intro x <;> simp at x <;> - simp [is_even_body, is_odd_body] - -- Prove the validity of the individual bodies - . split <;> simp - apply is_valid_p_bind <;> simp - . split <;> simp - apply is_valid_p_bind <;> simp - - theorem body_is_valid' : is_valid body := by - -- Split the proof into proofs of validity of the individual bodies - rw [is_valid] - simp [body] + simp only [body] intro k apply (Funs.is_valid_p_is_valid_p [Int, Int] [Bool, Bool]) simp [Funs.is_valid_p] (repeat (apply And.intro)) <;> intro x <;> simp at x <;> - simp [is_even_body, is_odd_body] + simp only [is_even_body, is_odd_body] -- Prove the validity of the individual bodies . split <;> simp apply is_valid_p_bind <;> simp @@ -1088,11 +1076,15 @@ namespace Ex5 theorem id_body_is_valid : ∀ k x, is_valid_p k (λ k => @id_body a k x) := by intro k x - simp [id_body] + simp only [id_body] split <;> simp - apply is_valid_p_bind <;> simp_all + apply is_valid_p_bind <;> simp [*] -- We have to show that `map k tl` is valid - apply map_is_valid; simp + apply map_is_valid; + -- Remark: if we don't do the intro, then the last step is expensive: + -- "typeclass inference of Nonempty took 119ms" + intro k x + simp only [is_valid_p_same, is_valid_p_rec] noncomputable def id (t : Tree a) := fix id_body t -- cgit v1.2.3 From 19bde89b84619defc2a822c3bf96bdca9c97eee7 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 28 Jun 2023 12:16:10 +0200 Subject: Reorganize backends/lean/Base --- backends/lean/Base/Diverge.lean | 1104 +-------------------------------------- 1 file changed, 2 insertions(+), 1102 deletions(-) (limited to 'backends/lean/Base/Diverge.lean') diff --git a/backends/lean/Base/Diverge.lean b/backends/lean/Base/Diverge.lean index c62e6dd5..c9a2eec2 100644 --- a/backends/lean/Base/Diverge.lean +++ b/backends/lean/Base/Diverge.lean @@ -3,1105 +3,5 @@ import Lean.Meta.Tactic.Simp import Init.Data.List.Basic import Mathlib.Tactic.RunCmd import Mathlib.Tactic.Linarith - -/- -TODO: -- we want an easier to use cases: - - keeps in the goal an equation of the shape: `t = case` - - if called on Prop terms, uses Classical.em - Actually, the cases from mathlib seems already quite powerful - (https://leanprover-community.github.io/mathlib_docs/tactics.html#cases) - For instance: cases h : e - Also: cases_matching -- better split tactic -- we need conversions to operate on the head of applications. - Actually, something like this works: - ``` - conv at Hl => - apply congr_fun - simp [fix_fuel_P] - ``` - Maybe we need a rpt ... ; focus? -- simplifier/rewriter have a strange behavior sometimes --/ - - -/- TODO: this is very useful, but is there more? -/ -set_option profiler true -set_option profiler.threshold 100 - -namespace Diverge - -namespace Primitives -/-! # Copy-pasting from Primitives to make the file self-contained -/ - -inductive Error where - | assertionFailure: Error - | integerOverflow: Error - | divisionByZero: Error - | arrayOutOfBounds: Error - | maximumSizeExceeded: Error - | panic: Error -deriving Repr, BEq - -open Error - -inductive Result (α : Type u) where - | ret (v: α): Result α - | fail (e: Error): Result α - | div -deriving Repr, BEq - -open Result - -def bind (x: Result α) (f: α -> Result β) : Result β := - match x with - | ret v => f v - | fail v => fail v - | div => div - -@[simp] theorem bind_ret (x : α) (f : α → Result β) : bind (.ret x) f = f x := by simp [bind] -@[simp] theorem bind_fail (x : Error) (f : α → Result β) : bind (.fail x) f = .fail x := by simp [bind] -@[simp] theorem bind_div (f : α → Result β) : bind .div f = .div := by simp [bind] - --- Allows using Result in do-blocks -instance : Bind Result where - bind := bind - --- Allows using return x in do-blocks -instance : Pure Result where - pure := fun x => ret x - -@[simp] theorem bind_tc_ret (x : α) (f : α → Result β) : - (do let y ← .ret x; f y) = f x := by simp [Bind.bind, bind] - -@[simp] theorem bind_tc_fail (x : Error) (f : α → Result β) : - (do let y ← fail x; f y) = fail x := by simp [Bind.bind, bind] - -@[simp] theorem bind_tc_div (f : α → Result β) : - (do let y ← div; f y) = div := by simp [Bind.bind, bind] - -def div? {α: Type} (r: Result α): Bool := - match r with - | div => true - | ret _ | fail _ => false - -end Primitives - -namespace Fix - - open Primitives - open Result - - variable {a : Type} {b : a → Type} - variable {c d : Type} - - /-! # The least fixed point definition and its properties -/ - - def least_p (p : Nat → Prop) (n : Nat) : Prop := p n ∧ (∀ m, m < n → ¬ p m) - noncomputable def least (p : Nat → Prop) : Nat := - Classical.epsilon (least_p p) - - -- Auxiliary theorem for [least_spec]: if there exists an `n` satisfying `p`, - -- there there exists a least `m` satisfying `p`. - theorem least_spec_aux (p : Nat → Prop) : ∀ (n : Nat), (hn : p n) → ∃ m, least_p p m := by - apply Nat.strongRec' - intros n hi hn - -- Case disjunction on: is n the smallest n satisfying p? - match Classical.em (∀ m, m < n → ¬ p m) with - | .inl hlt => - -- Yes: trivial - exists n - | .inr hlt => - simp at * - let ⟨ m, ⟨ hmlt, hm ⟩ ⟩ := hlt - have hi := hi m hmlt hm - apply hi - - -- The specification of [least]: either `p` is never satisfied, or it is satisfied - -- by `least p` and no `n < least p` satisfies `p`. - theorem least_spec (p : Nat → Prop) : (∀ n, ¬ p n) ∨ (p (least p) ∧ ∀ n, n < least p → ¬ p n) := by - -- Case disjunction on the existence of an `n` which satisfies `p` - match Classical.em (∀ n, ¬ p n) with - | .inl h => - -- There doesn't exist: trivial - apply (Or.inl h) - | .inr h => - -- There exists: we simply use `least_spec_aux` in combination with the property - -- of the epsilon operator - simp at * - let ⟨ n, hn ⟩ := h - apply Or.inr - have hl := least_spec_aux p n hn - have he := Classical.epsilon_spec hl - apply he - - /-! # The fixed point definitions -/ - - def fix_fuel (n : Nat) (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : - Result (b x) := - match n with - | 0 => .div - | n + 1 => - f (fix_fuel n f) x - - @[simp] def fix_fuel_pred (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (x : a) (n : Nat) := - not (div? (fix_fuel n f x)) - - def fix_fuel_P (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (x : a) (n : Nat) : Prop := - fix_fuel_pred f x n - - noncomputable - def fix (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : Result (b x) := - fix_fuel (least (fix_fuel_P f x)) f x - - /-! # The validity property -/ - - -- Monotonicity relation over results - -- TODO: generalize (we should parameterize the definition by a relation over `a`) - def result_rel {a : Type u} (x1 x2 : Result a) : Prop := - match x1 with - | div => True - | fail _ => x2 = x1 - | ret _ => x2 = x1 -- TODO: generalize - - -- Monotonicity relation over monadic arrows (i.e., Kleisli arrows) - def karrow_rel (k1 k2 : (x:a) → Result (b x)) : Prop := - ∀ x, result_rel (k1 x) (k2 x) - - -- Monotonicity property for function bodies - def is_mono (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := - ∀ {{k1 k2}}, karrow_rel k1 k2 → karrow_rel (f k1) (f k2) - - -- "Continuity" property. - -- We need this, and this looks a lot like continuity. Also see this paper: - -- https://inria.hal.science/file/index/docid/216187/filename/tarski.pdf - -- We define our "continuity" criteria so that it gives us what we need to - -- prove the fixed-point equation, and we can also easily manipulate it. - def is_cont (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := - ∀ x, (Hdiv : ∀ n, fix_fuel (.succ n) f x = div) → f (fix f) x = div - - /-! # The proof of the fixed-point equation -/ - theorem fix_fuel_mono {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} - (Hmono : is_mono f) : - ∀ {{n m}}, n ≤ m → karrow_rel (fix_fuel n f) (fix_fuel m f) := by - intros n - induction n - case zero => simp [karrow_rel, fix_fuel, result_rel] - case succ n1 Hi => - intros m Hle x - simp [result_rel] - match m with - | 0 => - exfalso - zify at * - linarith - | Nat.succ m1 => - simp_arith at Hle - simp [fix_fuel] - have Hi := Hi Hle - have Hmono := Hmono Hi x - simp [result_rel] at Hmono - apply Hmono - - @[simp] theorem neg_fix_fuel_P - {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} {x : a} {n : Nat} : - ¬ fix_fuel_P f x n ↔ (fix_fuel n f x = div) := by - simp [fix_fuel_P, div?] - cases fix_fuel n f x <;> simp - - theorem fix_fuel_fix_mono {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} (Hmono : is_mono f) : - ∀ n, karrow_rel (fix_fuel n f) (fix f) := by - intros n x - simp [result_rel] - have Hl := least_spec (fix_fuel_P f x) - simp at Hl - match Hl with - | .inl Hl => simp [*] - | .inr ⟨ Hl, Hn ⟩ => - match Classical.em (fix_fuel n f x = div) with - | .inl Hd => - simp [*] - | .inr Hd => - have Hineq : least (fix_fuel_P f x) ≤ n := by - -- Proof by contradiction - cases Classical.em (least (fix_fuel_P f x) ≤ n) <;> simp [*] - simp at * - rename_i Hineq - have Hn := Hn n Hineq - contradiction - have Hfix : ¬ (fix f x = div) := by - simp [fix] - -- By property of the least upper bound - revert Hd Hl - -- TODO: there is no conversion to select the head of a function! - conv => lhs; apply congr_fun; apply congr_fun; apply congr_fun; simp [fix_fuel_P, div?] - cases fix_fuel (least (fix_fuel_P f x)) f x <;> simp - have Hmono := fix_fuel_mono Hmono Hineq x - simp [result_rel] at Hmono - simp [fix] at * - cases Heq: fix_fuel (least (fix_fuel_P f x)) f x <;> - cases Heq':fix_fuel n f x <;> - simp_all - - theorem fix_fuel_P_least {f : ((x:a) → Result (b x)) → (x:a) → Result (b x)} (Hmono : is_mono f) : - ∀ {{x n}}, fix_fuel_P f x n → fix_fuel_P f x (least (fix_fuel_P f x)) := by - intros x n Hf - have Hfmono := fix_fuel_fix_mono Hmono n x - -- TODO: there is no conversion to select the head of a function! - conv => apply congr_fun; simp [fix_fuel_P] - simp [fix_fuel_P] at Hf - revert Hf Hfmono - simp [div?, result_rel, fix] - cases fix_fuel n f x <;> simp_all - - -- Prove the fixed point equation in the case there exists some fuel for which - -- the execution terminates - theorem fix_fixed_eq_terminates (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (Hmono : is_mono f) - (x : a) (n : Nat) (He : fix_fuel_P f x n) : - fix f x = f (fix f) x := by - have Hl := fix_fuel_P_least Hmono He - -- TODO: better control of simplification - conv at Hl => - apply congr_fun - simp [fix_fuel_P] - -- The least upper bound is > 0 - have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by - revert Hl - simp [div?] - cases least (fix_fuel_P f x) <;> simp [fix_fuel] - simp [Hsucc] at Hl - revert Hl - simp [*, div?, fix, fix_fuel] - -- Use the monotonicity - have Hfixmono := fix_fuel_fix_mono Hmono n - have Hvm := Hmono Hfixmono x - -- Use functional extensionality - simp [result_rel, fix] at Hvm - revert Hvm - split <;> simp [*] <;> intros <;> simp [*] - - theorem fix_fixed_eq_forall {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hmono : is_mono f) (Hcont : is_cont f) : - ∀ x, fix f x = f (fix f) x := by - intros x - -- Case disjunction: is there a fuel such that the execution successfully execute? - match Classical.em (∃ n, fix_fuel_P f x n) with - | .inr He => - -- No fuel: the fixed point evaluates to `div` - --simp [fix] at * - simp at * - conv => lhs; simp [fix] - have Hel := He (Nat.succ (least (fix_fuel_P f x))); simp [*, fix_fuel] at *; clear Hel - -- Use the "continuity" of `f` - have He : ∀ n, fix_fuel (.succ n) f x = div := by intros; simp [*] - have Hcont := Hcont x He - simp [Hcont] - | .inl ⟨ n, He ⟩ => apply fix_fixed_eq_terminates f Hmono x n He - - -- The final fixed point equation - theorem fix_fixed_eq {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hmono : is_mono f) (Hcont : is_cont f) : - fix f = f (fix f) := by - have Heq := fix_fixed_eq_forall Hmono Hcont - have Heq1 : fix f = (λ x => fix f x) := by simp - rw [Heq1] - conv => lhs; ext; simp [Heq] - - /-! # Making the proofs of validity manageable (and automatable) -/ - - -- Monotonicity property for expressions - def is_mono_p (e : ((x:a) → Result (b x)) → Result c) : Prop := - ∀ {{k1 k2}}, karrow_rel k1 k2 → result_rel (e k1) (e k2) - - theorem is_mono_p_same (x : Result c) : - @is_mono_p a b c (λ _ => x) := by - simp [is_mono_p, karrow_rel, result_rel] - split <;> simp - - theorem is_mono_p_rec (x : a) : - @is_mono_p a b (b x) (λ f => f x) := by - simp_all [is_mono_p, karrow_rel, result_rel] - - -- The important lemma about `is_mono_p` - theorem is_mono_p_bind - (g : ((x:a) → Result (b x)) → Result c) - (h : c → ((x:a) → Result (b x)) → Result d) : - is_mono_p g → - (∀ y, is_mono_p (h y)) → - @is_mono_p a b d (λ k => do let y ← g k; h y k) := by - intro hg hh - simp [is_mono_p] - intro fg fh Hrgh - simp [karrow_rel, result_rel] - have hg := hg Hrgh; simp [result_rel] at hg - cases heq0: g fg <;> simp_all - rename_i y _ - have hh := hh y Hrgh; simp [result_rel] at hh - simp_all - - -- Continuity property for expressions - note that we take the continuation - -- as parameter - def is_cont_p (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (e : ((x:a) → Result (b x)) → Result c) : Prop := - (Hc : ∀ n, e (fix_fuel n k) = .div) → - e (fix k) = .div - - theorem is_cont_p_same (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (x : Result c) : - is_cont_p k (λ _ => x) := by - simp [is_cont_p] - - theorem is_cont_p_rec (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : - is_cont_p f (λ f => f x) := by - simp_all [is_cont_p, fix] - - -- The important lemma about `is_cont_p` - theorem is_cont_p_bind - (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (Hkmono : is_mono k) - (g : ((x:a) → Result (b x)) → Result c) - (h : c → ((x:a) → Result (b x)) → Result d) : - is_mono_p g → - is_cont_p k g → - (∀ y, is_mono_p (h y)) → - (∀ y, is_cont_p k (h y)) → - is_cont_p k (λ k => do let y ← g k; h y k) := by - intro Hgmono Hgcont Hhmono Hhcont - simp [is_cont_p] - intro Hdiv - -- Case on `g (fix... k)`: is there an n s.t. it terminates? - cases Classical.em (∀ n, g (fix_fuel n k) = .div) <;> rename_i Hn - . -- Case 1: g diverges - have Hgcont := Hgcont Hn - simp_all - . -- Case 2: g doesn't diverge - simp at Hn - let ⟨ n, Hn ⟩ := Hn - have Hdivn := Hdiv n - have Hffmono := fix_fuel_fix_mono Hkmono n - have Hgeq := Hgmono Hffmono - simp [result_rel] at Hgeq - cases Heq: g (fix_fuel n k) <;> rename_i y <;> simp_all - -- Remains the .ret case - -- Use Hdiv to prove that: ∀ n, h y (fix_fuel n f) = div - -- We do this in two steps: first we prove it for m ≥ n - have Hhdiv: ∀ m, h y (fix_fuel m k) = .div := by - have Hhdiv : ∀ m, n ≤ m → h y (fix_fuel m k) = .div := by - -- We use the fact that `g (fix_fuel n f) = .div`, combined with Hdiv - intro m Hle - have Hdivm := Hdiv m - -- Monotonicity of g - have Hffmono := fix_fuel_mono Hkmono Hle - have Hgmono := Hgmono Hffmono - -- We need to clear Hdiv because otherwise simp_all rewrites Hdivm with Hdiv - clear Hdiv - simp_all [result_rel] - intro m - -- TODO: we shouldn't need the excluded middle here because it is decidable - cases Classical.em (n ≤ m) <;> rename_i Hl - . apply Hhdiv; assumption - . simp at Hl - -- Make a case disjunction on `h y (fix_fuel m k)`: if it is not equal - -- to div, use the monotonicity of `h y` - have Hle : m ≤ n := by linarith - have Hffmono := fix_fuel_mono Hkmono Hle - have Hmono := Hhmono y Hffmono - simp [result_rel] at Hmono - cases Heq: h y (fix_fuel m k) <;> simp_all - -- We can now use the continuity hypothesis for h - apply Hhcont; assumption - - -- The validity property for an expression - def is_valid_p (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) - (e : ((x:a) → Result (b x)) → Result c) : Prop := - is_mono_p e ∧ - (is_mono k → is_cont_p k e) - - @[simp] theorem is_valid_p_same - (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : Result c) : - is_valid_p k (λ _ => x) := by - simp [is_valid_p, is_mono_p_same, is_cont_p_same] - - @[simp] theorem is_valid_p_rec - (k : ((x:a) → Result (b x)) → (x:a) → Result (b x)) (x : a) : - is_valid_p k (λ k => k x) := by - simp_all [is_valid_p, is_mono_p_rec, is_cont_p_rec] - - -- Lean is good at unification: we can write a very general version - -- (in particular, it will manage to figure out `g` and `h` when we - -- apply the lemma) - theorem is_valid_p_bind - {{k : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - {{g : ((x:a) → Result (b x)) → Result c}} - {{h : c → ((x:a) → Result (b x)) → Result d}} - (Hgvalid : is_valid_p k g) - (Hhvalid : ∀ y, is_valid_p k (h y)) : - is_valid_p k (λ k => do let y ← g k; h y k) := by - let ⟨ Hgmono, Hgcont ⟩ := Hgvalid - simp [is_valid_p, forall_and] at Hhvalid - have ⟨ Hhmono, Hhcont ⟩ := Hhvalid - simp [← imp_forall_iff] at Hhcont - simp [is_valid_p]; constructor - . -- Monotonicity - apply is_mono_p_bind <;> assumption - . -- Continuity - intro Hkmono - have Hgcont := Hgcont Hkmono - have Hhcont := Hhcont Hkmono - apply is_cont_p_bind <;> assumption - - def is_valid (f : ((x:a) → Result (b x)) → (x:a) → Result (b x)) : Prop := - ∀ k x, is_valid_p k (λ k => f k x) - - theorem is_valid_p_imp_is_valid {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hvalid : is_valid f) : - is_mono f ∧ is_cont f := by - have Hmono : is_mono f := by - intro f h Hr x - have Hmono := Hvalid (λ _ _ => .div) x - have Hmono := Hmono.left - apply Hmono; assumption - have Hcont : is_cont f := by - intro x Hdiv - have Hcont := (Hvalid f x).right Hmono - simp [is_cont_p] at Hcont - apply Hcont - intro n - have Hdiv := Hdiv n - simp [fix_fuel] at Hdiv - simp [*] - simp [*] - - theorem is_valid_fix_fixed_eq {{f : ((x:a) → Result (b x)) → (x:a) → Result (b x)}} - (Hvalid : is_valid f) : - fix f = f (fix f) := by - have ⟨ Hmono, Hcont ⟩ := is_valid_p_imp_is_valid Hvalid - exact fix_fixed_eq Hmono Hcont - -end Fix - -namespace FixI - /- Indexed fixed-point: definitions with indexed types, convenient to use for mutually - recursive definitions. We simply port the definitions and proofs from Fix to a more - specific case. - -/ - open Primitives Fix - - -- The index type - variable {id : Type} - - -- The input/output types - variable {a b : id → Type} - - -- Monotonicity relation over monadic arrows (i.e., Kleisli arrows) - def karrow_rel (k1 k2 : (i:id) → a i → Result (b i)) : Prop := - ∀ i x, result_rel (k1 i x) (k2 i x) - - def kk_to_gen (k : (i:id) → a i → Result (b i)) : - (x: (i:id) × a i) → Result (b x.fst) := - λ ⟨ i, x ⟩ => k i x - - def kk_of_gen (k : (x: (i:id) × a i) → Result (b x.fst)) : - (i:id) → a i → Result (b i) := - λ i x => k ⟨ i, x ⟩ - - def k_to_gen (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : - ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst) := - λ kk => kk_to_gen (k (kk_of_gen kk)) - - def k_of_gen (k : ((x: (i:id) × a i) → Result (b x.fst)) → (x: (i:id) × a i) → Result (b x.fst)) : - ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i) := - λ kk => kk_of_gen (k (kk_to_gen kk)) - - def e_to_gen (e : ((i:id) → a i → Result (b i)) → Result c) : - ((x: (i:id) × a i) → Result (b x.fst)) → Result c := - λ k => e (kk_of_gen k) - - def is_valid_p (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) - (e : ((i:id) → a i → Result (b i)) → Result c) : Prop := - Fix.is_valid_p (k_to_gen k) (e_to_gen e) - - def is_valid (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : Prop := - ∀ k i x, is_valid_p k (λ k => f k i x) - - noncomputable def fix - (f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) : - (i:id) → a i → Result (b i) := - kk_of_gen (Fix.fix (k_to_gen f)) - - theorem is_valid_fix_fixed_eq - {{f : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)}} - (Hvalid : is_valid f) : - fix f = f (fix f) := by - have Hvalid' : Fix.is_valid (k_to_gen f) := by - intro k x - simp only [is_valid, is_valid_p] at Hvalid - let ⟨ i, x ⟩ := x - have Hvalid := Hvalid (k_of_gen k) i x - simp only [k_to_gen, k_of_gen, kk_to_gen, kk_of_gen] at Hvalid - refine Hvalid - have Heq := Fix.is_valid_fix_fixed_eq Hvalid' - simp [fix] - conv => lhs; rw [Heq] - - /- Some utilities to define the mutually recursive functions -/ - - -- TODO: use more - @[simp] def kk_ty (id : Type) (a b : id → Type) := (i:id) → a i → Result (b i) - @[simp] def k_ty (id : Type) (a b : id → Type) := kk_ty id a b → kk_ty id a b - - -- Initially, we had left out the parameters id, a and b. - -- However, by parameterizing Funs with those parameters, we can state - -- and prove lemmas like Funs.is_valid_p_is_valid_p - inductive Funs (id : Type) (a b : id → Type) : - List (Type u) → List (Type u) → Type (u + 1) := - | Nil : Funs id a b [] [] - | Cons {ity oty : Type u} {itys otys : List (Type u)} - (f : kk_ty id a b → ity → Result oty) (tl : Funs id a b itys otys) : - Funs id a b (ity :: itys) (oty :: otys) - - theorem Funs.length_eq {itys otys : List (Type)} (fl : Funs id a b itys otys) : - otys.length = itys.length := - match fl with - | .Nil => by simp - | .Cons f tl => - have h:= Funs.length_eq tl - by simp [h] - - def fin_cast {n m : Nat} (h : m = n) (i : Fin n) : Fin m := - ⟨ i.val, by have h1:= i.isLt; simp_all ⟩ - - @[simp] def Funs.cast_fin {itys otys : List (Type)} - (fl : Funs id a b itys otys) (i : Fin itys.length) : Fin otys.length := - fin_cast (fl.length_eq) i - - def get_fun {itys otys : List (Type)} (fl : Funs id a b itys otys) : - (i : Fin itys.length) → kk_ty id a b → itys.get i → Result (otys.get (fl.cast_fin i)) := - match fl with - | .Nil => λ i => by have h:= i.isLt; simp at h - | @Funs.Cons id a b ity oty itys1 otys1 f tl => - λ i => - if h: i.val = 0 then - Eq.mp (by cases i; simp_all [List.get]) f - else - let j := i.val - 1 - have Hj: j < itys1.length := by - have Hi := i.isLt - simp at Hi - revert Hi - cases Heq: i.val <;> simp_all - simp_arith - let j: Fin itys1.length := ⟨ j, Hj ⟩ - Eq.mp - (by - cases Heq: i; rename_i val isLt; - cases Heq': j; rename_i val' isLt; - cases val <;> simp_all [List.get, fin_cast]) - (get_fun tl j) - - -- TODO: move - theorem add_one_le_iff_le_ne (n m : Nat) (h1 : m ≤ n) (h2 : m ≠ n) : m + 1 ≤ n := by - -- Damn, those proofs on natural numbers are hard - I wish Omega was in mathlib4... - simp [Nat.add_one_le_iff] - simp [Nat.lt_iff_le_and_ne] - simp_all - - def for_all_fin_aux {n : Nat} (f : Fin n → Prop) (m : Nat) (h : m ≤ n) : Prop := - if heq: m = n then True - else - f ⟨ m, by simp_all [Nat.lt_iff_le_and_ne] ⟩ ∧ - for_all_fin_aux f (m + 1) (by simp_all [add_one_le_iff_le_ne]) - termination_by for_all_fin_aux n _ m h => n - m - decreasing_by - simp_wf - apply Nat.sub_add_lt_sub <;> simp - simp_all [add_one_le_iff_le_ne] - - def for_all_fin {n : Nat} (f : Fin n → Prop) := for_all_fin_aux f 0 (by simp) - - theorem for_all_fin_aux_imp_forall {n : Nat} (f : Fin n → Prop) (m : Nat) : - (h : m ≤ n) → - for_all_fin_aux f m h → ∀ i, m ≤ i.val → f i - := by - generalize h: (n - m) = k - revert m - induction k -- TODO: induction h rather? - case zero => - simp_all - intro m h1 h2 - have h: n = m := by - linarith - unfold for_all_fin_aux; simp_all - simp_all - -- There is no i s.t. m ≤ i - intro i h3; cases i; simp_all - linarith - case succ k hi => - simp_all - intro m hk hmn - intro hf i hmi - have hne: m ≠ n := by - have hineq := Nat.lt_of_sub_eq_succ hk - linarith - -- m = i? - if heq: m = i then - -- Yes: simply use the `for_all_fin_aux` hyp - unfold for_all_fin_aux at hf - simp_all - tauto - else - -- No: use the induction hypothesis - have hlt: m < i := by simp_all [Nat.lt_iff_le_and_ne] - have hineq: m + 1 ≤ n := by - have hineq := Nat.lt_of_sub_eq_succ hk - simp [*, Nat.add_one_le_iff] - have heq1: n - (m + 1) = k := by - -- TODO: very annoying arithmetic proof - simp [Nat.sub_eq_iff_eq_add hineq] - have hineq1: m ≤ n := by linarith - simp [Nat.sub_eq_iff_eq_add hineq1] at hk - simp_arith [hk] - have hi := hi (m + 1) heq1 hineq - apply hi <;> simp_all - . unfold for_all_fin_aux at hf - simp_all - . simp_all [add_one_le_iff_le_ne] - - -- TODO: this is not necessary anymore - theorem for_all_fin_imp_forall (n : Nat) (f : Fin n → Prop) : - for_all_fin f → ∀ i, f i - := by - intro Hf i - apply for_all_fin_aux_imp_forall <;> try assumption - simp - - /- Automating the proofs -/ - @[simp] theorem is_valid_p_same - (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) (x : Result c) : - is_valid_p k (λ _ => x) := by - simp [is_valid_p, k_to_gen, e_to_gen] - - @[simp] theorem is_valid_p_rec - (k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)) (i : id) (x : a i) : - is_valid_p k (λ k => k i x) := by - simp [is_valid_p, k_to_gen, e_to_gen, kk_to_gen, kk_of_gen] - - theorem is_valid_p_bind - {{k : ((i:id) → a i → Result (b i)) → (i:id) → a i → Result (b i)}} - {{g : ((i:id) → a i → Result (b i)) → Result c}} - {{h : c → ((i:id) → a i → Result (b i)) → Result d}} - (Hgvalid : is_valid_p k g) - (Hhvalid : ∀ y, is_valid_p k (h y)) : - is_valid_p k (λ k => do let y ← g k; h y k) := by - apply Fix.is_valid_p_bind - . apply Hgvalid - . apply Hhvalid - - def Funs.is_valid_p - (k : k_ty id a b) - (fl : Funs id a b itys otys) : - Prop := - match fl with - | .Nil => True - | .Cons f fl => (∀ x, FixI.is_valid_p k (λ k => f k x)) ∧ fl.is_valid_p k - - def Funs.is_valid_p_is_valid_p_aux - {k : k_ty id a b} - {itys otys : List Type} - (Heq : List.length otys = List.length itys) - (fl : Funs id a b itys otys) (Hvalid : is_valid_p k fl) : - ∀ (i : Fin itys.length) (x : itys.get i), FixI.is_valid_p k (fun k => get_fun fl i k x) := by - -- Prepare the induction - have ⟨ n, Hn ⟩ : { n : Nat // itys.length = n } := ⟨ itys.length, by rfl ⟩ - revert itys otys Heq fl Hvalid - induction n - -- - case zero => - intro itys otys Heq fl Hvalid Hlen; - have Heq: itys = [] := by cases itys <;> simp_all - have Heq: otys = [] := by cases otys <;> simp_all - intro i x - simp_all - have Hi := i.isLt - simp_all - case succ n Hn => - intro itys otys Heq fl Hvalid Hlen i x; - cases fl <;> simp at Hlen i x Heq Hvalid - rename_i ity oty itys otys f fl - have ⟨ Hvf, Hvalid ⟩ := Hvalid - have Hvf1: is_valid_p k fl := by - simp [Hvalid, Funs.is_valid_p] - have Hn := @Hn itys otys (by simp[*]) fl Hvf1 (by simp [*]) - -- Case disjunction on i - match i with - | ⟨ 0, _ ⟩ => - simp at x - simp [get_fun] - apply (Hvf x) - | ⟨ .succ j, HiLt ⟩ => - simp_arith at HiLt - simp at x - let j : Fin (List.length itys) := ⟨ j, by simp_arith [HiLt] ⟩ - have Hn := Hn j x - apply Hn - - def Funs.is_valid_p_is_valid_p - (itys otys : List (Type)) (Heq: otys.length = itys.length := by decide) - (k : k_ty (Fin (List.length itys)) (List.get itys) fun i => List.get otys (fin_cast Heq i)) - (fl : Funs (Fin itys.length) itys.get (λ i => otys.get (fin_cast Heq i)) itys otys) : - fl.is_valid_p k → - ∀ (i : Fin itys.length) (x : itys.get i), FixI.is_valid_p k (fun k => get_fun fl i k x) - := by - intro Hvalid - apply is_valid_p_is_valid_p_aux <;> simp [*] - -end FixI - -namespace Ex1 - /- An example of use of the fixed-point -/ - open Primitives Fix - - variable {a : Type} (k : (List a × Int) → Result a) - - def list_nth_body (x : (List a × Int)) : Result a := - let (ls, i) := x - match ls with - | [] => .fail .panic - | hd :: tl => - if i = 0 then .ret hd - else k (tl, i - 1) - - theorem list_nth_body_is_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by - intro k x - simp [list_nth_body] - split <;> simp - split <;> simp - - noncomputable - def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) - - -- The unfolding equation - diverges if `i < 0` - theorem list_nth_eq (ls : List a) (i : Int) : - list_nth ls i = - match ls with - | [] => .fail .panic - | hd :: tl => - if i = 0 then .ret hd - else list_nth tl (i - 1) - := by - have Heq := is_valid_fix_fixed_eq (@list_nth_body_is_valid a) - simp [list_nth] - conv => lhs; rw [Heq] - -end Ex1 - -namespace Ex2 - /- Same as Ex1, but we make the body of nth non tail-rec (this is mostly - to see what happens when there are let-bindings) -/ - open Primitives Fix - - variable {a : Type} (k : (List a × Int) → Result a) - - def list_nth_body (x : (List a × Int)) : Result a := - let (ls, i) := x - match ls with - | [] => .fail .panic - | hd :: tl => - if i = 0 then .ret hd - else - do - let y ← k (tl, i - 1) - .ret y - - theorem list_nth_body_is_valid: ∀ k x, is_valid_p k (λ k => @list_nth_body a k x) := by - intro k x - simp [list_nth_body] - split <;> simp - split <;> simp - apply is_valid_p_bind <;> intros <;> simp_all - - noncomputable - def list_nth (ls : List a) (i : Int) : Result a := fix list_nth_body (ls, i) - - -- The unfolding equation - diverges if `i < 0` - theorem list_nth_eq (ls : List a) (i : Int) : - (list_nth ls i = - match ls with - | [] => .fail .panic - | hd :: tl => - if i = 0 then .ret hd - else - do - let y ← list_nth tl (i - 1) - .ret y) - := by - have Heq := is_valid_fix_fixed_eq (@list_nth_body_is_valid a) - simp [list_nth] - conv => lhs; rw [Heq] - -end Ex2 - -namespace Ex3 - /- Mutually recursive functions - first encoding (see Ex4 for a better encoding) -/ - open Primitives Fix - - /- Because we have mutually recursive functions, we use a sum for the inputs - and the output types: - - inputs: the sum allows to select the function to call in the recursive - calls (and the functions may not have the same input types) - - outputs: this case is degenerate because `even` and `odd` have the same - return type `Bool`, but generally speaking we need a sum type because - the functions in the mutually recursive group may have different - return types. - -/ - variable (k : (Int ⊕ Int) → Result (Bool ⊕ Bool)) - - def is_even_is_odd_body (x : (Int ⊕ Int)) : Result (Bool ⊕ Bool) := - match x with - | .inl i => - -- Body of `is_even` - if i = 0 - then .ret (.inl true) -- We use .inl because this is `is_even` - else - do - let b ← - do - -- Call `odd`: we need to wrap the input value in `.inr`, then - -- extract the output value - let r ← k (.inr (i- 1)) - match r with - | .inl _ => .fail .panic -- Invalid output - | .inr b => .ret b - -- Wrap the return value - .ret (.inl b) - | .inr i => - -- Body of `is_odd` - if i = 0 - then .ret (.inr false) -- We use .inr because this is `is_odd` - else - do - let b ← - do - -- Call `is_even`: we need to wrap the input value in .inr, then - -- extract the output value - let r ← k (.inl (i- 1)) - match r with - | .inl b => .ret b - | .inr _ => .fail .panic -- Invalid output - -- Wrap the return value - .ret (.inr b) - - theorem is_even_is_odd_body_is_valid: - ∀ k x, is_valid_p k (λ k => is_even_is_odd_body k x) := by - intro k x - simp [is_even_is_odd_body] - split <;> simp <;> split <;> simp - apply is_valid_p_bind; simp - intros; split <;> simp - apply is_valid_p_bind; simp - intros; split <;> simp - - noncomputable - def is_even (i : Int): Result Bool := - do - let r ← fix is_even_is_odd_body (.inl i) - match r with - | .inl b => .ret b - | .inr _ => .fail .panic - - noncomputable - def is_odd (i : Int): Result Bool := - do - let r ← fix is_even_is_odd_body (.inr i) - match r with - | .inl _ => .fail .panic - | .inr b => .ret b - - -- The unfolding equation for `is_even` - diverges if `i < 0` - theorem is_even_eq (i : Int) : - is_even i = (if i = 0 then .ret true else is_odd (i - 1)) - := by - have Heq := is_valid_fix_fixed_eq is_even_is_odd_body_is_valid - simp [is_even, is_odd] - conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp - -- Very annoying: we need to swap the matches - -- Doing this with rewriting lemmas is hard generally speaking - -- (especially as we may have to generate lemmas for user-defined - -- inductives on the fly). - -- The simplest is to repeatedly split then simplify (we identify - -- the outer match or monadic let-binding, and split on its scrutinee) - split <;> simp - cases H0 : fix is_even_is_odd_body (Sum.inr (i - 1)) <;> simp - rename_i v - split <;> simp - - -- The unfolding equation for `is_odd` - diverges if `i < 0` - theorem is_odd_eq (i : Int) : - is_odd i = (if i = 0 then .ret false else is_even (i - 1)) - := by - have Heq := is_valid_fix_fixed_eq is_even_is_odd_body_is_valid - simp [is_even, is_odd] - conv => lhs; rw [Heq]; simp; rw [is_even_is_odd_body]; simp - -- Same remark as for `even` - split <;> simp - cases H0 : fix is_even_is_odd_body (Sum.inl (i - 1)) <;> simp - rename_i v - split <;> simp - -end Ex3 - -namespace Ex4 - /- Mutually recursive functions - 2nd encoding -/ - open Primitives FixI - - attribute [local simp] List.get - - /- We make the input type and output types dependent on a parameter -/ - @[simp] def input_ty (i : Fin 2) : Type := - [Int, Int].get i - - @[simp] def output_ty (i : Fin 2) : Type := - [Bool, Bool].get i - - /- The continuation -/ - variable (k : (i : Fin 2) → input_ty i → Result (output_ty i)) - - /- The bodies are more natural -/ - def is_even_body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i : Int) : Result Bool := - if i = 0 - then .ret true - else do - let b ← k 1 (i - 1) - .ret b - - def is_odd_body (i : Int) : Result Bool := - if i = 0 - then .ret false - else do - let b ← k 0 (i - 1) - .ret b - - @[simp] def bodies : - Funs (Fin 2) input_ty output_ty [Int, Int] [Bool, Bool] := - Funs.Cons (is_even_body) (Funs.Cons (is_odd_body) Funs.Nil) - - def body (k : (i : Fin 2) → input_ty i → Result (output_ty i)) (i: Fin 2) : - input_ty i → Result (output_ty i) := get_fun bodies i k - - theorem body_is_valid : is_valid body := by - -- Split the proof into proofs of validity of the individual bodies - rw [is_valid] - simp only [body] - intro k - apply (Funs.is_valid_p_is_valid_p [Int, Int] [Bool, Bool]) - simp [Funs.is_valid_p] - (repeat (apply And.intro)) <;> intro x <;> simp at x <;> - simp only [is_even_body, is_odd_body] - -- Prove the validity of the individual bodies - . split <;> simp - apply is_valid_p_bind <;> simp - . split <;> simp - apply is_valid_p_bind <;> simp - - theorem body_fix_eq : fix body = body (fix body) := - is_valid_fix_fixed_eq body_is_valid - - noncomputable def is_even (i : Int) : Result Bool := fix body 0 i - noncomputable def is_odd (i : Int) : Result Bool := fix body 1 i - - theorem is_even_eq (i : Int) : is_even i = - (if i = 0 - then .ret true - else do - let b ← is_odd (i - 1) - .ret b) := by - simp [is_even, is_odd]; - conv => lhs; rw [body_fix_eq] - - theorem is_odd_eq (i : Int) : is_odd i = - (if i = 0 - then .ret false - else do - let b ← is_even (i - 1) - .ret b) := by - simp [is_even, is_odd]; - conv => lhs; rw [body_fix_eq] - -end Ex4 - -namespace Ex5 - /- Higher-order example -/ - open Primitives Fix - - variable {a b : Type} - - /- An auxiliary function, which doesn't require the fixed-point -/ - def map (f : a → Result b) (ls : List a) : Result (List b) := - match ls with - | [] => .ret [] - | hd :: tl => - do - let hd ← f hd - let tl ← map f tl - .ret (hd :: tl) - - /- The validity theorem for `map`, generic in `f` -/ - theorem map_is_valid - {{f : (a → Result b) → a → Result c}} - (Hfvalid : ∀ k x, is_valid_p k (λ k => f k x)) - (k : (a → Result b) → a → Result b) - (ls : List a) : - is_valid_p k (λ k => map (f k) ls) := by - induction ls <;> simp [map] - apply is_valid_p_bind <;> simp_all - intros - apply is_valid_p_bind <;> simp_all - - /- An example which uses map -/ - inductive Tree (a : Type) := - | leaf (x : a) - | node (tl : List (Tree a)) - - def id_body (k : Tree a → Result (Tree a)) (t : Tree a) : Result (Tree a) := - match t with - | .leaf x => .ret (.leaf x) - | .node tl => - do - let tl ← map k tl - .ret (.node tl) - - theorem id_body_is_valid : - ∀ k x, is_valid_p k (λ k => @id_body a k x) := by - intro k x - simp only [id_body] - split <;> simp - apply is_valid_p_bind <;> simp [*] - -- We have to show that `map k tl` is valid - apply map_is_valid; - -- Remark: if we don't do the intro, then the last step is expensive: - -- "typeclass inference of Nonempty took 119ms" - intro k x - simp only [is_valid_p_same, is_valid_p_rec] - - noncomputable def id (t : Tree a) := fix id_body t - - -- The unfolding equation - theorem id_eq (t : Tree a) : - (id t = - match t with - | .leaf x => .ret (.leaf x) - | .node tl => - do - let tl ← map id tl - .ret (.node tl)) - := by - have Heq := is_valid_fix_fixed_eq (@id_body_is_valid a) - simp [id] - conv => lhs; rw [Heq]; simp; rw [id_body] - -end Ex5 - -end Diverge +import Base.Diverge.Base +import Base.Diverge.Elab -- cgit v1.2.3