From a68c231db4edf97c4f007724969aec7dd60941a1 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 2 Feb 2024 20:48:26 +0100 Subject: Update lean to v4.6.0-rc1 and start fixing the proofs --- backends/lean/Base/Diverge/Base.lean | 37 ++++++---- backends/lean/Base/Diverge/Elab.lean | 11 ++- backends/lean/Base/Diverge/ElabBase.lean | 16 +++-- backends/lean/Base/Extensions.lean | 10 +-- backends/lean/Base/Progress/Base.lean | 14 ++-- backends/lean/Base/Utils.lean | 6 +- backends/lean/lake-manifest.json | 117 ++++++++++++++++++------------- backends/lean/lean-toolchain | 2 +- 8 files changed, 120 insertions(+), 93 deletions(-) (limited to 'backends/lean') diff --git a/backends/lean/Base/Diverge/Base.lean b/backends/lean/Base/Diverge/Base.lean index 9458c926..e40432bd 100644 --- a/backends/lean/Base/Diverge/Base.lean +++ b/backends/lean/Base/Diverge/Base.lean @@ -21,7 +21,7 @@ namespace Lemmas else f ⟨ m, by simp_all [Nat.lt_iff_le_and_ne] ⟩ ∧ for_all_fin_aux f (m + 1) (by simp_all [Arith.add_one_le_iff_le_ne]) - termination_by for_all_fin_aux n _ m h => n - m + termination_by n - m decreasing_by simp_wf apply Nat.sub_add_lt_sub <;> try simp @@ -240,8 +240,8 @@ namespace Fix 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?] + conv => lhs; rw [fix_fuel_P] + simp [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 @@ -255,7 +255,7 @@ namespace Fix 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] + rw [fix_fuel_P] simp [fix_fuel_P] at Hf revert Hf Hfmono simp [div?, result_rel, fix] @@ -268,9 +268,7 @@ namespace Fix 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] + rw [fix_fuel_P] at Hl; simp at Hl -- The least upper bound is > 0 have ⟨ n, Hsucc ⟩ : ∃ n, least (fix_fuel_P f x) = Nat.succ n := by revert Hl @@ -618,12 +616,16 @@ namespace FixI @[simp] theorem is_valid_p_same (k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x)) (x : Result c) : is_valid_p k (λ _ => x) := by - simp [is_valid_p, k_to_gen, e_to_gen] + simp [is_valid_p] + unfold k_to_gen e_to_gen + simp @[simp] theorem is_valid_p_rec (k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x)) (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] + simp [is_valid_p] + unfold k_to_gen e_to_gen kk_to_gen kk_of_gen + simp theorem is_valid_p_ite (k : ((i:id) → (x:a i) → Result (b i x)) → (i:id) → (x:a i) → Result (b i x)) @@ -826,12 +828,16 @@ namespace FixII @[simp] theorem is_valid_p_same (k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t)) (x : Result c) : is_valid_p k (λ _ => x) := by - simp [is_valid_p, k_to_gen, e_to_gen] + simp [is_valid_p] + unfold k_to_gen e_to_gen + simp @[simp] theorem is_valid_p_rec (k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t)) (i : id) (t : ty i) (x : a i t) : is_valid_p k (λ k => k i t x) := by - simp [is_valid_p, k_to_gen, e_to_gen, kk_to_gen, kk_of_gen] + simp [is_valid_p] + unfold k_to_gen e_to_gen kk_to_gen kk_of_gen + simp theorem is_valid_p_ite (k : ((i:id) → (t:ty i) → a i t → Result (b i t)) → (i:id) → (t:ty i) → a i t → Result (b i t)) @@ -1531,10 +1537,11 @@ namespace Ex9 intro k a x simp only [id_body] split <;> try simp - apply is_valid_p_bind <;> try simp [*] - -- We have to show that `map k tl` is valid - -- Remark: `map_is_valid` doesn't work here, we need the specialized version - apply map_is_valid_simple + . apply is_valid_p_same + . apply is_valid_p_bind <;> try simp [*] + -- We have to show that `map k tl` is valid + -- Remark: `map_is_valid` doesn't work here, we need the specialized version + apply map_is_valid_simple def body (k : (i : Fin 1) → (t : ty i) → (x : input_ty i t) → Result (output_ty i t)) (i: Fin 1) : (t : ty i) → (x : input_ty i t) → Result (output_ty i t) := get_fun bodies i k diff --git a/backends/lean/Base/Diverge/Elab.lean b/backends/lean/Base/Diverge/Elab.lean index 6115b13b..3c2ea877 100644 --- a/backends/lean/Base/Diverge/Elab.lean +++ b/backends/lean/Base/Diverge/Elab.lean @@ -383,10 +383,7 @@ def mkFin (n : Nat) : Expr := def mkFinVal (n i : Nat) : MetaM Expr := do let n_lit : Expr := .lit (.natVal (n - 1)) let i_lit : Expr := .lit (.natVal i) - -- We could use `trySynthInstance`, but as we know the instance that we are - -- going to use, we can save the lookup - let ofNat ← mkAppOptM ``Fin.instOfNatFinHAddNatInstHAddInstAddNatOfNat #[n_lit, i_lit] - mkAppOptM ``OfNat.ofNat #[none, none, ofNat] + mkAppOptM ``Fin.ofNat #[.some n_lit, .some i_lit] /- Information about the type of a function in a declaration group. @@ -654,8 +651,8 @@ partial def proveExprIsValid (k_var kk_var : Expr) (e : Expr) : MetaM Expr := do -- Normalize to eliminate the lambdas - TODO: this is slightly dangerous. let e ← do if e.isLet ∧ normalize_let_bindings then do - let updt_config config := - { config with transparency := .reducible, zetaNonDep := false } + let updt_config (config : Lean.Meta.Config) := + { config with transparency := .reducible } let e ← withConfig updt_config (whnf e) trace[Diverge.def.valid] "e (after normalization): {e}" pure e @@ -929,7 +926,7 @@ partial def proveAppIsValidApplyThms (k_var kk_var : Expr) (e : Expr) -- We sometimes need to reduce the term - TODO: this is really dangerous let e ← do let updt_config config := - { config with transparency := .reducible, zetaNonDep := false } + { config with transparency := .reducible } withConfig updt_config (whnf e) trace[Diverge.def.valid] "e (after normalization): {e}" let e_valid ← proveExprIsValid k_var kk_var e diff --git a/backends/lean/Base/Diverge/ElabBase.lean b/backends/lean/Base/Diverge/ElabBase.lean index 0d33e9d2..08ef96f7 100644 --- a/backends/lean/Base/Diverge/ElabBase.lean +++ b/backends/lean/Base/Diverge/ElabBase.lean @@ -27,12 +27,12 @@ initialize registerTraceClass `Diverge.attr -- divspec attribute structure DivSpecAttr where attr : AttributeImpl - ext : DiscrTreeExtension Name true + ext : DiscrTreeExtension Name deriving Inhabited /- The persistent map from expressions to divspec theorems. -/ initialize divspecAttr : DivSpecAttr ← do - let ext ← mkDiscrTreeExtention `divspecMap true + let ext ← mkDiscrTreeExtention `divspecMap let attrImpl : AttributeImpl := { name := `divspec descr := "Marks theorems to use with the `divergent` encoding" @@ -44,7 +44,7 @@ initialize divspecAttr : DivSpecAttr ← do -- Lookup the theorem let env ← getEnv let thDecl := env.constants.find! thName - let fKey : Array (DiscrTree.Key true) ← MetaM.run' (do + let fKey : Array (DiscrTree.Key) ← MetaM.run' (do /- The theorem should have the shape: `∀ ..., is_valid_p k (λ k => ...)` @@ -59,7 +59,9 @@ initialize divspecAttr : DivSpecAttr ← do let (_, _, fExpr) ← lambdaMetaTelescope fExpr.consumeMData trace[Diverge] "Registering divspec theorem for {fExpr}" -- Convert the function expression to a discrimination tree key - DiscrTree.mkPath fExpr) + -- We use the default configuration + let config : WhnfCoreConfig := {} + DiscrTree.mkPath fExpr config) let env := ext.addEntry env (fKey, thName) setEnv env trace[Diverge] "Saved the environment" @@ -69,9 +71,11 @@ initialize divspecAttr : DivSpecAttr ← do pure { attr := attrImpl, ext := ext } def DivSpecAttr.find? (s : DivSpecAttr) (e : Expr) : MetaM (Array Name) := do - (s.ext.getState (← getEnv)).getMatch e + -- We use the default configuration + let config : WhnfCoreConfig := {} + (s.ext.getState (← getEnv)).getMatch e config -def DivSpecAttr.getState (s : DivSpecAttr) : MetaM (DiscrTree Name true) := do +def DivSpecAttr.getState (s : DivSpecAttr) : MetaM (DiscrTree Name) := do pure (s.ext.getState (← getEnv)) def showStoredDivSpec : MetaM Unit := do diff --git a/backends/lean/Base/Extensions.lean b/backends/lean/Base/Extensions.lean index b34f41dc..c0e80861 100644 --- a/backends/lean/Base/Extensions.lean +++ b/backends/lean/Base/Extensions.lean @@ -31,13 +31,13 @@ def mkMapDeclarationExtension [Inhabited α] (name : Name := by exact decl_name% store the keys from *after* the transformation (i.e., the `DiscrTreeKey` below). The transformation itself can be done elsewhere. -/ -abbrev DiscrTreeKey (simpleReduce : Bool) := Array (DiscrTree.Key simpleReduce) +abbrev DiscrTreeKey := Array DiscrTree.Key -abbrev DiscrTreeExtension (α : Type) (simpleReduce : Bool) := - SimplePersistentEnvExtension (DiscrTreeKey simpleReduce × α) (DiscrTree α simpleReduce) +abbrev DiscrTreeExtension (α : Type) := + SimplePersistentEnvExtension (DiscrTreeKey × α) (DiscrTree α) -def mkDiscrTreeExtention [Inhabited α] [BEq α] (name : Name := by exact decl_name%) (simpleReduce : Bool) : - IO (DiscrTreeExtension α simpleReduce) := +def mkDiscrTreeExtention [Inhabited α] [BEq α] (name : Name := by exact decl_name%) : + IO (DiscrTreeExtension α) := registerSimplePersistentEnvExtension { name := name, addImportedFn := fun a => a.foldl (fun s a => a.foldl (fun s (k, v) => s.insertCore k v) s) DiscrTree.empty, diff --git a/backends/lean/Base/Progress/Base.lean b/backends/lean/Base/Progress/Base.lean index a64212a5..03c80a42 100644 --- a/backends/lean/Base/Progress/Base.lean +++ b/backends/lean/Base/Progress/Base.lean @@ -139,12 +139,12 @@ def getPSpecFunArgsExpr (isGoal : Bool) (th : Expr) : MetaM Expr := -- pspec attribute structure PSpecAttr where attr : AttributeImpl - ext : DiscrTreeExtension Name true + ext : DiscrTreeExtension Name deriving Inhabited /- The persistent map from expressions to pspec theorems. -/ initialize pspecAttr : PSpecAttr ← do - let ext ← mkDiscrTreeExtention `pspecMap true + let ext ← mkDiscrTreeExtention `pspecMap let attrImpl : AttributeImpl := { name := `pspec descr := "Marks theorems to use with the `progress` tactic" @@ -160,7 +160,9 @@ initialize pspecAttr : PSpecAttr ← do let fExpr ← getPSpecFunArgsExpr false thDecl.type trace[Progress] "Registering spec theorem for {fExpr}" -- Convert the function expression to a discrimination tree key - DiscrTree.mkPath fExpr) + -- We use the default configuration + let config : WhnfCoreConfig := {} + DiscrTree.mkPath fExpr config) let env := ext.addEntry env (fKey, thName) setEnv env trace[Progress] "Saved the environment" @@ -170,9 +172,11 @@ initialize pspecAttr : PSpecAttr ← do pure { attr := attrImpl, ext := ext } def PSpecAttr.find? (s : PSpecAttr) (e : Expr) : MetaM (Array Name) := do - (s.ext.getState (← getEnv)).getMatch e + -- We use the default configuration + let config : WhnfCoreConfig := {} + (s.ext.getState (← getEnv)).getMatch e config -def PSpecAttr.getState (s : PSpecAttr) : MetaM (DiscrTree Name true) := do +def PSpecAttr.getState (s : PSpecAttr) : MetaM (DiscrTree Name) := do pure (s.ext.getState (← getEnv)) def showStoredPSpec : MetaM Unit := do diff --git a/backends/lean/Base/Utils.lean b/backends/lean/Base/Utils.lean index b0032281..eacfe72b 100644 --- a/backends/lean/Base/Utils.lean +++ b/backends/lean/Base/Utils.lean @@ -1,6 +1,5 @@ import Lean import Mathlib.Tactic.Core -import Mathlib.Tactic.LeftRight import Base.UtilsBase /- @@ -503,9 +502,8 @@ elab "split_disj " n:ident : tactic => do example (x y : Int) (h0 : x ≤ y ∨ x ≥ y) : x ≤ y ∨ x ≥ y := by split_disj h0 - . left; assumption - . right; assumption - + . apply Or.inl; assumption + . apply Or.inr; assumption -- Tactic to split on an exists. -- `h` must be an FVar diff --git a/backends/lean/lake-manifest.json b/backends/lean/lake-manifest.json index 934ee2d9..3a18466f 100644 --- a/backends/lean/lake-manifest.json +++ b/backends/lean/lake-manifest.json @@ -1,51 +1,68 @@ -{"version": 5, - "packagesDir": "lake-packages", +{"version": 7, + "packagesDir": ".lake/packages", "packages": - [{"git": - {"url": "https://github.com/EdAyers/ProofWidgets4", - "subDir?": null, - "rev": "a0c2cd0ac3245a0dade4f925bcfa97e06dd84229", - "opts": {}, - "name": "proofwidgets", - "inputRev?": "v0.0.13", - "inherited": true}}, - {"git": - {"url": "https://github.com/mhuisi/lean4-cli.git", - "subDir?": null, - "rev": "21dac2e9cc7e3cf7da5800814787b833e680b2fd", - "opts": {}, - "name": "Cli", - "inputRev?": "nightly", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover-community/mathlib4.git", - "subDir?": null, - "rev": "226948a52f8e19ad95ff6025a96784d7e7ed6ed0", - "opts": {}, - "name": "mathlib", - "inputRev?": null, - "inherited": false}}, - {"git": - {"url": "https://github.com/gebner/quote4", - "subDir?": null, - "rev": "e75daed95ad1c92af4e577fea95e234d7a8401c1", - "opts": {}, - "name": "Qq", - "inputRev?": "master", - "inherited": true}}, - {"git": - {"url": "https://github.com/JLimperg/aesop", - "subDir?": null, - "rev": "1a0cded2be292b5496e659b730d2accc742de098", - "opts": {}, - "name": "aesop", - "inputRev?": "master", - "inherited": true}}, - {"git": - {"url": "https://github.com/leanprover/std4", - "subDir?": null, - "rev": "ba5e5e3af519b4fc5221ad0fa4b2c87276f1d323", - "opts": {}, - "name": "std", - "inputRev?": "main", - "inherited": true}}]} + [{"url": "https://github.com/leanprover/std4", + "type": "git", + "subDir": null, + "rev": "276953b13323ca151939eafaaec9129bf7970306", + "name": "std", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/quote4", + "type": "git", + "subDir": null, + "rev": "1c88406514a636d241903e2e288d21dc6d861e01", + "name": "Qq", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/aesop", + "type": "git", + "subDir": null, + "rev": "6beed82dcfbb7731d173cd517675df27d62ad0f4", + "name": "aesop", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/ProofWidgets4", + "type": "git", + "subDir": null, + "rev": "af1e86cf7a37389632a02f4a111e6b501b2b818f", + "name": "proofwidgets", + "manifestFile": "lake-manifest.json", + "inputRev": "v0.0.27", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover/lean4-cli", + "type": "git", + "subDir": null, + "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", + "name": "Cli", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/import-graph.git", + "type": "git", + "subDir": null, + "rev": "8079d2d1d0e073bde42eab159c24f4c2d0d3a871", + "name": "importGraph", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/mathlib4.git", + "type": "git", + "subDir": null, + "rev": "056cc4b21e25e8d1daaeef3a6e3416872c9fc12c", + "name": "mathlib", + "manifestFile": "lake-manifest.json", + "inputRev": null, + "inherited": false, + "configFile": "lakefile.lean"}], + "name": "base", + "lakeDir": ".lake"} diff --git a/backends/lean/lean-toolchain b/backends/lean/lean-toolchain index fbca4d37..cfcdd327 100644 --- a/backends/lean/lean-toolchain +++ b/backends/lean/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.0.0 \ No newline at end of file +leanprover/lean4:v4.6.0-rc1 -- cgit v1.2.3