summaryrefslogtreecommitdiff
path: root/backends/lean/Base/Diverge
diff options
context:
space:
mode:
authorSon Ho2024-02-02 20:48:26 +0100
committerSon Ho2024-02-02 20:48:26 +0100
commita68c231db4edf97c4f007724969aec7dd60941a1 (patch)
tree7d6767e8f5ba5bda991baf02844afa29e14150ac /backends/lean/Base/Diverge
parent0960ad16838a43da3746f47cf5b640bfbb783d84 (diff)
Update lean to v4.6.0-rc1 and start fixing the proofs
Diffstat (limited to '')
-rw-r--r--backends/lean/Base/Diverge/Base.lean37
-rw-r--r--backends/lean/Base/Diverge/Elab.lean11
-rw-r--r--backends/lean/Base/Diverge/ElabBase.lean16
3 files changed, 36 insertions, 28 deletions
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