From d0b5f5d687eb5279a3a36a6595cfc107beeff18f Mon Sep 17 00:00:00 2001 From: Abdalrhman Mohamed Date: Tue, 22 Oct 2024 23:12:38 -0700 Subject: [PATCH 01/25] Delete profile messages. More elaborate profiling can be done using `set_option trace.profiler true`. --- Auto/Tactic.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 6575aee..b3d7061 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -478,7 +478,6 @@ def runAuto @[tactic auto] def evalAuto : Tactic | `(auto | auto $instr $hints $[$uords]*) => withMainContext do - let startTime ← IO.monoMsNow -- Suppose the goal is `∀ (x₁ x₂ ⋯ xₙ), G` -- First, apply `intros` to put `x₁ x₂ ⋯ xₙ` into the local context, -- now the goal is just `G` @@ -494,7 +493,6 @@ def evalAuto : Tactic let (lemmas, inhFacts) ← collectAllLemmas hints uords (goalBinders.push ngoal) let declName? ← Elab.Term.getDeclName? let proof ← runAuto declName? lemmas inhFacts - IO.println s!"Auto found proof. Time spent by auto : {(← IO.monoMsNow) - startTime}ms" absurd.assign proof | .useSorry => let proof ← Meta.mkAppM ``sorryAx #[Expr.const ``False [], Expr.const ``false []] @@ -564,7 +562,6 @@ def runNativeProverWithAuto @[tactic mononative] def evalMonoNative : Tactic | `(mononative | mononative $hints $[$uords]*) => withMainContext do - let startTime ← IO.monoMsNow -- Suppose the goal is `∀ (x₁ x₂ ⋯ xₙ), G` -- First, apply `intros` to put `x₁ x₂ ⋯ xₙ` into the local context, -- now the goal is just `G` @@ -576,7 +573,6 @@ def evalMonoNative : Tactic withMainContext do let (lemmas, inhFacts) ← collectAllLemmas hints uords (goalBinders.push ngoal) let proof ← monoInterface lemmas inhFacts Solver.Native.queryNative - IO.println s!"Auto found proof. Time spent by auto : {(← IO.monoMsNow) - startTime}ms" absurd.assign proof | _ => throwUnsupportedSyntax From 420bc92e96aad66f922f3e5a80748e7c8e601d14 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 6 Nov 2024 01:03:32 -0800 Subject: [PATCH 02/25] fix minor issue in monomorphization --- Auto/Translation/Monomorphization.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Auto/Translation/Monomorphization.lean b/Auto/Translation/Monomorphization.lean index 70442e1..ee9c6c9 100644 --- a/Auto/Translation/Monomorphization.lean +++ b/Auto/Translation/Monomorphization.lean @@ -381,9 +381,8 @@ def ConstInsts.canonicalize? (cis : ConstInsts) (ci : ConstInst) : MetaM (Option private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) : Expr → MetaM (HashSet LemmaInst) | .bvar _ => throwError "MLemmaInst.matchConstInst :: Loose bound variable" | e@(.app ..) => do - let fn := e.getAppFn let args := e.getAppArgs - let mut ret ← MLemmaInst.matchConstInst ci mi fn + let mut ret := HashSet.empty for arg in args do ret := mergeHashSet ret (← MLemmaInst.matchConstInst ci mi arg) let s ← saveState From 91e58201ad7715aab780781b6626a4141383aef9 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 20 Nov 2024 23:30:25 -0800 Subject: [PATCH 03/25] remove reorderForallInstDep --- Auto/Tactic.lean | 2 -- Auto/Translation/Assumptions.lean | 5 ++++- Auto/Translation/Monomorphization.lean | 3 +++ Test/SmtTranslation/Trigger.lean | 5 +++++ 4 files changed, 12 insertions(+), 3 deletions(-) diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 6575aee..1e9c509 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -184,7 +184,6 @@ def unfoldConstAndPreprocessLemma (unfolds : Array Prep.ConstUnfoldInfo) (lem : let type := Prep.unfoldConsts unfolds type let type ← Core.betaReduce (← instantiateMVars type) let lem := {lem with type := type} - let lem ← lem.reorderForallInstDep return lem /-- @@ -202,7 +201,6 @@ def unfoldConstAndprepReduceDefeq (unfolds : Array Prep.ConstUnfoldInfo) (lem : let type := Prep.unfoldConsts unfolds type let type ← Core.betaReduce (← instantiateMVars type) let lem := {lem with type := type} - let lem ← lem.reorderForallInstDep return lem def traceLemmas (pre : String) (lemmas : Array Lemma) : TacticM Unit := do diff --git a/Auto/Translation/Assumptions.lean b/Auto/Translation/Assumptions.lean index 81d5fd1..f0bcb56 100644 --- a/Auto/Translation/Assumptions.lean +++ b/Auto/Translation/Assumptions.lean @@ -97,7 +97,9 @@ def Lemma.equivQuick (lem₁ lem₂ : Lemma) : MetaM Bool := do let s₂₁ ← Lemma.subsumeQuick lem₂ lem₁ return s₁₂ && s₂₁ -/-- Reorder top-level `∀` so that (non-prop / dependent) ones precede other ones -/ +/- **TODO:** Determine whether this is useful -/ +/- Reorder top-level `∀` so that (non-prop / dependent) ones precede other ones -/ +/- def Lemma.reorderForallInstDep (lem : Lemma) : MetaM Lemma := do let depargs := HashSet.empty.insertMany (Expr.depArgs lem.type) Meta.forallTelescope lem.type fun xs body => do @@ -112,6 +114,7 @@ def Lemma.reorderForallInstDep (lem : Lemma) : MetaM Lemma := do let proof ← Meta.mkLambdaFVars prec (← Meta.mkLambdaFVars trail proof) let type ← Meta.mkForallFVars prec (← Meta.mkForallFVars trail body) return ⟨⟨proof, type, lem.deriv⟩, lem.params⟩ +-/ /-- Rewrite using a universe-monomorphic rigid equality diff --git a/Auto/Translation/Monomorphization.lean b/Auto/Translation/Monomorphization.lean index ee9c6c9..b204b50 100644 --- a/Auto/Translation/Monomorphization.lean +++ b/Auto/Translation/Monomorphization.lean @@ -17,6 +17,7 @@ initialize registerTraceClass `auto.mono.printLemmaInst registerTraceClass `auto.mono.printConstInst registerTraceClass `auto.mono.printResult + registerTraceClass `auto.mono.printInputLemmas register_option auto.mono.saturationThreshold : Nat := { defValue := 250 @@ -903,6 +904,8 @@ def intromono (lemmas : Array Lemma) (mvarId : MVarId) : MetaM MVarId := do return mvar.mvarId!) def monomorphize (lemmas : Array Lemma) (inhFacts : Array Lemma) (k : Reif.State → MetaM α) : MetaM α := do + for h in lemmas do + trace[auto.mono.printInputLemmas] "Monomorphization got input lemma :: {h.type}" let (inductiveVals, monoSt) ← monoMAction.run {} MetaState.runWithIntroducedFVars (metaStateMAction inductiveVals monoSt) k where diff --git a/Test/SmtTranslation/Trigger.lean b/Test/SmtTranslation/Trigger.lean index 8999f07..45cb596 100644 --- a/Test/SmtTranslation/Trigger.lean +++ b/Test/SmtTranslation/Trigger.lean @@ -17,3 +17,8 @@ axiom fGreater : forall x, trigger (f x) (f x > x) set_option trace.auto.lamFOL2SMT.nameSuggestion true theorem fPlusOneGreater : forall x, (f x) + 1 > x := by auto [fGreater] u[] + +axiom fTrueGreater : true → ∀ x, trigger (f x) (f x > x) + +set_option trace.auto.printLemmas true +theorem fTrueGreaterTr : true → forall x, f x > x := by auto [fTrueGreater] u[] From 4c4bcdd3d52cd1d850389f45986825ac0d0c4324 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Sat, 23 Nov 2024 17:57:34 -0800 Subject: [PATCH 04/25] update readme --- README.md | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/README.md b/README.md index f3e69fa..89e851f 100644 --- a/README.md +++ b/README.md @@ -4,7 +4,17 @@ Lean-auto is an interface between Lean and automated theorem provers. Up to now, lean-auto is maintained and developed primarily by Yicheng Qian (GitHub: PratherConid). It is currently in active development, and pull requests/issues are welcome. For more information, feel free to reach out to Yicheng Qian on [Lean Zulip](https://leanprover.zulipchat.com). -Lean-auto is based on a monomorphization procedure from dependent type theory to higher-order logic and a deep embedding of higher-order logic into dependent type theory. It is capable of handling dependently-typed and/or universe-polymorphic input terms. Currently, proof reconstruction can be handled by [Duper](https://github.com/leanprover-community/duper), a higher-order superposition prover written in Lean. +Lean-auto is based on a monomorphization procedure from dependent type theory to higher-order logic and a deep embedding of higher-order logic into dependent type theory. It is capable of handling dependently-typed and/or universe-polymorphic input terms. Currently, proof reconstruction can be handled by [Duper](https://github.com/leanprover-community/duper), a higher-order superposition prover written in Lean. To enable Duper, please import the duper repo in your project, and set the following options: +```lean +open Lean Auto in +def Auto.duperRaw (lemmas : Array Lemma) (_ : Array Lemma) : MetaM Expr := do + let lemmas : Array (Expr × Expr × Array Name × Bool) ← lemmas.mapM + (fun ⟨⟨proof, ty, _⟩, _⟩ => do return (ty, ← Meta.mkAppM ``eq_true #[proof], #[], true)) + runDuper lemmas.data 0 + +attribute [rebind Auto.Native.solverFunc] Auto.duperRaw +set_option auto.native true +``` Although Lean-auto is still under development, it's already able to solve nontrivial problems. For example the first part of the "snake lemma" in category theory can be solved by a direct invocation to ``auto`` (and the second part can also be partly automated): From 52f3d98f1741982d19e6c506d26e3d8943fff85e Mon Sep 17 00:00:00 2001 From: PratherConid Date: Sat, 23 Nov 2024 19:49:41 -0800 Subject: [PATCH 05/25] update to v4.13.0 --- Auto/Embedding/Lam2Base.lean | 12 +- Auto/Embedding/LamBase.lean | 2 +- Auto/Embedding/LamBitVec.lean | 12 +- Auto/Embedding/LamChecker.lean | 14 +- Auto/Embedding/LamConv.lean | 4 +- Auto/Embedding/LamInhReasoning.lean | 8 +- Auto/Embedding/LamSystem.lean | 8 +- Auto/Embedding/LamTermInterp.lean | 14 +- Auto/Embedding/Lift.lean | 4 +- Auto/IR/SMT.lean | 44 +++--- Auto/LemDB.lean | 20 +-- Auto/Lib/AbstractMVars.lean | 16 +- Auto/Lib/BinTree.lean | 6 +- Auto/Lib/Containers.lean | 4 +- Auto/Lib/DeCompile.lean | 14 +- Auto/Lib/ExprExtra.lean | 4 +- Auto/Lib/LevelExtra.lean | 6 +- Auto/Lib/MessageData.lean | 4 +- Auto/Lib/MonadUtils.lean | 6 +- Auto/Lib/Pos.lean | 6 +- Auto/Parser/LeanLex.lean | 72 ++++----- Auto/Parser/NDFA.lean | 132 ++++++++--------- Auto/Parser/SMTSexp.lean | 2 +- Auto/Parser/TPTP.lean | 20 +-- Auto/Translation/Assumptions.lean | 16 +- Auto/Translation/Inductive.lean | 4 +- Auto/Translation/Inhabitation.lean | 4 +- Auto/Translation/Lam2DAtomAsFVar.lean | 6 +- Auto/Translation/Lam2TH0.lean | 2 +- Auto/Translation/LamFOL2SMT.lean | 18 +-- Auto/Translation/LamReif.lean | 195 +++++++++++++------------ Auto/Translation/LamUtils.lean | 110 +++++++------- Auto/Translation/Monomorphization.lean | 50 +++---- Auto/Translation/Preprocessing.lean | 16 +- Auto/Translation/Reduction.lean | 8 +- Auto/Translation/ReifM.lean | 8 +- README.md | 3 + TODO.md | 1 - Test/SmtTranslation/Inductive.lean | 2 - Test/SmtTranslation/String.lean | 1 + Test/Test_Regression.lean | 2 +- lean-toolchain | 2 +- 42 files changed, 442 insertions(+), 440 deletions(-) diff --git a/Auto/Embedding/Lam2Base.lean b/Auto/Embedding/Lam2Base.lean index 27de546..5ca43ac 100644 --- a/Auto/Embedding/Lam2Base.lean +++ b/Auto/Embedding/Lam2Base.lean @@ -121,9 +121,9 @@ def Lam₂Type.check_iff_interp simp [check, interp] cases Nat.decLt n (List.length lctx) case isTrue h => - simp [h]; simp [List.get?_eq_get h] + simp [h] case isFalse h => - simp [h]; simp at h; simp [List.get?_len_le h] + simp [h, List.getElem?_eq] case func fn arg IHfn IHarg => revert IHfn IHarg simp [check, interp] @@ -134,9 +134,9 @@ def Lam₂Type.check_iff_interp simp; match cifn : interp val lctx fn, ciarg : interp val lctx arg with | .some ⟨0, _⟩, .some ⟨0, _⟩ => simp - | .some ⟨0, _⟩, .some ⟨n + 1, _⟩ => simp; simp_arith + | .some ⟨0, _⟩, .some ⟨n + 1, _⟩ => simp | .some ⟨0, _⟩, .none => simp - | .some ⟨n + 1, _⟩, _ => simp; simp_arith + | .some ⟨n + 1, _⟩, _ => simp | .none , _ => simp | .some 0, .some (n + 1) => simp; @@ -168,7 +168,7 @@ def Lam₂Type.check_iff_interp simp; match cifn : interp val lctx fn, ciarg : interp val lctx arg with | .some ⟨n + 1, _⟩, .some ⟨0, _⟩ => simp - | .some ⟨n + 1, _⟩, .some ⟨m + 1, _⟩ => simp; simp_arith + | .some ⟨n + 1, _⟩, .some ⟨m + 1, _⟩ => simp | .some ⟨n + 1, _⟩, .none => simp | .some ⟨0, _⟩, _ => simp | .none , _ => simp @@ -185,7 +185,7 @@ def Lam₂Type.check_iff_interp | .some 0, _ => simp; match cifn : interp val lctx fn with - | .some ⟨n + 1, _⟩ => simp; simp_arith + | .some ⟨n + 1, _⟩ => simp | .some ⟨0, _⟩ => simp | .none => simp | .none, _ => diff --git a/Auto/Embedding/LamBase.lean b/Auto/Embedding/LamBase.lean index 4e72c09..55b7fc3 100644 --- a/Auto/Embedding/LamBase.lean +++ b/Auto/Embedding/LamBase.lean @@ -3780,7 +3780,7 @@ def LamWF.bvarApps rw [List.length_reverse]; apply Nat.le_refl dsimp [lctxr]; rw [← List.reverseAux, List.reverseAux_eq] rw [pushLCtxs_lt (by rw [List.length_append]; apply Nat.le_trans exlt (Nat.le_add_right _ _))] - rw [List.getD_eq_getElem?_getD]; rw [List.getElem?_append exlt]; + rw [List.getD_eq_getElem?_getD]; rw [List.getElem?_append_left exlt]; rw [List.getElem?_reverse (by dsimp [List.length]; apply Nat.le_refl _)] dsimp [List.length]; simp conv => enter [2, 3]; rw [tyeq] diff --git a/Auto/Embedding/LamBitVec.lean b/Auto/Embedding/LamBitVec.lean index d64048d..81a6cb8 100644 --- a/Auto/Embedding/LamBitVec.lean +++ b/Auto/Embedding/LamBitVec.lean @@ -55,10 +55,10 @@ namespace BVLems unfold BitVec.ushiftRight BitVec.toNat dsimp; rw [Nat.shiftRight_eq_div_pow] - theorem toNat_zeroExtend' {a : BitVec n} (le : n ≤ m) : (a.zeroExtend' le).toNat = a.toNat := rfl + theorem toNat_zeroExtend' {a : BitVec n} (le : n ≤ m) : (a.setWidth' le).toNat = a.toNat := rfl theorem toNat_zeroExtend {a : BitVec n} (i : Nat) : (a.zeroExtend i).toNat = a.toNat % (2 ^ i) := by - unfold BitVec.zeroExtend; cases hdec : decide (n ≤ i) + unfold BitVec.zeroExtend BitVec.setWidth; cases hdec : decide (n ≤ i) case false => have hnle := of_decide_eq_false hdec rw [Bool.dite_eq_false (proof:=hnle)]; rfl @@ -101,9 +101,9 @@ namespace BVLems apply Nat.le_trans (toNat_le _) (Nat.pow_le_pow_of_le_right (.step .refl) h) theorem msb_equiv_lt (a : BitVec n) : !a.msb ↔ a.toNat < 2 ^ (n - 1) := by - dsimp [BitVec.msb, BitVec.getMsb, BitVec.getLsb] + dsimp [BitVec.msb, BitVec.getMsbD, BitVec.getLsbD] cases n - case zero => simp + case zero => cases a <;> simp case succ n => have dtrue : decide (0 < n + 1) = true := by simp rw [dtrue, Bool.not_eq_true', Bool.true_and, Nat.succ_sub_one, Nat.testBit_false_iff] @@ -112,7 +112,7 @@ namespace BVLems theorem msb_equiv_lt' (a : BitVec n) : !a.msb ↔ 2 * a.toNat < 2 ^ n := by rw [msb_equiv_lt] cases n - case zero => simp + case zero => cases a <;> simp case succ n => rw [Nat.succ_sub_one, Nat.pow_succ, Nat.mul_comm (m:=2)] apply Iff.symm; apply Nat.mul_lt_mul_left @@ -133,7 +133,7 @@ namespace BVLems rw [Nat.sub_one, Nat.pred_lt_iff_le (Nat.two_pow_pos _)] apply Nat.le_trans (Nat.sub_le _ _) (Nat.pow_le_pow_of_le_right (.step .refl) h) apply eq_of_val_eq; rw [toNat_ofNatLt, hzero] - rw [toNat_neg, Int.mod_def', Int.emod]; dsimp only; + rw [toNat_neg, Int.mod_def', Int.emod] rw [Nat.zero_mod, Int.natAbs_ofNat, Nat.succ_eq_add_one, Nat.zero_add] rw [Int.subNatNat_of_sub_eq_zero ((Nat.sub_eq_zero_iff_le).mpr (Nat.two_pow_pos _))] rw [Int.toNat_ofNat, BitVec.toNat_ofNat] diff --git a/Auto/Embedding/LamChecker.lean b/Auto/Embedding/LamChecker.lean index 0fdbd66..bf12dce 100644 --- a/Auto/Embedding/LamChecker.lean +++ b/Auto/Embedding/LamChecker.lean @@ -2460,9 +2460,9 @@ theorem Checker.getValidExport_indirectReduce_reflection structure RTableStatus where rTable : Array REntry := #[] rTableTree : BinTree REntry := BinTree.leaf - nonemptyMap : HashMap LamSort Nat := {} - wfMap : HashMap (List LamSort × LamSort × LamTerm) Nat := {} - validMap : HashMap (List LamSort × LamTerm) Nat := {} + nonemptyMap : Std.HashMap LamSort Nat := {} + wfMap : Std.HashMap (List LamSort × LamSort × LamTerm) Nat := {} + validMap : Std.HashMap (List LamSort × LamTerm) Nat := {} -- maxEVarSucc maxEVarSucc : Nat := 0 -- lamEVarTy @@ -2470,7 +2470,7 @@ structure RTableStatus where -- This works as a cache for `BinTree.ofListGet lamEVarTy.data` lamEVarTyTree : BinTree LamSort := BinTree.leaf -- `chkMap.find?[re]` returns the checkstep which proves `re` - chkMap : HashMap REntry ChkStep := {} + chkMap : Std.HashMap REntry ChkStep := {} deriving Inhabited def RTableStatus.push (rs : RTableStatus) (re : REntry) := @@ -2494,8 +2494,8 @@ def RTableStatus.newEtomWithValid (rs : RTableStatus) (c : ChkStep) (re : REntry def RTableStatus.findPos? (rs : RTableStatus) (re : REntry) := match re with - | .nonempty s => rs.nonemptyMap.find? s - | .wf lctx s t => rs.wfMap.find? (lctx, s, t) - | .valid lctx t => rs.validMap.find? (lctx, t) + | .nonempty s => rs.nonemptyMap.get? s + | .wf lctx s t => rs.wfMap.get? (lctx, s, t) + | .valid lctx t => rs.validMap.get? (lctx, t) end Auto.Embedding.Lam diff --git a/Auto/Embedding/LamConv.lean b/Auto/Embedding/LamConv.lean index 2c78aa2..950a022 100644 --- a/Auto/Embedding/LamConv.lean +++ b/Auto/Embedding/LamConv.lean @@ -1957,8 +1957,8 @@ section UnsafeOps /-- Turn `ts[i]` into `.bvar i` -/ def LamTerm.abstractsImp (t : LamTerm) (ts : Array LamTerm) := let ts := ts.mapIdx (fun i x => (x, LamTerm.bvar i.val)) - let tmap := @Lean.HashMap.ofList _ _ inferInstance inferInstance ts.data - t.replace (fun x => tmap.find? x) 0 + let tmap := @Std.HashMap.ofList _ _ inferInstance inferInstance ts.toList + t.replace (fun x => tmap.get? x) 0 def LamTerm.abstractsRevImp (t : LamTerm) (ts : Array LamTerm) := t.abstractsImp ts.reverse diff --git a/Auto/Embedding/LamInhReasoning.lean b/Auto/Embedding/LamInhReasoning.lean index e3ec8cf..242219d 100644 --- a/Auto/Embedding/LamInhReasoning.lean +++ b/Auto/Embedding/LamInhReasoning.lean @@ -10,18 +10,18 @@ namespace Auto.Embedding.Lam def Inhabitation.subsumeQuick (s₁ s₂ : LamSort) : Bool := let s₁args := s₁.getArgTys let s₁res := s₁.getResTy - let s₂args := HashSet.empty.insertMany s₂.getArgTys + let s₂args := Std.HashSet.empty.insertMany s₂.getArgTys let s₂res := s₂.getResTy s₂args.contains s₂res || (s₁res == s₂res && ( s₁args.all (fun arg => s₂args.contains arg) )) /-- Run a quick test on whether the inhabitation of `s` is trivial -/ -def Inhabitation.trivialQuick := go HashSet.empty -where go (argTys : HashSet LamSort) (s : LamSort) : Bool := +def Inhabitation.trivialQuick := go Std.HashSet.empty +where go (argTys : Std.HashSet LamSort) (s : LamSort) : Bool := match s with | .func argTy resTy => argTys.contains s || go (argTys.insert argTy) resTy | _ => argTys.contains s -end Auto.Embedding.Lam \ No newline at end of file +end Auto.Embedding.Lam diff --git a/Auto/Embedding/LamSystem.lean b/Auto/Embedding/LamSystem.lean index 1f9a6b8..195d41b 100644 --- a/Auto/Embedding/LamSystem.lean +++ b/Auto/Embedding/LamSystem.lean @@ -209,7 +209,7 @@ theorem LamTerm.rwGenAllWith_lam : rwGenAllWith conv rty (.lam s body) = | .none => match rty with | .func _ resTy => (rwGenAllWith conv resTy body).bind (LamTerm.lam s ·) - | _ => .none := by simp [rwGenAllWith] + | _ => .none := by cases rty <;> simp [rwGenAllWith] theorem LamTerm.rwGenAllWith_app : rwGenAllWith conv rty (.app s fn arg) = match conv rty (.app s fn arg) with @@ -1492,7 +1492,7 @@ theorem LamTerm.evarBounded_rwGenAllWith (H : ∀ s, evarBounded (conv s) bound) case none.refl => apply Nat.le_max_right case some.refl => apply H _ _ _ h case lam s' body IH => - simp [LamTerm.rwGenAllWith] + rw [LamTerm.rwGenAllWith_lam] match h₁ : conv s (.lam s' body) with | .some t' => intro h₂; cases h₂; apply H _ _ _ h₁ | .none => @@ -1547,7 +1547,7 @@ theorem LamTerm.evarEquiv_rwGenAllWith (H : ∀ s, evarEquiv (conv s)) : case none.refl => rfl case some.refl => apply H _ _ _ h case lam s' body IH => - simp [LamTerm.rwGenAllWith] + rw [LamTerm.rwGenAllWith_lam] match h₁ : conv s (.lam s' body) with | .some t' => intro h₂; cases h₂; apply H _ _ _ h₁ | .none => @@ -1592,7 +1592,7 @@ theorem LamGenConvWith.rwGenAllWith (H : LamGenConvWith lval conv) : LamGenConvW case none.refl => apply LamGenEquivWith.refl case some.refl => apply H _ _ _ h case lam s' body IH => - simp [LamTerm.rwGenAllWith] + rw [LamTerm.rwGenAllWith_lam] match h₁ : conv s (.lam s' body) with | .some t' => intro h₂; cases h₂; apply H _ _ _ h₁ | .none => diff --git a/Auto/Embedding/LamTermInterp.lean b/Auto/Embedding/LamTermInterp.lean index 93c18f2..69de7de 100644 --- a/Auto/Embedding/LamTermInterp.lean +++ b/Auto/Embedding/LamTermInterp.lean @@ -164,7 +164,7 @@ theorem LamTerm.lamCheck?Eq'_bvar (h : lctx.get? n = .some ⟨s, val⟩) : LamTerm.lamCheck?Eq' lval lctx (.bvar n) s := by dsimp [lamCheck?Eq', lamCheck?]; have ⟨hlt, _⟩ := List.get?_eq_some.mp h - rw [pushLCtxs_lt hlt, List.getD_eq_get?, h]; rfl + rw [pushLCtxs_lt hlt, List.getD_eq_getElem?_getD, ← List.get?_eq_getElem?, h]; rfl theorem LamTerm.lamCheck?Eq'_lam (h : LamTerm.lamCheck?Eq' lval (⟨argTy, val⟩ :: lctx) body s) : @@ -201,7 +201,7 @@ theorem LamTerm.interpEq_bvar (s : LamSort) (val : s.interp lval.tyVal) (h : lctx.get? n = .some ⟨s, val⟩) : LamTerm.interpEq lval lctx (.bvar n) val := by dsimp [interpEq, interp]; have ⟨hlt, _⟩ := List.get?_eq_some.mp h - rw [pushLCtxs_lt hlt, List.getD_eq_get?, h]; rfl + rw [pushLCtxs_lt hlt, List.getD_eq_getElem?_getD, ← List.get?_eq_getElem?, h]; rfl theorem LamTerm.interpEq_lam (lval : LamValuation) (lctx : List ((s : LamSort) × s.interp lval.tyVal)) @@ -243,7 +243,7 @@ namespace Interp structure State where sortFVars : Array FVarId := #[] - sortMap : HashMap LamSort FVarId := {} + sortMap : Std.HashMap LamSort FVarId := {} -- Let `n := lctxTyRev.size` -- Reversed lctxTyRev : Array LamSort := #[] @@ -253,11 +253,11 @@ structure State where -- Required : `lctxTyDrop[i] ≝ lctxTyRev[:(i+1)].data ≝ lctxTerm[(n-i-1):]` lctxTyDrop : Array FVarId := #[] -- Required : `tyEqFact[i][j] : lctxTy[i:].drop j = lctxTy[i+j:]` - typeEqFact : HashMap (Nat × Nat) FVarId := {} + typeEqFact : Std.HashMap (Nat × Nat) FVarId := {} -- Required : `lctxTermDrop[i] ≝ lctxTermRev[:(i+1)].data ≝ lctxTerm[(n-i-1):]` lctxTermDrop : Array FVarId := #[] -- Required : `termEqFact[i][j] : lctxTerm[i:].drop j = lctxTerm[i+j:]` - termEqFact : HashMap (Nat × Nat) FVarId := {} + termEqFact : Std.HashMap (Nat × Nat) FVarId := {} -- Required : `lctxCon[i] : lctxTerm[i].map Sigma.snd = lctxTy[i]` lctxCon : Array FVarId := #[] @@ -282,7 +282,7 @@ def getLCtxTy! (idx : Nat) : InterpM LamSort := do def sort2FVarId (s : LamSort) : InterpM FVarId := do let sortMap ← getSortMap let userName := (`interpsf).appendIndexAfter (← getSortMap).size - match sortMap.find? s with + match sortMap.get? s with | .some id => return id | .none => match s with @@ -336,4 +336,4 @@ where withLCtxTy {α : Type} (s : LamSort) (k : InterpM α) : InterpM α := do end Interp -end Auto.Embedding.Lam \ No newline at end of file +end Auto.Embedding.Lam diff --git a/Auto/Embedding/Lift.lean b/Auto/Embedding/Lift.lean index 8c7db2e..b0695b1 100644 --- a/Auto/Embedding/Lift.lean +++ b/Auto/Embedding/Lift.lean @@ -85,10 +85,10 @@ def imulLift.{u} (m n : GLift.{1, u} Int) := GLift.up (Int.mul m.down n.down) def idivLift.{u} (m n : GLift.{1, u} Int) := - GLift.up (Int.div m.down n.down) + GLift.up (Int.tdiv m.down n.down) def imodLift.{u} (m n : GLift.{1, u} Int) := - GLift.up (Int.mod m.down n.down) + GLift.up (Int.tmod m.down n.down) def iedivLift.{u} (m n : GLift.{1, u} Int) := GLift.up (Int.ediv m.down n.down) diff --git a/Auto/IR/SMT.lean b/Auto/IR/SMT.lean index 41f8aac..1384e37 100644 --- a/Auto/IR/SMT.lean +++ b/Auto/IR/SMT.lean @@ -33,7 +33,7 @@ deriving BEq, Hashable, Inhabited def SIdent.toString : SIdent → String | .symb s => if isSimpleSymbol s then s else "|" ++ s ++ "|" | .indexed s idx => - s!"(_ {s} " ++ String.intercalate " " (idx.data.map (fun idx => + s!"(_ {s} " ++ String.intercalate " " (idx.toList.map (fun idx => match idx with | .inl idx => s!"{idx}" | .inr idx => s!"{idx}")) ++ ")" @@ -63,7 +63,7 @@ where go : List SSort → List SIdent → List String | a :: as, binders => SSort.toStringAux a binders :: go as binders def SSort.toString (s : SSort) (binders : Array SIdent) : String := - SSort.toStringAux s binders.data + SSort.toStringAux s binders.toList /-- Caution : Do not use this in define-sort, because sort there might contain bvars -/ instance : ToString SSort where @@ -181,7 +181,7 @@ private partial def STerm.toStringAux : STerm → List SIdent → String intro ++ body | .attr t attrs, binders => let intro := "(! " ++ STerm.toStringAux t binders ++ " " - let sattrs := String.intercalate " " (attrs.data.map (attrToStringAux · binders)) + let sattrs := String.intercalate " " (attrs.toList.map (attrToStringAux · binders)) intro ++ sattrs ++ ")" where goQIdApp : List STerm → List SIdent → List String @@ -194,9 +194,9 @@ where let pattern := "(" ++ (ToString.toString (SIdent.symb constr)) pattern ++ body else - let binders := args.data.map .symb ++ binders + let binders := args.toList.map .symb ++ binders let body := " " ++ STerm.toStringAux body binders ++ ")" - let args := args.data.map (fun x => ToString.toString (SIdent.symb x)) + let args := args.toList.map (fun x => ToString.toString (SIdent.symb x)) let pattern := "((" ++ String.intercalate " " (ToString.toString (SIdent.symb constr) :: args) ++ ")" pattern ++ body goMatchBody : List (MatchCase STerm) → List SIdent → List String @@ -206,16 +206,16 @@ where | .none s, _ => ":" ++ s | .spec s sc, _ => s!":{s} {sc.toString}" | .symb s s', _ => s!":{s} {s'}" - | .sexpr s ts, binders => s!":{s} (" ++ String.intercalate " " (ts.data.map (STerm.toStringAux · binders)) ++ ")" + | .sexpr s ts, binders => s!":{s} (" ++ String.intercalate " " (ts.toList.map (STerm.toStringAux · binders)) ++ ")" def STerm.toString (t : STerm) (binders : Array SIdent) : String := - STerm.toStringAux t binders.data + STerm.toStringAux t binders.toList instance : ToString STerm where toString t := STerm.toString t #[] def Attribute.toString (attr : Attribute) (binders : Array SIdent) : String := - SMT.STerm.toStringAux.attrToStringAux attr binders.data + SMT.STerm.toStringAux.attrToStringAux attr binders.toList instance : ToString Attribute where toString attr := Attribute.toString attr #[] @@ -232,7 +232,7 @@ private def ConstrDecl.toString : ConstrDecl → Array SIdent → String | ⟨name, selDecls⟩, binders => let pre := s!"({SIdent.symb name}" let selDecls := selDecls.map (fun (name, sort) => s!"({SIdent.symb name} " ++ SSort.toString sort binders ++ ")") - String.intercalate " " (pre :: selDecls.data) ++ ")" + String.intercalate " " (pre :: selDecls.toList) ++ ")" /-- 〈datatype_dec〉 ::= ( 〈constructor_dec〉+ ) | ( par ( 〈symbol 〉+ ) ( 〈constructor_dec〉+ ) ) @@ -243,11 +243,11 @@ structure DatatypeDecl where private def DatatypeDecl.toString : DatatypeDecl → String := fun ⟨params, cstrDecls⟩ => let scstrDecls := cstrDecls.map (fun d => ConstrDecl.toString d (params.map SIdent.symb)) - let scstrDecls := "(" ++ String.intercalate " " scstrDecls.data ++ ")" + let scstrDecls := "(" ++ String.intercalate " " scstrDecls.toList ++ ")" if params.size == 0 then scstrDecls else - "(par (" ++ String.intercalate " " params.data ++ ") " ++ scstrDecls ++ ")" + "(par (" ++ String.intercalate " " params.toList ++ ") " ++ scstrDecls ++ ")" inductive SMTOption where | diagnosticOC : String → SMTOption @@ -344,26 +344,26 @@ def Command.toString : Command → String | .checkSat => "(check-sat)" | .declFun name argSorts resSort => let pre := s!"(declare-fun {SIdent.symb name} (" - let argSorts := String.intercalate " " (argSorts.map ToString.toString).data ++ ") " + let argSorts := String.intercalate " " (argSorts.map ToString.toString).toList ++ ") " let trail := s!"{resSort})" pre ++ argSorts ++ trail | .declSort name arity => s!"(declare-sort {SIdent.symb name} {arity})" | .defFun isRec name args resTy body => let pre := if isRec then "(define-fun-rec " else "(define-fun " let pre := pre ++ ToString.toString (SIdent.symb name) ++ " " - let binders := "(" ++ String.intercalate " " (args.map (fun (name, sort) => s!"({SIdent.symb name} {sort})")).data ++ ") " + let binders := "(" ++ String.intercalate " " (args.map (fun (name, sort) => s!"({SIdent.symb name} {sort})")).toList ++ ") " let trail := s!"{resTy} " ++ STerm.toString body (args.map (fun (name, _) => SIdent.symb name)) ++ ")" pre ++ binders ++ trail | .defSort name args body => let pre := s!"(define-sort {SIdent.symb name} (" - let sargs := String.intercalate " " args.data ++ ") " + let sargs := String.intercalate " " args.toList ++ ") " let trail := SSort.toString body (args.map SIdent.symb) ++ ")" pre ++ sargs ++ trail | .declDtype name ddecl => s!"(declare-datatype {SIdent.symb name} {ddecl.toString})" | .declDtypes infos => - let sort_decs := String.intercalate " " (infos.data.map (fun (name, args, _) => s!"({SIdent.symb name} {args})")) - let datatype_decs := String.intercalate " " (infos.data.map (fun (_, _, ddecl) => ddecl.toString)) + let sort_decs := String.intercalate " " (infos.toList.map (fun (name, args, _) => s!"({SIdent.symb name} {args})")) + let datatype_decs := String.intercalate " " (infos.toList.map (fun (_, _, ddecl) => ddecl.toString)) s!"(declare-datatypes ({sort_decs}) ({datatype_decs}))" | .exit => "(exit)" @@ -383,10 +383,10 @@ section -/ structure State where -- Map from high-level construct to symbol - h2lMap : HashMap ω String := {} + h2lMap : Std.HashMap ω String := {} -- Inverse of `h2lMap` -- Map from symbol to high-level construct - l2hMap : HashMap String ω := {} + l2hMap : Std.HashMap String ω := {} -- We allow two types of names -- · `"_" ++ s`, -- · `"_" ++ s ++ "_" ++ ` @@ -400,7 +400,7 @@ section -- been used, we return `n'` as the final name. If `n'` has -- been used for `k` times (`k > 0`), return `n' ++ s!"_{k - 1}"`. -- `usedNames` records the `k - 1` for each `n'` - usedNames : HashMap String Nat := {} + usedNames : Std.HashMap String Nat := {} -- List of commands commands : Array Command := #[] @@ -441,7 +441,7 @@ section preName := "pl_" ++ preName if preName.back.isDigit then preName := preName ++ "_" - if let .some idx := (← getUsedNames).find? preName then + if let .some idx := (← getUsedNames).get? preName then -- Used setUsedNames ((← getUsedNames).insert preName (idx + 1)) return "_" ++ preName ++ s!"_{idx}" @@ -455,7 +455,7 @@ section "~!@$%^&*_-+=<>.?/" ++ "ΑαΒβΓγΔδΕεΖζΗηΘθΙιΚκΛλΜμΝνΞξΟοΠπΡρΣσςΤτΥυΦφΧχΨψΩω" ++ "₀₁₂₃₄₅₆₇₈₉" - let allowedSet : HashSet UInt32 := HashSet.insertMany HashSet.empty (List.map Char.val allowedStr.toList) + let allowedSet : Std.HashSet UInt32 := Std.HashSet.insertMany Std.HashSet.empty (List.map Char.val allowedStr.toList) c.isAlphanum || allowedSet.contains c.val @@ -469,7 +469,7 @@ section partial def h2Symb (cstr : ω) (nameSuggestion : Option String) : TransM ω String := do let l2hMap ← getL2hMap let h2lMap ← getH2lMap - if let .some name := h2lMap.find? cstr then + if let .some name := h2lMap.get? cstr then return name let .some nameSuggestion := nameSuggestion | throwError "IR.SMT.h2Symb :: Fresh high-level constraint {cstr} without name suggestion" diff --git a/Auto/LemDB.lean b/Auto/LemDB.lean index 6fb183b..9ab8ad6 100644 --- a/Auto/LemDB.lean +++ b/Auto/LemDB.lean @@ -20,7 +20,7 @@ inductive LemDB where deriving BEq, Hashable, Inhabited, Repr abbrev LemDBExtension := - PersistentEnvExtension (Name × LemDB) (Name × LemDB) (HashMap Name LemDB) + PersistentEnvExtension (Name × LemDB) (Name × LemDB) (Std.HashMap Name LemDB) initialize lemDBExt : LemDBExtension ← registerPersistentEnvExtension { name := `LemDBExt @@ -28,20 +28,20 @@ initialize lemDBExt : LemDBExtension ← registerPersistentEnvExtension { addEntryFn := fun s n => s.insert n.1 n.2 -- **Note** We suppose that, if module `a` imports module `b`, -- then the index of `a` within the `arr` is greater than the index of `b` in `arr` - addImportedFn := fun arr => pure <| HashMap.ofList (arr.concatMap id).toList, + addImportedFn := fun arr => pure <| Std.HashMap.ofList (arr.concatMap id).toList, exportEntriesFn := fun s => s.toArray } -partial def LemDB.toHashSet : LemDB → AttrM (HashSet Name) -| .empty => pure HashSet.empty +partial def LemDB.toHashSet : LemDB → AttrM (Std.HashSet Name) +| .empty => pure Std.HashSet.empty | .addLemma lem hdb => do let hset ← hdb.toHashSet return hset.insert lem | .compose hdbs => do let state := lemDBExt.getState (← getEnv) - let mut ret := HashSet.empty + let mut ret := Std.HashSet.empty for hdb in hdbs do - let some hdb := state.find? hdb + let some hdb := state.get? hdb | throwError "LemDB.toHashSet :: Unknown lemma database {hdb}" let hset ← hdb.toHashSet ret := ret.insertMany hset @@ -60,7 +60,7 @@ def registerAddLemToDB : IO Unit := add := fun decl stx _ => do let dbname := (← Attribute.Builtin.getIdent stx).getId let state := lemDBExt.getState (← getEnv) - if let some db := state.find? dbname then + if let some db := state.get? dbname then let state' := state.insert dbname (.addLemma decl db) modifyEnv fun env => lemDBExt.modifyState env fun _ => state' else @@ -74,7 +74,7 @@ initialize registerAddLemToDB def findLemDB (dbname : Name) : CoreM (Option LemDB) := do let dbname := dbname let state := lemDBExt.getState (← getEnv) - if let some db := state.find? dbname then + if let some db := state.get? dbname then return .some db else return .none @@ -91,7 +91,7 @@ def elabdeclarelemdb : CommandElab := fun stx => do | `(declarelemdb | #declare_lemdb $dbname) => let dbname := dbname.getId let state := lemDBExt.getState (← getEnv) - if let some db := state.find? dbname then + if let some db := state.get? dbname then throwError "Lemma database {repr db} has already been declared" else let state' := state.insert dbname .empty @@ -118,7 +118,7 @@ def elabcomposelemdb : CommandElab := fun stx => do for db in dbs do if !state.contains db then throwError "Unknown lemma database {repr db}" - if let some db := state.find? db then + if let some db := state.get? db then throwError "Lemma database {repr db} has already been declared" else let state' := state.insert db (.compose dbs) diff --git a/Auto/Lib/AbstractMVars.lean b/Auto/Lib/AbstractMVars.lean index 0e5a3e6..a4f7a02 100644 --- a/Auto/Lib/AbstractMVars.lean +++ b/Auto/Lib/AbstractMVars.lean @@ -18,8 +18,8 @@ structure State where nextParamIdx : Nat := 0 paramNames : Array Name := #[] fvars : Array Expr := #[] - lmap : HashMap LMVarId Level := {} - emap : HashMap MVarId Expr := {} + lmap : Std.HashMap LMVarId Level := {} + emap : Std.HashMap MVarId Expr := {} abbrev M := StateM State @@ -58,7 +58,7 @@ private partial def abstractLevelMVars (u : Level) : M Level := do if l != lNew then abstractLevelMVars lNew else - match s.lmap.find? mvarId with + match s.lmap.get? mvarId with | some u => pure u | none => let paramId := Name.mkNum `_abstMVar s.nextParamIdx @@ -91,7 +91,7 @@ partial def abstractExprMVars (e : Expr) : M Expr := do if e != eNew then abstractExprMVars eNew else - match (← get).emap.find? mvarId with + match (← get).emap.get? mvarId with | some e => return e | none => @@ -155,20 +155,20 @@ def abstractMVarsLambdaWithIds (e : Expr) : MetaM (Expr × Array Expr × Array N let e := s.lctx.mkLambda s.fvars e -- Restore the corresponding expr mvarId let sfvars := s.fvars - let mut fvarpos : HashMap FVarId Nat := {} + let mut fvarpos : Std.HashMap FVarId Nat := {} for i in [:sfvars.size] do fvarpos := fvarpos.insert sfvars[i]!.fvarId! i let mut mvars := sfvars for (mid, fvar) in s.emap.toList do - mvars := mvars.set! (fvarpos.find! fvar.fvarId!) (mkMVar mid) + mvars := mvars.set! (fvarpos.get! fvar.fvarId!) (mkMVar mid) -- Restore the corresponding level mvarId let sparams := s.paramNames - let mut parampos : HashMap Name Nat := {} + let mut parampos : Std.HashMap Name Nat := {} for i in [:sparams.size] do parampos := parampos.insert sparams[i]! i let mut lmvars := sparams.map (Level.param) for (lmid, paramname) in s.lmap.toList do - lmvars := lmvars.set! (parampos.find! (getParamLevelName! paramname)) (mkLevelMVar lmid) + lmvars := lmvars.set! (parampos.get! (getParamLevelName! paramname)) (mkLevelMVar lmid) pure (e, mvars, sparams, lmvars) where getParamLevelName! : Level → Name | .param name => name diff --git a/Auto/Lib/BinTree.lean b/Auto/Lib/BinTree.lean index 030a42f..531d73d 100644 --- a/Auto/Lib/BinTree.lean +++ b/Auto/Lib/BinTree.lean @@ -260,7 +260,7 @@ theorem insert'Aux.equiv (bt : BinTree α) (n : Nat) (x : α) (rd : Nat) : case succ rd' IH => match n with | 0 => rw [insert'Aux, insert'WF] - | 1 => rw [insert'Aux, insert'WF] + | 1 => cases bt <;> rw [insert'Aux, insert'WF] | n' + 2 => dsimp [insert'Aux]; rw [insert'WF.succSucc]; @@ -313,7 +313,7 @@ theorem insert'.correct₁ (bt : BinTree β) (n : Nat) (x : β) : n ≠ 0 → ge have hne' : (n + 2) / 2 ≠ 0 := by rw [Nat.add_div_right _ (.step .refl)]; intro h; cases h let IH' := fun bt => IH bt hne' - rw [get?'_succSucc, insert'.succSucc, left!, right!] + rw [get?'_succSucc, insert'.succSucc, left!.eq_def, right!.eq_def] cases (n + 2) % 2 <;> cases bt <;> dsimp <;> rw [IH'] theorem insert.correct₁ (bt : BinTree β) (n : Nat) (x : β) : get? (insert bt n x) n = .some x := @@ -338,7 +338,7 @@ theorem insert'.correct₂ (bt : BinTree β) (n₁ n₂ : Nat) (x : β) : n₁ | 0 => cases bt <;> rfl | 1 => cases bt <;> rfl | n₁ + 2 => - rw [insert'.succSucc, left!, right!] + rw [insert'.succSucc, left!.eq_def, right!.eq_def] have hne' : (n₁ + 2) % 2 = (n₂ + 2) % 2 → (n₁ + 2) / 2 ≠ (n₂ + 2) / 2 := by intro heq h; apply hne; rw [← Nat.div_add_mod (n₁ + 2) 2, ← Nat.div_add_mod (n₂ + 2) 2] diff --git a/Auto/Lib/Containers.lean b/Auto/Lib/Containers.lean index 2f22643..ee5d660 100644 --- a/Auto/Lib/Containers.lean +++ b/Auto/Lib/Containers.lean @@ -3,7 +3,7 @@ open Lean namespace Auto -def mergeHashSet {α : Type u} [BEq α] [Hashable α] (a1 a2 : HashSet α) := +def mergeHashSet {α : Type u} [BEq α] [Hashable α] (a1 a2 : Std.HashSet α) := if a1.size < a2.size then a2.insertMany a1.toArray else @@ -22,4 +22,4 @@ class Container {α : Type u} (C : Type v → Type w) where insertCorrect₁ : ∀ (c : C β) (k : α) (v : β), get? (insert c k v) k = .some val insertCorrect₂ : ∀ (c : C β) (k₁ k₂ : α) (v : β), k₁ ≠ k₂ → get? (insert c k₁ v) k₂ = get? c k₂ -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Lib/DeCompile.lean b/Auto/Lib/DeCompile.lean index b89acea..4d79c70 100644 --- a/Auto/Lib/DeCompile.lean +++ b/Auto/Lib/DeCompile.lean @@ -13,7 +13,7 @@ namespace ExprDeCompile -- as explicit structure Context where - usedNames : HashSet String + usedNames : Std.HashSet String structure State where binderNames : Array String := #[] @@ -69,7 +69,7 @@ private partial def exprDeCompileAux (final : Bool) (e : Expr) : ExprDeCompM Str return ret let fs ← exprDeCompileAux final f let argss ← args.mapM (exprDeCompileAux final) - let argssC := String.intercalate " " argss.data + let argssC := String.intercalate " " argss.toList return s!"({fs} {argssC})" | Expr.proj _ idx b => let bs ← exprDeCompileAux final b @@ -94,7 +94,7 @@ private partial def exprDeCompileAux (final : Bool) (e : Expr) : ExprDeCompM Str | Expr.bvar n => let bn := (← get).binderNames return bn[bn.size - n - 1]! -where +where constFinalDeComp (name : Name) (us : List Level) := if List.length us == 0 then "@" ++ name.toString @@ -118,7 +118,7 @@ where e := .lam `_ (ty.instantiateLevelParams lparams cst.constLevels!) e .default return e -def levelCollectNames (l : Level) : StateT (HashSet String) MetaM Unit := do +def levelCollectNames (l : Level) : StateT (Std.HashSet String) MetaM Unit := do match l with | .zero => return | .succ l => levelCollectNames l @@ -127,7 +127,7 @@ def levelCollectNames (l : Level) : StateT (HashSet String) MetaM Unit := do | .max l1 l2 => levelCollectNames l1; levelCollectNames l2 | .imax l1 l2 => levelCollectNames l1; levelCollectNames l2 -def exprCollectNames (e : Expr) : StateT (HashSet String) MetaM Unit := do +def exprCollectNames (e : Expr) : StateT (Std.HashSet String) MetaM Unit := do match e with | Expr.forallE _ d b _ => exprCollectNames d; exprCollectNames b | Expr.lam _ d b _ => exprCollectNames d; exprCollectNames b @@ -143,7 +143,7 @@ def exprCollectNames (e : Expr) : StateT (HashSet String) MetaM Unit := do | Expr.lit _ => return | Expr.sort l => levelCollectNames l | Expr.mvar .. => return - | Expr.fvar f => + | Expr.fvar f => let some decl := (← getLCtx).find? f | throwError "Unknown free variable {e}" let name := decl.userName.toString @@ -212,4 +212,4 @@ universe u --/ -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Lib/ExprExtra.lean b/Auto/Lib/ExprExtra.lean index adf6730..a83f10b 100644 --- a/Auto/Lib/ExprExtra.lean +++ b/Auto/Lib/ExprExtra.lean @@ -46,7 +46,7 @@ def Expr.forallBinders (e : Expr) : Array (Name × Expr × BinderInfo) := | _ => [] Array.mk (aux e) -def Expr.collectRawM [Monad m] (p : Expr → m Bool) : Expr → m (HashSet Expr) +def Expr.collectRawM [Monad m] (p : Expr → m Bool) : Expr → m (Std.HashSet Expr) | e@(.forallE _ d b _) => do let hd ← collectRawM p d let hb ← collectRawM p b @@ -70,7 +70,7 @@ def Expr.collectRawM [Monad m] (p : Expr → m Bool) : Expr → m (HashSet Expr) | e@(.proj _ _ b) => do let hb ← collectRawM p b addp? p e hb -| e => addp? p e HashSet.empty +| e => addp? p e Std.HashSet.empty where addp? p e hs := do if ← p e then return hs.insert e diff --git a/Auto/Lib/LevelExtra.lean b/Auto/Lib/LevelExtra.lean index 58a8d0d..3be8e09 100644 --- a/Auto/Lib/LevelExtra.lean +++ b/Auto/Lib/LevelExtra.lean @@ -17,13 +17,13 @@ def Level.collectLevelMVars (l : Level) : MetaM (Array LMVarId) := do let hset := go l return hset.toArray where - go : Level → HashSet LMVarId + go : Level → Std.HashSet LMVarId | .zero => {} | .succ l => go l | .max l₁ l₂ => mergeHashSet (go l₁) (go l₂) | .imax l₁ l₂ => mergeHashSet (go l₁) (go l₂) | .param _ => {} - | .mvar id => HashSet.empty.insert id + | .mvar id => Std.HashSet.empty.insert id def Level.findParam? (p : Name → Bool) : Level → Option Name | .zero => .none @@ -36,4 +36,4 @@ def Level.findParam? (p : Name → Bool) : Level → Option Name | false => .none | .mvar _ => .none -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Lib/MessageData.lean b/Auto/Lib/MessageData.lean index a0d9324..f9b9fec 100644 --- a/Auto/Lib/MessageData.lean +++ b/Auto/Lib/MessageData.lean @@ -17,6 +17,6 @@ def MessageData.list (as : List α) (f : α → MessageData) : MessageData := .compose m!"[" (.compose (MessageData.intercalate m!", " (as.map f)) m!"]") def MessageData.array (as : Array α) (f : α → MessageData) : MessageData := - MessageData.list as.data f + MessageData.list as.toList f -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Lib/MonadUtils.lean b/Auto/Lib/MonadUtils.lean index 2f22fe2..ba33e99 100644 --- a/Auto/Lib/MonadUtils.lean +++ b/Auto/Lib/MonadUtils.lean @@ -62,9 +62,9 @@ def structureProjs (structTy : Expr) : CoreM (Name × Expr × Array (Name × Exp | throwError s!"structureProjs :: {structName} is not a structure" let structMkExpr := mkAppN (Expr.const structDotMk lvls) structTy.getAppArgs let tyArgs := structTy.getAppArgs - let nameMap : HashMap Name StructureFieldInfo := HashMap.ofList - (structInfo.fieldInfo.map (fun sfi => (sfi.fieldName, sfi))).data - let sorted := structInfo.fieldNames.map (fun name => nameMap.find! name) + let nameMap : Std.HashMap Name StructureFieldInfo := Std.HashMap.ofList + (structInfo.fieldInfo.map (fun sfi => (sfi.fieldName, sfi))).toList + let sorted := structInfo.fieldNames.map (fun name => nameMap.get! name) let fieldInfos := sorted.map (fun i => -- Field name, Projection function (i.fieldName, mkAppN (Expr.const i.projFn lvls) tyArgs)) diff --git a/Auto/Lib/Pos.lean b/Auto/Lib/Pos.lean index 91ba509..3eaacfe 100644 --- a/Auto/Lib/Pos.lean +++ b/Auto/Lib/Pos.lean @@ -105,7 +105,7 @@ theorem ofNat'WF.doubleSucc_xI (n : Nat) : theorem ofNat'WF_toNat' (p : Pos) : ofNat'WF (toNat' p) = p := by induction p - case xH => rfl + case xH => rw [toNat', ofNat'WF] case xO p' IH => dsimp [toNat']; rw [ofNat'WF.double_xO]; @@ -118,7 +118,7 @@ theorem ofNat'WF_toNat' (p : Pos) : ofNat'WF (toNat' p) = p := by theorem toNat'_ofNat'WF (n : Nat) : n ≠ 0 → toNat' (ofNat'WF n) = n := by revert n; apply ofNat'WF.induction case base₀ => intro H; contradiction - case base₁ => intro _; rfl + case base₁ => intro _; rw [ofNat'WF, toNat'] case ind => intro x IH _ have hne : (x + 2) / 2 ≠ 0 := by @@ -163,7 +163,7 @@ theorem ofNat'.equivAux (rd n : Nat) : rd ≥ n → ofNat'WF n = ofNat'RD rd n : induction rd generalizing n <;> intro H case zero => have hzero : n = 0 := Nat.eq_zero_of_le_zero H - rw [hzero]; rfl + rw [hzero]; rw [ofNat'WF, ofNat'RD] case succ rd' IH => dsimp [ofNat'RD] match n with diff --git a/Auto/Parser/LeanLex.lean b/Auto/Parser/LeanLex.lean index 59ec179..087d4b9 100644 --- a/Auto/Parser/LeanLex.lean +++ b/Auto/Parser/LeanLex.lean @@ -8,7 +8,7 @@ namespace Auto -- **TODO**: Parser for POSIX ERE namespace Regex -private def sort : List Nat → List Nat := +private def sort : List Nat → List Nat := have : DecidableRel Nat.le := fun (x y : Nat) => inferInstanceAs (Decidable (x <= y)) List.mergeSort Nat.le @@ -82,17 +82,17 @@ deriving BEq, Hashable, Inhabited def EREBracket.neg (b : EREBracket) : EREBracket := .minus b (.cc .all) -- **TODO**: Why does this need `partial`? -partial def EREBracket.toHashSet : EREBracket → HashSet Char - | .cc cty => HashSet.empty.insertMany (toString cty).toList - | .inStr s => HashSet.empty.insertMany s.toList +partial def EREBracket.toHashSet : EREBracket → Std.HashSet Char + | .cc cty => Std.HashSet.empty.insertMany (toString cty).toList + | .inStr s => Std.HashSet.empty.insertMany s.toList | .plus ⟨bl⟩ => go bl | .minus b1 b2 => let hb := b2.toHashSet let b1s := b1.toHashSet.toList - HashSet.empty.insertMany (b1s.filter (fun x => !hb.contains x)) + Std.HashSet.empty.insertMany (b1s.filter (fun x => !hb.contains x)) where - go : List EREBracket → HashSet Char - | [] => HashSet.empty + go : List EREBracket → Std.HashSet Char + | [] => Std.HashSet.empty | b :: bl => (go bl).insertMany (toHashSet b) def EREBracket.toString (e : EREBracket) := String.mk e.toHashSet.toList @@ -184,9 +184,9 @@ section structure CharGrouping where ngroup : Nat -- All relevant characters - all : HashSet σ + all : Std.HashSet σ -- Map from character to its corresponding group - charMap : HashMap σ Nat + charMap : Std.HashMap σ Nat -- A character is in `all` iff it's in `charMap`. -- Group number takes value in `0, 1, ..., ngroups - 1` -- For the intermediate `NFA` generated in `ERE.toADFA`, @@ -213,21 +213,21 @@ section -- Refer to `ERE.toADFA`, `ERE.ADFALex` and `DFA.run` cg : CharGrouping σ deriving Inhabited - + variable {σ : Type} [Hashable σ] [BEq σ] [ToString σ] - + def CharGrouping.wf : CharGrouping σ → Bool := fun ⟨ngroup, all, charMap⟩ => - let img := charMap.fold (fun hs _ n => hs.insert n) HashSet.empty + let img := charMap.fold (fun hs _ n => hs.insert n) Std.HashSet.empty let surj := (sort img.toList) == List.range ngroup let allInCharMap := all.toList.all (fun c => charMap.contains c) let sizeEq := all.size == charMap.size surj && allInCharMap && sizeEq - - def CharGrouping.groups : CharGrouping σ → Array (HashSet σ) := + + def CharGrouping.groups : CharGrouping σ → Array (Std.HashSet σ) := fun ⟨ngroup, _, charMap⟩ => Id.run <| do - let mut arr : Array (HashSet σ) := - Array.mk ((List.range ngroup).map (fun _ => HashSet.empty)) + let mut arr : Array (Std.HashSet σ) := + Array.mk ((List.range ngroup).map (fun _ => Std.HashSet.empty)) for (c, idx) in charMap.toList do arr := arr.modify idx (fun hs => hs.insert c) return arr @@ -244,7 +244,7 @@ section s!"Group representing beginning of string := {ngroup}" :: s!"Group representing end of string := {ngroup + 1}" :: s!"Group representing other utf-8 characters := {ngroup + 2}" :: - groups.data + groups.toList String.intercalate "\n " all ++ "\n⦘⦘" def CharGrouping.toString (cg : CharGrouping σ) : String := @@ -254,7 +254,7 @@ section toString := CharGrouping.toString def CharGrouping.getGroup (cg : CharGrouping σ) (c : σ) : Nat := - match cg.charMap.find? c with + match cg.charMap.get? c with | .some g => g -- Invalid character | .none => cg.ngroup + 2 @@ -278,14 +278,14 @@ section s!"Group representing beginning of string := {cg.ngroup}" :: s!"Group representing end of string := {cg.ngroup + 1}" :: s!"Group representing other utf-8 characters := {cg.ngroup + 2}" :: - "(GroupIdx, Group members):" :: cggroups.data ++ - s!"(State, GroupIdx → State'):" :: dtr.data ++ - s!"(State, Attributes)" :: attrs.data + "(GroupIdx, Group members):" :: cggroups.toList ++ + s!"(State, GroupIdx → State'):" :: dtr.toList ++ + s!"(State, Attributes)" :: attrs.toList String.intercalate "\n " all ++ "\n⦘⦘" def ADFA.toString (a : ADFA σ) : String := ADFA.toStringAux a (fun l => ToString.toString l.toList) - + instance : ToString (ADFA σ) where toString := ADFA.toString @@ -297,7 +297,7 @@ section end def CharGrouping.toStringForChar (cg : CharGrouping Char) : String := - CharGrouping.toStringAux cg (fun l => + CharGrouping.toStringAux cg (fun l => let sorted := sort (l.toList.map Char.toNat) let str := String.mk (sorted.map Char.ofNat) ToString.toString (repr str)) @@ -306,7 +306,7 @@ instance : ToString (CharGrouping Char) where toString := CharGrouping.toStringForChar def ADFA.toStringForChar (a : ADFA Char) : String := - ADFA.toStringAux a (fun l => + ADFA.toStringAux a (fun l => let sorted := sort (l.toList.map Char.toNat) let str := String.mk (sorted.map Char.ofNat) ToString.toString (repr str)) @@ -316,38 +316,38 @@ instance : ToString (ADFA Char) where def ERE.charGrouping (e : ERE) : CharGrouping Char := Id.run <| do let hsets := e.brackets.map EREBracket.toHashSet - let mut all := hsets.foldl (fun hs nhs => hs.insertMany nhs) HashSet.empty - let mut charMap := all.fold (fun hs c => hs.insert c 0) HashMap.empty + let mut all := hsets.foldl (fun hs nhs => hs.insertMany nhs) Std.HashSet.empty + let mut charMap := all.fold (fun hs c => hs.insert c 0) Std.HashMap.empty -- Current number of groups let mut curidx := 1 for hset in hsets do - let mut reloc : HashMap Nat Nat := {} + let mut reloc : Std.HashMap Nat Nat := {} for c in hset do - let cidx := charMap.find! c - match reloc.find? cidx with + let cidx := charMap.get! c + match reloc.get? cidx with | .some r => charMap := charMap.insert c r | .none => reloc := reloc.insert cidx curidx; charMap := charMap.insert c curidx; curidx := curidx + 1 let mut ridx := 0 - let mut reloc : HashMap Nat Nat := {} + let mut reloc : Std.HashMap Nat Nat := {} for (_, i) in charMap.toList do - match reloc.find? i with + match reloc.get? i with | .some _ => continue | .none => reloc := reloc.insert i ridx; ridx := ridx + 1 - charMap := HashMap.ofList (charMap.toList.map (fun (c, i) => (c, reloc.find! i))) + charMap := Std.HashMap.ofList (charMap.toList.map (fun (c, i) => (c, reloc.get! i))) return CharGrouping.mk ridx all charMap private partial def ERE.toNFAAux (cg : CharGrouping Char) : ERE → (NFA Nat) | .bracket b => let bs := toString b - let states := bs.foldl (fun hs c => hs.insert (cg.charMap.find! c)) HashSet.empty + let states := bs.foldl (fun hs c => hs.insert (cg.charMap.get! c)) Std.HashSet.empty NFA.ofSymbPlus states.toArray | .bracketN b => let bs := toString b -- All `utf-8` characters - let initHs := HashSet.empty.insertMany ((cg.ngroup + 2) :: List.range cg.ngroup) - let states := bs.foldl (fun hs c => hs.erase (cg.charMap.find! c)) initHs + let initHs := Std.HashSet.empty.insertMany ((cg.ngroup + 2) :: List.range cg.ngroup) + let states := bs.foldl (fun hs c => hs.erase (cg.charMap.get! c)) initHs NFA.ofSymbPlus states.toArray | .startp => NFA.ofSymb (cg.ngroup) | .endp => NFA.ofSymb (cg.ngroup + 1) @@ -587,4 +587,4 @@ private def testit (a : ADFA Char) (s : Substring) (short:=false) (strict:=false end Regex -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Parser/NDFA.lean b/Auto/Parser/NDFA.lean index b59c949..0ea5eb9 100644 --- a/Auto/Parser/NDFA.lean +++ b/Auto/Parser/NDFA.lean @@ -5,7 +5,7 @@ open Lean namespace Auto -private def sort : List Nat → List Nat := +private def sort : List Nat → List Nat := have : DecidableRel Nat.le := fun (x y : Nat) => inferInstanceAs (Decidable (x <= y)) List.mergeSort Nat.le @@ -38,11 +38,11 @@ section NFA However, the initial state might have incoming edges -/ structure NFA where - tr : Array (HashMap (Unit ⊕ σ) (Array Nat)) + tr : Array (Std.HashMap (Unit ⊕ σ) (Array Nat)) -- Each state (including the accept state) is associated -- with an array of attributes. So the length of `attrs` -- should be `tr.size + 1` - attrs : Array (HashSet String) + attrs : Array (Std.HashSet String) deriving Inhabited variable {σ : Type} [BEq σ] [Hashable σ] @@ -61,12 +61,12 @@ section NFA c.toArray.map (fun el => snatS idx el)) let tr := tr.concatMap id let attrs := n.attrs.mapIdx (fun idx attrs => s!"{idx} : {attrs.toList}") - let all := "NFA ⦗⦗" :: s!"Accept state := {n.tr.size}" :: tr.data ++ attrs.data + let all := "NFA ⦗⦗" :: s!"Accept state := {n.tr.size}" :: tr.toList ++ attrs.toList String.intercalate "\n " all ++ "\n⦘⦘" - + instance : ToString (NFA σ) where toString := NFA.toString - + private def NFA.nextStatesOfState (r : NFA σ) (s : Nat) (c : Unit ⊕ σ) : Array Nat := if h₁ : s > r.tr.size then panic! s!"NFA.nextStates :: State {s} is not valid for {r}" @@ -80,12 +80,12 @@ section NFA have h₄ : (s = Array.size r.tr) = False := eq_false h₂ simp [h₄] at h₃; simp [h₃] ) - match hmap.find? c with + match hmap.get? c with | .some arr => arr | .none => #[] - + -- Why this does not need `partial`? - def NFA.εClosureOfStates (r : NFA σ) (ss : HashSet Nat) := Id.run <| do + def NFA.εClosureOfStates (r : NFA σ) (ss : Std.HashSet Nat) := Id.run <| do let mut front := ss.toArray let mut cur := 0 let mut ret := ss @@ -99,15 +99,15 @@ section NFA ret := ret.insert n return ret - def NFA.move (r : NFA σ) (ss : HashSet Nat) (c : σ) := + def NFA.move (r : NFA σ) (ss : Std.HashSet Nat) (c : σ) := let sss := ss.toArray.map (fun s => NFA.nextStatesOfState r s (.inr c)) - sss.foldl (fun hs s => hs.insertMany s) HashSet.empty + sss.foldl (fun hs s => hs.insertMany s) Std.HashSet.empty -- Valid moves from a set of states `ss`, ignoring `ε` transitions -- Returns a hashmap from symbol to HashSet of states - def NFA.moves [ToString σ] (r : NFA σ) (ss : HashSet Nat) : HashMap σ (HashSet Nat) := + def NFA.moves [ToString σ] (r : NFA σ) (ss : Std.HashSet Nat) : Std.HashMap σ (Std.HashSet Nat) := Id.run <| do - let mut ret : HashMap σ (HashSet Nat) := {} + let mut ret : Std.HashMap σ (Std.HashSet Nat) := {} for i in ss do if i > r.tr.size then panic! s!"NFA.moves :: {i} from state set {ss.toList} is not a valid state of {r}" @@ -121,22 +121,22 @@ section NFA -- Ignore `ε` transitions | .inl .unit => continue | .inr c => - if let some d := ret.find? c then + if let some d := ret.get? c then ret := ret.insert c (d.insertMany dests) else - ret := ret.insert c (HashSet.empty.insertMany dests) + ret := ret.insert c (Std.HashSet.empty.insertMany dests) return ret - + -- Move, then compute ε-closure - def NFA.moveε (r : NFA σ) (ss : HashSet Nat) (c : σ) : HashSet Nat := + def NFA.moveε (r : NFA σ) (ss : Std.HashSet Nat) (c : σ) : Std.HashSet Nat := r.εClosureOfStates (r.move ss c) - def NFA.moveεMany (r : NFA σ) (ss : HashSet Nat) (cs : Array σ) := + def NFA.moveεMany (r : NFA σ) (ss : Std.HashSet Nat) (cs : Array σ) := cs.foldl (fun ss' c => r.moveε ss' c) ss def NFA.run (r : NFA σ) (cs : Array σ) := - r.moveεMany (r.εClosureOfStates (HashSet.empty.insert 0)) cs - + r.moveεMany (r.εClosureOfStates (Std.HashSet.empty.insert 0)) cs + end Run /-- Criterion : The destination of all transitions should be ≤ n.size -/ @@ -149,10 +149,10 @@ section NFA def NFA.normalize (n : NFA σ) : NFA σ := let size := n.tr.size let normEntry (x : _ × Array Nat) := - (x.fst, (HashSet.empty.insertMany (x.snd.filter (· <= size))).toArray) - let tr' := n.tr.map (fun hs => HashMap.ofList (hs.toList.map normEntry)) + (x.fst, (Std.HashSet.empty.insertMany (x.snd.filter (· <= size))).toArray) + let tr' := n.tr.map (fun hs => Std.HashMap.ofList (hs.toList.map normEntry)) let attrs' := n.attrs[0:size+1].toArray - let attrs' := attrs'.append ⟨(List.range (size + 1 - attrs'.size)).map (fun _ => HashSet.empty)⟩ + let attrs' := attrs'.append ⟨(List.range (size + 1 - attrs'.size)).map (fun _ => Std.HashSet.empty)⟩ NFA.mk tr' attrs' /-- Whether the NFA's initial state has incoming edges -/ @@ -162,11 +162,11 @@ section NFA private def NFA.relocateEntry (x : α × Array Nat) (off : Nat) := (x.fst, x.snd.map (· + off)) - private def NFA.relocateHMap (x : HashMap (Unit ⊕ σ) (Array Nat)) (off : Nat) := - HashMap.ofList (x.toList.map (relocateEntry · off)) + private def NFA.relocateHMap (x : Std.HashMap (Unit ⊕ σ) (Array Nat)) (off : Nat) := + Std.HashMap.ofList (x.toList.map (relocateEntry · off)) - private def NFA.addEdgesToHMap (x : HashMap (Unit ⊕ σ) (Array Nat)) (e : (Unit ⊕ σ) × Array Nat) := - x.insert e.fst (match x.find? e.fst with | some arr => arr ++ e.snd | none => e.snd) + private def NFA.addEdgesToHMap (x : Std.HashMap (Unit ⊕ σ) (Array Nat)) (e : (Unit ⊕ σ) × Array Nat) := + x.insert e.fst (match x.get? e.fst with | some arr => arr ++ e.snd | none => e.snd) /-- Add attribute to a specific state -/ def NFA.addAttrToState (n : NFA σ) (s : Nat) (attr : String) := @@ -185,15 +185,15 @@ section NFA NFA.mk n.tr new_attrs /-- Does not accept any string -/ - def NFA.zero : NFA σ := NFA.mk #[HashMap.empty] #[.empty, .empty] + def NFA.zero : NFA σ := NFA.mk #[Std.HashMap.empty] #[.empty, .empty] /-- Only accepts empty string -/ def NFA.epsilon : NFA σ := - NFA.mk #[HashMap.empty.insert (.inl .unit) #[1]] #[.empty, .empty] + NFA.mk #[Std.HashMap.empty.insert (.inl .unit) #[1]] #[.empty, .empty] /-- Accepts a character -/ def NFA.ofSymb (c : σ) : NFA σ := - NFA.mk #[HashMap.empty.insert (.inr c) #[1]] #[.empty, .empty] + NFA.mk #[Std.HashMap.empty.insert (.inr c) #[1]] #[.empty, .empty] /-- Produce an NFA whose language is the union of `m`'s and `n`'s -/ def NFA.plus (m n : NFA σ) : NFA σ := @@ -202,16 +202,16 @@ section NFA let off_n := m.tr.size + 2 -- `acc'` is the new accept state let acc' := m.tr.size + n.tr.size + 3 - let initTrans : HashMap (Unit ⊕ σ) (Array Nat) := - HashMap.empty.insert (Sum.inl .unit) #[off_m, off_n] + let initTrans : Std.HashMap (Unit ⊕ σ) (Array Nat) := + Std.HashMap.empty.insert (Sum.inl .unit) #[off_m, off_n] -- Move the states of `m` by `off_m` let new_mtr := m.tr.map (relocateHMap · off_m) - let new_mtr := new_mtr.push (HashMap.empty.insert (.inl .unit) #[acc']) + let new_mtr := new_mtr.push (Std.HashMap.empty.insert (.inl .unit) #[acc']) -- Move the states of `n` by `off_n` let new_ntr := n.tr.map (relocateHMap · off_n) - let new_ntr := new_ntr.push (HashMap.empty.insert (.inl .unit) #[acc']) + let new_ntr := new_ntr.push (Std.HashMap.empty.insert (.inl .unit) #[acc']) let new_tr := #[initTrans] ++ new_mtr ++ new_ntr - let new_attrs := #[HashSet.empty] ++ m.attrs ++ n.attrs ++ #[HashSet.empty] + let new_attrs := #[Std.HashSet.empty] ++ m.attrs ++ n.attrs ++ #[Std.HashSet.empty] NFA.mk new_tr new_attrs def NFA.multiPlus (as : Array (NFA σ)) := @@ -221,16 +221,16 @@ section NFA | _ => let (acc', offs) : Nat × Array Nat := as.foldl (fun (cur, acc) (arr : NFA σ) => (cur + arr.tr.size + 1, acc.push cur)) (1, #[]) - let initTrans : HashMap (Unit ⊕ σ) (Array Nat) := - HashMap.empty.insert (Sum.inl .unit) offs + let initTrans : Std.HashMap (Unit ⊕ σ) (Array Nat) := + Std.HashMap.empty.insert (Sum.inl .unit) offs let trs := (as.zip offs).map (fun ((a, off) : NFA σ × Nat) => let new_a := a.tr.map (relocateHMap · off) - new_a.push (HashMap.empty.insert (.inl .unit) #[acc']) + new_a.push (Std.HashMap.empty.insert (.inl .unit) #[acc']) ) let new_tr := (#[#[initTrans]] ++ trs).concatMap id - let new_attrs := #[HashSet.empty] ++ + let new_attrs := #[Std.HashSet.empty] ++ (as.map (fun (⟨_, attrs⟩ : NFA σ) => attrs)).concatMap id ++ - #[HashSet.empty] + #[Std.HashSet.empty] NFA.mk new_tr new_attrs def NFA.comp (m n : NFA σ) : NFA σ := @@ -267,11 +267,11 @@ section NFA let acc' := m.tr.size + 2 -- The new location of the original accept state of `m` -- let macc' := m.size + 1 - let initTrans : HashMap (Unit ⊕ σ) (Array Nat) := - HashMap.empty.insert (Sum.inl .unit) #[1, acc'] + let initTrans : Std.HashMap (Unit ⊕ σ) (Array Nat) := + Std.HashMap.empty.insert (Sum.inl .unit) #[1, acc'] -- Move the states of `m` by `1` let new_mtr := m.tr.map (relocateHMap · 1) - let new_mtr := new_mtr.push (HashMap.empty.insert (.inl .unit) #[1, acc']) + let new_mtr := new_mtr.push (Std.HashMap.empty.insert (.inl .unit) #[1, acc']) let new_tr := #[initTrans] ++ new_mtr let new_attrs := #[.empty] ++ m.attrs ++ #[.empty] NFA.mk new_tr new_attrs @@ -282,7 +282,7 @@ section NFA | .cons a .nil => a | a :: as => NFA.comp a (NFA.multiCompAux as) - def NFA.multiComp (a : Array (NFA σ)) := NFA.multiCompAux a.data + def NFA.multiComp (a : Array (NFA σ)) := NFA.multiCompAux a.toList def NFA.repeatN (r : NFA σ) (n : Nat) := NFA.multiComp ⟨(List.range n).map (fun _ => r)⟩ @@ -296,7 +296,7 @@ section NFA if r.hasEdgeToInit then -- Add a new state as the initial state so that the -- new initial state has no incoming edges - let new_tr := #[HashMap.empty.insert (.inl .unit) #[1]] ++ r.tr.map (relocateHMap · 1) + let new_tr := #[Std.HashMap.empty.insert (.inl .unit) #[1]] ++ r.tr.map (relocateHMap · 1) let new_attrs := #[.empty] ++ r.attrs NFA.mk new_tr new_attrs else @@ -309,7 +309,7 @@ section NFA new_r.modify 0 (fun hm => NFA.addEdgesToHMap hm (.inl .unit, #[acc'])) ) let new_tr := new_trs.concatMap id - let new_attrs : Array (HashSet String) := + let new_attrs : Array (Std.HashSet String) := ((Array.mk (List.range (n - 1))).map (fun _ => r.attrs[:r.tr.size].toArray)).concatMap id ++ r.attrs NFA.mk new_tr new_attrs @@ -322,11 +322,11 @@ section NFA /-- Accepts all characters in an array of characters -/ def NFA.ofSymbPlus (cs : Array σ) : NFA σ := - NFA.mk #[HashMap.ofList (cs.map (fun c => (.inr c,#[1]))).data] #[.empty, .empty] + NFA.mk #[Std.HashMap.ofList (cs.map (fun c => (.inr c,#[1]))).toList] #[.empty, .empty] /-- An `NFA UInt32` that accepts exactly a string -/ def NFA.ofSymbComp (s : Array σ) : NFA σ := - let tr := (Array.mk s.data).mapIdx (fun idx c => HashMap.empty.insert (.inr c) #[idx.val + 1]) + let tr := (Array.mk s.toList).mapIdx (fun idx c => Std.HashMap.empty.insert (.inr c) #[idx.val + 1]) let attrs := Array.mk ((List.range (s.size + 1)).map (fun _ => .empty)) NFA.mk tr attrs @@ -339,7 +339,7 @@ section NFA HashMap.ofList [(.inr "a", #[5]), (.inr "b", #[1, 0])], HashMap.ofList [(.inl .unit, #[1]), (.inr "c", #[2, 4]), (.inr "a", #[6,1,2])] ], #[]⟩ - + def test₂ : NFA String := test₁.normalize #eval IO.println test₁ @@ -370,13 +370,13 @@ section NFA end NFA section DFA - + -- Alphabet of DFA variable (σ : Type) [BEq σ] [Hashable σ] structure DFA where -- Array of accept states - accepts : HashSet Nat + accepts : Std.HashSet Nat -- Transition function -- `0` is the initial statet -- `{0, 1, ⋯, tr.size}` are the set of allowed states, where @@ -385,13 +385,13 @@ section DFA -- If the transition map of state `i` does not include -- an entry for character `c`, then the transition from -- `i` to `c` ends in `malformed input` state - tr : Array (HashMap σ Nat) + tr : Array (Std.HashMap σ Nat) -- Each state (except for the `malformed input` state) -- is associated with an array of attributes. -- So, we should have `attrs.size == tr.size` - attrs : Array (HashSet String) + attrs : Array (Std.HashSet String) deriving Inhabited - + variable {σ : Type} [BEq σ] [Hashable σ] [ToString σ] def DFA.toString (d : DFA σ) : String := @@ -402,7 +402,7 @@ section DFA let all := "DFA ⦗⦗" :: s!"Accept states := {d.accepts.toList}" :: s!"Size (Malformed-input state) = {d.tr.size}" :: - tr.data ++ attrs.data + tr.toList ++ attrs.toList String.intercalate "\n " all ++ "\n⦘⦘" instance : ToString (DFA σ) where @@ -422,7 +422,7 @@ section DFA have h₄ : (s = Array.size _) = False := eq_false h₂ simp [h₄] at h₃; simp [h₃] ) - match hmap.find? c with + match hmap.get? c with | .some s => s -- `malformed input` | .none => d.tr.size @@ -431,17 +431,17 @@ section DFA if !n.wf then panic! s!"DFA.ofNFA :: {n} is not well-formed" -- Array of states - let mut dstates : Array (List Nat) := #[sort (n.εClosureOfStates (HashSet.empty.insert 0)).toList] + let mut dstates : Array (List Nat) := #[sort (n.εClosureOfStates (Std.HashSet.empty.insert 0)).toList] -- Map from state to idx of state - let mut idxmap : HashMap (List Nat) Nat := - HashMap.empty.insert dstates[0]! 0 + let mut idxmap : Std.HashMap (List Nat) Nat := + Std.HashMap.empty.insert dstates[0]! 0 -- `Unit` represents the `malformed input` state - let mut tr : Array (HashMap σ (Nat ⊕ Unit)) := #[{}] + let mut tr : Array (Std.HashMap σ (Nat ⊕ Unit)) := #[{}] -- Next state to process let mut cur := 0 while h : cur < dstates.size do let st := dstates[cur] - let moves := n.moves (HashSet.empty.insertMany st) + let moves := n.moves (Std.HashSet.empty.insertMany st) for (c, st) in moves.toList do -- If `st` is empty, then the move ends in `malformed input` state if st.size == 0 then @@ -454,11 +454,11 @@ section DFA idxmap := idxmap.insert εst idxmap.size tr := tr.push {} -- Now `idxmap` contains `εst` - let idx := idxmap.find! εst + let idx := idxmap.get! εst tr := tr.modify cur (fun hmap => hmap.insert c (.inl idx)) cur := cur + 1 let rettr := tr.map ( - fun hmap => HashMap.ofList (hmap.toList.map ( + fun hmap => Std.HashMap.ofList (hmap.toList.map ( fun (s, nu) => match nu with | .inl n => (s, n) @@ -466,9 +466,9 @@ section DFA )) ) let accepts := dstates.mapIdx (fun idx l => if l.contains n.tr.size then some idx.val else none) - let accepts := accepts.foldl (fun hs o => if let some x := o then hs.insert x else hs) HashSet.empty - let attrs := dstates.map (fun l => - (Array.mk l).foldl (fun acc s => if let some x := n.attrs[s]? then acc.insertMany x else acc) HashSet.empty) + let accepts := accepts.foldl (fun hs o => if let some x := o then hs.insert x else hs) Std.HashSet.empty + let attrs := dstates.map (fun l => + (Array.mk l).foldl (fun acc s => if let some x := n.attrs[s]? then acc.insertMany x else acc) Std.HashSet.empty) return DFA.mk accepts rettr attrs /- diff --git a/Auto/Parser/SMTSexp.lean b/Auto/Parser/SMTSexp.lean index e34becd..d78179e 100644 --- a/Auto/Parser/SMTSexp.lean +++ b/Auto/Parser/SMTSexp.lean @@ -89,7 +89,7 @@ deriving Inhabited, BEq, Hashable partial def Sexp.toString : Sexp → String | .atom l => ToString.toString l -| .app ls => "(" ++ String.intercalate " " (ls.map toString).data ++ ")" +| .app ls => "(" ++ String.intercalate " " (ls.map toString).toList ++ ")" instance : ToString Sexp where toString e := Sexp.toString e diff --git a/Auto/Parser/TPTP.lean b/Auto/Parser/TPTP.lean index 468b5e6..15351c7 100644 --- a/Auto/Parser/TPTP.lean +++ b/Auto/Parser/TPTP.lean @@ -33,11 +33,11 @@ def tokens := [ "~", ",", "(", ")", "*", "!", "?", "^", ":", "[", "]", "!>", ".", "*" ] -def tokenHashMap : HashSet String := - HashSet.empty.insertMany tokens +def tokenHashMap : Std.HashSet String := + Std.HashSet.empty.insertMany tokens -def tokenPrefixes : HashSet String := - HashSet.empty.insertMany $ tokens.bind (fun t => Id.run do +def tokenPrefixes : Std.HashSet String := + Std.HashSet.empty.insertMany $ tokens.bind (fun t => Id.run do let mut res := [] let mut pref := "" for c in t.data do @@ -605,7 +605,7 @@ open Embedding.Lam in structure ParsingInfo where lamVarTy : Array LamSort lamEVarTy : Array LamSort - proverSkolem : HashMap String (LamSort × Nat) := {} + proverSkolem : Std.HashMap String (LamSort × Nat) := {} open Embedding.Lam in def ParsingInfo.toLamTyVal (pi : ParsingInfo) : LamTyVal := @@ -623,7 +623,7 @@ def LamConstr.ofBaseTerm (pi : ParsingInfo) (b : LamBaseTerm) : LamConstr := open Embedding.Lam in def ident2LamConstr (pi : ParsingInfo) (s : String) : LamConstr := - match pi.proverSkolem.find? s with + match pi.proverSkolem.get? s with | .some (s, n) => .term s (.etom (n + pi.lamEVarTy.size)) | .none => match s.get ⟨0⟩ with @@ -668,7 +668,7 @@ def ident2LamConstr (pi : ParsingInfo) (s : String) : LamConstr := open Embedding.Lam in /-- This function is only for zipperposition's output -/ -partial def term2LamTerm (pi : ParsingInfo) (t : Term) (lctx : HashMap String (Nat × LamSort) := {}) : LamConstr := +partial def term2LamTerm (pi : ParsingInfo) (t : Term) (lctx : Std.HashMap String (Nat × LamSort) := {}) : LamConstr := match t with | ⟨.ident "$i", []⟩ => .error "Does not have iota" | ⟨.ident "$tType", []⟩ => .kind @@ -811,7 +811,7 @@ partial def term2LamTerm (pi : ParsingInfo) (t : Term) (lctx : HashMap String (N | _, r => .error s!"Unexpected term {t} produces ({r})" | _ => .error s!"term2LamTerm :: Could not translate to Lean Expr: {t}" where - processVar (pi : ParsingInfo) (var : Term) (lctx : HashMap String (Nat × LamSort)) : Option (String × LamSort) := + processVar (pi : ParsingInfo) (var : Term) (lctx : Std.HashMap String (Nat × LamSort)) : Option (String × LamSort) := match var with | ⟨.ident v, ty⟩ => match ty with @@ -821,8 +821,8 @@ where | _ => .none | _ => .none | _ => .none - deBruijnAndSort? (lctx : HashMap String (Nat × LamSort)) (id : String) : Option (Nat × LamSort) := - match lctx.find? id with + deBruijnAndSort? (lctx : Std.HashMap String (Nat × LamSort)) (id : String) : Option (Nat × LamSort) := + match lctx.get? id with | .some (n, s) => (lctx.size - 1 - n, s) | .none => none lamConstrMkAppN (head : LamConstr) (args : List LamConstr) := diff --git a/Auto/Translation/Assumptions.lean b/Auto/Translation/Assumptions.lean index f0bcb56..b338b07 100644 --- a/Auto/Translation/Assumptions.lean +++ b/Auto/Translation/Assumptions.lean @@ -17,7 +17,7 @@ deriving Inhabited, Hashable, BEq partial def DTr.toString : DTr → String | .node s dtrs => - s ++ " [" ++ String.intercalate ", " (dtrs.map DTr.toString).data ++ "]" + s ++ " [" ++ String.intercalate ", " (dtrs.map DTr.toString).toList ++ "]" | .leaf s => s instance : ToString DTr where @@ -278,7 +278,7 @@ deriving Inhabited, Hashable, BEq instance : ToMessageData MLemmaInst where toMessageData mi := MessageData.compose m!"MLemmaInst ⦗⦗ {mi.origProof} " (.compose - (MessageData.intercalate " " (mi.args.data.map (fun e => m!"({e})"))) + (MessageData.intercalate " " (mi.args.toList.map (fun e => m!"({e})"))) m!" : {mi.type} ⦘⦘") def MLemmaInst.ofLemmaInst (li : LemmaInst) : MetaM (Array Level × Array Expr × MLemmaInst) := do @@ -312,13 +312,13 @@ def LemmaInst.ofMLemmaInst (mi : MLemmaInst) : MetaM LemmaInst := do let lem : Lemma := ⟨⟨proof, type, deriv⟩, s.paramNames⟩ return ⟨lem, nbinders, nargs⟩ -partial def collectUniverseLevels : Expr → MetaM (HashSet Level) -| .bvar _ => return HashSet.empty +partial def collectUniverseLevels : Expr → MetaM (Std.HashSet Level) +| .bvar _ => return Std.HashSet.empty | e@(.fvar _) => do collectUniverseLevels (← instantiateMVars (← Meta.inferType e)) | e@(.mvar _) => do collectUniverseLevels (← instantiateMVars (← Meta.inferType e)) -| .sort u => return HashSet.empty.insert u +| .sort u => return Std.HashSet.empty.insert u | e@(.const _ us) => do - let hus := HashSet.empty.insertMany us + let hus := Std.HashSet.empty.insertMany us let tys ← collectUniverseLevels (← instantiateMVars (← Meta.inferType e)) return mergeHashSet hus tys | .app fn arg => do @@ -338,14 +338,14 @@ partial def collectUniverseLevels : Expr → MetaM (HashSet Level) let vs ← collectUniverseLevels v let bodys ← collectUniverseLevels body return mergeHashSet (mergeHashSet tys vs) bodys -| .lit _ => return HashSet.empty.insert (.succ .zero) +| .lit _ => return Std.HashSet.empty.insert (.succ .zero) | .mdata _ e' => collectUniverseLevels e' | .proj .. => throwError "Please unfold projections before collecting universe levels" def computeMaxLevel (facts : Array UMonoFact) : MetaM Level := do let levels ← facts.foldlM (fun hs ⟨_, ty, _⟩ => do let tyUs ← collectUniverseLevels ty - return mergeHashSet tyUs hs) HashSet.empty + return mergeHashSet tyUs hs) Std.HashSet.empty -- Compute the universe level that we need to lift to -- Use `.succ` two times to reveal bugs let level := Level.succ (.succ (levels.fold (fun l l' => Level.max l l') Level.zero)) diff --git a/Auto/Translation/Inductive.lean b/Auto/Translation/Inductive.lean index b376a7f..52b6083 100644 --- a/Auto/Translation/Inductive.lean +++ b/Auto/Translation/Inductive.lean @@ -80,7 +80,7 @@ def SimpleIndVal.zetaReduce (si : SimpleIndVal) : MetaM SimpleIndVal := do is an array of `(instantiated_tyctor, [SimpleIndVal associated to tyctor])` -/ structure CollectInduct.State where - recorded : HashMap Name (Array Expr) := {} + recorded : Std.HashMap Name (Array Expr) := {} sis : Array (Array SimpleIndVal) := #[] abbrev IndCollectM := StateRefT CollectInduct.State MetaM @@ -127,7 +127,7 @@ mutual return if !(← getRecorded).contains tyctor then setRecorded ((← getRecorded).insert tyctor #[]) - let .some arr := (← getRecorded).find? tyctor + let .some arr := (← getRecorded).get? tyctor | throwError "collectAppInstSimpleInduct :: Unexpected error" for e' in arr do if ← Meta.isDefEq e e' then diff --git a/Auto/Translation/Inhabitation.lean b/Auto/Translation/Inhabitation.lean index b401db5..14cf56e 100644 --- a/Auto/Translation/Inhabitation.lean +++ b/Auto/Translation/Inhabitation.lean @@ -5,7 +5,7 @@ open Lean namespace Auto.Inhabitation -private def logicalConsts : HashSet Name := HashSet.empty.insertMany +private def logicalConsts : Std.HashSet Name := Std.HashSet.empty.insertMany #[``Eq, ``Exists, ``And, ``Or, ``Iff, ``Not] /-- @@ -28,7 +28,7 @@ where go (pos : Array Bool) (e : Expr) : Array (Array Bool) := | _ => #[pos] def getExprNonImpPosition (pos : Array Bool) (e : Expr) : Option Expr := - go pos.data e + go pos.toList e where go (pos : List Bool) (e : Expr) : Option Expr := match pos with | .nil => .some e diff --git a/Auto/Translation/Lam2DAtomAsFVar.lean b/Auto/Translation/Lam2DAtomAsFVar.lean index b471ed7..ce5f137 100644 --- a/Auto/Translation/Lam2DAtomAsFVar.lean +++ b/Auto/Translation/Lam2DAtomAsFVar.lean @@ -50,11 +50,11 @@ structure State where -- Etoms to be abstracted etomsToAbstract : Array (FVarId × Nat) := #[] -- Type atoms that are used in the expressions sent to external prover - typeAtomFVars : HashMap Nat Expr := {} + typeAtomFVars : Std.HashMap Nat Expr := {} -- Term atoms that are used in the expressions sent to external prover - termAtomFVars : HashMap Nat Expr := {} + termAtomFVars : Std.HashMap Nat Expr := {} -- Etoms that are used in the expression sent to external prover - etomFVars : HashMap Nat Expr := {} + etomFVars : Std.HashMap Nat Expr := {} abbrev ExternM := StateRefT State MetaStateM diff --git a/Auto/Translation/Lam2TH0.lean b/Auto/Translation/Lam2TH0.lean index 585890c..b1996fa 100644 --- a/Auto/Translation/Lam2TH0.lean +++ b/Auto/Translation/Lam2TH0.lean @@ -40,6 +40,6 @@ def lam2TH0 (lamVarTy : Array LamSort) (lamEVarTy : Array LamSort) (facts : Arra | .ok ts => return s!"thf(fact{i}, axiom, {ts})." | .error e => throwError e) let sep := "\n" - return s!"{String.intercalate sep sorts}\n\n{String.intercalate sep types}\n\n{String.intercalate sep facts.data}" + return s!"{String.intercalate sep sorts}\n\n{String.intercalate sep types}\n\n{String.intercalate sep facts.toList}" end Auto diff --git a/Auto/Translation/LamFOL2SMT.lean b/Auto/Translation/LamFOL2SMT.lean index 4f2a337..29f5090 100644 --- a/Auto/Translation/LamFOL2SMT.lean +++ b/Auto/Translation/LamFOL2SMT.lean @@ -50,19 +50,19 @@ instance : ToString LamAtomic where toString := LamAtomic.toString def LamAtomic.toLeanExpr - (tyValMap varValMap etomValMap : HashMap Nat Expr) + (tyValMap varValMap etomValMap : Std.HashMap Nat Expr) (atomic : LamAtomic) : MetaM Expr:= match atomic with | .sort n => do - let .some e := tyValMap.find? n + let .some e := tyValMap.get? n | throwError "SMT.printValuation :: Unknown sort atom {n}" return e | .term n => do - let .some e := varValMap.find? n + let .some e := varValMap.get? n | throwError "SMT.printValuation :: Unknown term atom {n}" return e | .etom n => do - let .some e := etomValMap.find? n + let .some e := etomValMap.get? n | throwError "SMT.printValuation :: Unknown etom {n}" return e | .bvOfNat n => do @@ -525,11 +525,11 @@ def termAuxDecls : Array IR.SMT.Command := `printFn : (tyValMap : _) → (varValMap : _) → (etomValMap : _) → MetaM α` -/ def withExprValuation - {α : Type} [Inhabited α] (sni : SMTNamingInfo) (h2lMap : HashMap LamAtomic String) - (printFn : HashMap Nat Expr → HashMap Nat Expr → HashMap Nat Expr → MetaM α) : + {α : Type} [Inhabited α] (sni : SMTNamingInfo) (h2lMap : Std.HashMap LamAtomic String) + (printFn : Std.HashMap Nat Expr → Std.HashMap Nat Expr → Std.HashMap Nat Expr → MetaM α) : MetaM α := do - let tyValMap := HashMap.ofList (sni.tyVal.zipWithIndex.map (fun ((e, _), n) => (n, e))).data - let varValMap := HashMap.ofList (sni.varVal.zipWithIndex.map (fun ((e, _), n) => (n, e))).data + let tyValMap := Std.HashMap.ofList (sni.tyVal.zipWithIndex.map (fun ((e, _), n) => (n, e))).toList + let varValMap := Std.HashMap.ofList (sni.varVal.zipWithIndex.map (fun ((e, _), n) => (n, e))).toList let etomsWithName := h2lMap.toArray.filterMap (fun (atomic, name) => match atomic with | .etom n => .some (n, name) | _ => .none) let declInfos ← etomsWithName.mapM (fun (n, name) => do @@ -538,7 +538,7 @@ def withExprValuation let type ← Lam2D.interpLamSortAsUnlifted tyValMap s return (Name.mkSimple name, .default, fun _ => pure type)) Meta.withLocalDecls declInfos (fun etomFVars => do - let etomValMap := HashMap.ofList ((etomsWithName.zip etomFVars).map (fun ((n, _), e) => (n, e))).data + let etomValMap := Std.HashMap.ofList ((etomsWithName.zip etomFVars).map (fun ((n, _), e) => (n, e))).toList printFn tyValMap varValMap etomValMap) end SMT diff --git a/Auto/Translation/LamReif.lean b/Auto/Translation/LamReif.lean index 8350e58..aabe7a5 100644 --- a/Auto/Translation/LamReif.lean +++ b/Auto/Translation/LamReif.lean @@ -33,13 +33,13 @@ open Embedding.Lam -/ structure State where -- Maps previously reified type to their type atom index - tyVarMap : HashMap Expr Nat := {} + tyVarMap : Std.HashMap Expr Nat := {} -- Maps previously reified expressions to their term atom index -- Whenever we encounter an atomic expression, we look up -- `varMap`. If it's already reified, then `varMap` -- tells us about its index. If it's not reified, insert -- it to `varMap`. - varMap : HashMap Expr Nat := {} + varMap : Std.HashMap Expr Nat := {} -- `tyVal` is the inverse of `tyVarMap` -- The `e : Expr` is the un-lifted valuation of the type atom -- The `lvl : level` is the sort level of `e` @@ -54,7 +54,7 @@ structure State where -- lamILTy lamILTy : Array LamSort := #[] -- Inverse of `lamILTy` - isomTyMap : HashMap LamSort Nat := {} + isomTyMap : Std.HashMap LamSort Nat := {} /- This hashmap contains assertions that have external (lean) proof · The key `t : LamTerm` is the corresponding λ term, @@ -67,7 +67,7 @@ structure State where This field also corresponds to the `ImportTable` in `LamChecker.lean` -/ - assertions : HashMap LamTerm (Expr × DTr × LamTerm × Nat) := {} + assertions : Std.HashMap LamTerm (Expr × DTr × LamTerm × Nat) := {} /- This hashmap contains lamsorts who has external proof of inhabitation · The key `s : LamSort` is the corresponding `λ` sort @@ -75,12 +75,12 @@ structure State where · The second `DTr` is the derivation of `e` from the context · The third `n : Nat` is the position of the inhabitation fact in the `ImportTable` -/ - inhabitations : HashMap LamSort (Expr × DTr × Nat) := {} + inhabitations : Std.HashMap LamSort (Expr × DTr × Nat) := {} -- Records the EvalResult and etoms generated by checksteps -- that produce etoms - chkStepCache : HashMap ChkStep (EvalResult × Array Nat) := {} + chkStepCache : Std.HashMap ChkStep (EvalResult × Array Nat) := {} -- The check step that produces a given etom - etomChkStep : HashMap Nat ChkStep := {} + etomChkStep : Std.HashMap Nat ChkStep := {} -- We insert entries into `rTable` through two different ways -- 1. Calling `newChkStep` -- 2. Validness facts from the ImportTable are treated in `newAssertions` @@ -110,7 +110,7 @@ def getLamEVarTy : ReifM (Array LamSort) := do return (← getRst).lamEVarTy def getLamEVarTyTree : ReifM (BinTree LamSort) := do return (← getRst).lamEVarTyTree -def getChkMap : ReifM (HashMap REntry ChkStep) := do return (← getRst).chkMap +def getChkMap : ReifM (Std.HashMap REntry ChkStep) := do return (← getRst).chkMap def checkerStats : ReifM (Array (String × Nat)) := do return #[("tyVal", (← getTyVal).size), ("varVal", (← getVarVal).size), @@ -120,7 +120,7 @@ def checkerStats : ReifM (Array (String × Nat)) := do def printCheckerStats : ReifM Unit := do let stats := (← checkerStats).map (fun (s, n) => " " ++ s ++ s!": {n}") let ss := #["Checker Statistics:"] ++ stats - trace[auto.buildChecker] (String.intercalate "\n" ss.data) + trace[auto.buildChecker] (String.intercalate "\n" ss.toList) def printValuation : ReifM Unit := do let tyVal ← getTyVal @@ -139,30 +139,30 @@ def printValuation : ReifM Unit := do def printProofs : ReifM Unit := do let chkMap ← getChkMap for re in (← getRTable) do - if let .some cs := chkMap.find? re then + if let .some cs := chkMap.get? re then let .some n := (← getRst).findPos? re | throwError "printProofs :: Unexpected error" let etoms := - match (← getChkStepCache).find? cs with + match (← getChkStepCache).get? cs with | .some (_, arr) => - if arr.size == 0 then m!"" else m!" and produces etoms {arr.data}" + if arr.size == 0 then m!"" else m!" and produces etoms {arr.toList}" | .none => m!"" trace[auto.lamReif.printProofs] "{n} : ChkStep ⦗⦗{cs}⦘⦘ proves ⦗⦗{re}⦘⦘{etoms}" else match re with | .valid [] t => - let .some (expr, _, n) := (← getAssertions).find? t + let .some (expr, _, n) := (← getAssertions).get? t | throwError "printProofs :: Unable to find assertion associated with {t}" trace[auto.lamReif.printProofs] "{n} : External fact ⦗⦗{← Meta.inferType expr}⦘⦘ proves ⦗⦗{re}⦘⦘" | .nonempty s => - let .some (expr, _, n) := (← getInhabitations).find? s + let .some (expr, _, n) := (← getInhabitations).get? s | throwError "printProofs :: Unable to find inhabitation fact associated with {s}" trace[auto.lamReif.printProofs] "{n} : Inhabitation fact ⦗⦗{← Meta.inferType expr}⦘⦘ proves ⦗⦗{re}⦘⦘" | _ => throwError "printProofs :: Unexpected entry {re}" def sort2LamILTyIdx (s : LamSort) : ReifM Nat := do let isomTyMap ← getIsomTyMap - match isomTyMap.find? s with + match isomTyMap.get? s with | .some n => return n | .none => let lamILTy ← getLamILTy @@ -191,7 +191,7 @@ def lookupLamILTy! (idx : Nat) : ReifM LamSort := do throwError "lookupLamILTy! :: Unknown index {idx}" def lookupAssertion! (t : LamTerm) : ReifM (Expr × DTr × LamTerm × Nat) := do - if let .some r := (← getAssertions).find? t then + if let .some r := (← getAssertions).get? t then return r else throwError "lookupAssertion! :: Unknown assertion {t}" @@ -208,11 +208,11 @@ def lookupREntryPos! (re : REntry) : ReifM Nat := do | .none => match re with | .valid [] t => - match (← getAssertions).find? t with + match (← getAssertions).get? t with | .some (_, _, _, n) => return n | .none => throwError "lookupREntryPos! :: Unknown REntry {re}" | .nonempty s => - match (← getInhabitations).find? s with + match (← getInhabitations).get? s with | .some (_, _, n) => return n | .none => throwError "lookupREntryPos! :: Unknown REntry {re}" | _ => throwError "lookupREntryPos! :: Unknown REntry {re}" @@ -223,16 +223,16 @@ inductive REntryProof where | assertion : Expr → DTr → LamTerm → REntryProof def lookupREntryProof? (re : REntry) : ReifM (Option REntryProof) := do - match (← getChkMap).find? re with + match (← getChkMap).get? re with | .some cs => return .some (.chkStep cs) | .none => match re with | .valid [] t => - match (← getAssertions).find? t with + match (← getAssertions).get? t with | .some (e, deriv, t, _) => return .some (.assertion e deriv t) | .none => return .none | .nonempty s => - match (← getInhabitations).find? s with + match (← getInhabitations).get? s with | .some (e, deriv, _) => return .some (.inhabitation e deriv s) | .none => return .none | _ => return .none @@ -249,19 +249,19 @@ def lookupLamEVarTy! (idx : Nat) : ReifM LamSort := do throwError "lookupLamEVarTy! :: Unknown etom {idx}" def lookupChkStepEtom! (cs : ChkStep) : ReifM (Array Nat) := do - if let .some (_, arr) := (← getChkStepCache).find? cs then + if let .some (_, arr) := (← getChkStepCache).get? cs then return arr else throwError "lookupChkStepEtom! :: ChkStep {cs} did not produce new etom" def lookupChkStepResult! (cs : ChkStep) : ReifM EvalResult := do - if let .some (er, _) := (← getChkStepCache).find? cs then + if let .some (er, _) := (← getChkStepCache).get? cs then return er else throwError "lookupChkStepEtom! :: ChkStep {cs} did not produce new etom" def lookupEtomChkStep! (eidx : Nat) : ReifM ChkStep := do - if let .some c := (← getEtomChkStep).find? eidx then + if let .some c := (← getEtomChkStep).get? eidx then return c else throwError "lookupEtomChkStep! :: Unknown etom {eidx}" @@ -337,7 +337,7 @@ def newChkStep (c : ChkStep) (res? : Option EvalResult) : ReifM (Bool × EvalRes -- twice may yield different results. So, we lookup -- `chkStepCache` to see whether this checkstep has -- been evaluated before - if let .some (res, _) := (← getChkStepCache).find? c then + if let .some (res, _) := (← getChkStepCache).get? c then return (false, res) let ltv ← getLamTyValAtMeta let res := c.eval ltv.lamVarTy ltv.lamILTy ⟨← getRTableTree, ← getMaxEVarSucc, ← getLamEVarTyTree⟩ @@ -694,7 +694,7 @@ section Checker def validOfBVarLowers (v : REntry) (ns : Array REntry) : ReifM REntry := do let pv ← lookupREntryPos! v let pns ← ns.mapM lookupREntryPos! - let (_, .addEntry re) ← newChkStep (.i (.validOfBVarLowers pv pns.data)) .none + let (_, .addEntry re) ← newChkStep (.i (.validOfBVarLowers pv pns.toList)) .none | throwError "validOfBVarLowers :: Unexpected evaluation result" return re @@ -716,20 +716,20 @@ section Checker def validOfImps (impV : REntry) (hypVs : Array REntry) : ReifM REntry := do let imp ← lookupREntryPos! impV let ps ← hypVs.mapM lookupREntryPos! - let (_, .addEntry re) ← newChkStep (.i (.validOfImps imp ps.data)) .none + let (_, .addEntry re) ← newChkStep (.i (.validOfImps imp ps.toList)) .none | throwError "validOfImps :: Unexpected evaluation result" return re /-- Repeated instantiation -/ def validOfInstantiate (v : REntry) (args : Array LamTerm) : ReifM REntry := do let p ← lookupREntryPos! v - let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiate p args.data)) .none + let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiate p args.toList)) .none | throwError "validOfInstantiate :: Unexpected evaluation result" return re def validOfInstantiateRev (v : REntry) (args : Array LamTerm) : ReifM REntry := do let p ← lookupREntryPos! v - let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiateRev p args.data)) .none + let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiateRev p args.toList)) .none | throwError "validOfInstantiateRev :: Unexpected evaluation result" return re @@ -980,7 +980,7 @@ def newTypeExpr (e : Expr) : ReifM LamSort := do def processTypeExpr (e : Expr) : ReifM LamSort := do let tyVarMap ← getTyVarMap let e ← Reif.resolveTy e - if let .some idx := tyVarMap.find? e then + if let .some idx := tyVarMap.get? e then return .atom idx match e with | .sort lvl => @@ -1019,16 +1019,16 @@ def newTermExpr (e : Expr) : ReifM LamTerm := do let lamTy ← reifType eTy newTermExprAux e lamTy -def reifMapIL : HashMap Name (LamSort → LamBaseTerm) := - HashMap.ofList [ +def reifMapIL : Std.HashMap Name (LamSort → LamBaseTerm) := + Std.HashMap.ofList [ (``Eq, .eq), (``Embedding.forallF, .forallE), (``Exists, .existE), (``Bool.ite', .ite) ] -def reifMapConstNilLvl : HashMap Name LamTerm := - HashMap.ofList [ +def reifMapConstNilLvl : Std.HashMap Name LamTerm := + Std.HashMap.ofList [ (``True, .base .trueE), (``False, .base .falseE), (``Not, .base .not), @@ -1057,8 +1057,8 @@ def reifMapConstNilLvl : HashMap Name LamTerm := (``Int.add, .base .iadd), (``Int.sub, .base .isub), (``Int.mul, .base .imul), - (``Int.div, .base .idiv), - (``Int.mod, .base .imod), + (``Int.tdiv, .base .idiv), + (``Int.tmod, .base .imod), (``Int.ediv, .base .iediv), (``Int.emod, .base .iemod), (``Int.le, .base .ile), @@ -1070,8 +1070,8 @@ def reifMapConstNilLvl : HashMap Name LamTerm := (``String.replace, .base .srepall) ] -def reifMapBVConst1 : HashMap Name (Nat → LamTerm) := - HashMap.ofList [ +def reifMapBVConst1 : Std.HashMap Name (Nat → LamTerm) := + Std.HashMap.ofList [ (``BitVec.ofNat, fun n => .base (.bvofNat n)), (``BitVec.toNat, fun n => .base (.bvtoNat n)), (``BitVec.ofInt, fun n => .base (.bvofInt n)), @@ -1102,21 +1102,22 @@ def reifMapBVConst1 : HashMap Name (Nat → LamTerm) := (``BitVec.rotateRight, fun n => .bvrotateRight n) ] -def reifMapBVConst2 : HashMap Name (Nat → Nat → LamTerm) := - HashMap.ofList [ +def reifMapBVConst2 : Std.HashMap Name (Nat → Nat → LamTerm) := + Std.HashMap.ofList [ (``BitVec.append, fun n m => .base (.bvappend n m)), (``BitVec.replicate, fun w i => .base (.bvrepeat w i)), (``BitVec.zeroExtend, fun w v => .base (.bvzeroExtend w v)), + (``BitVec.setWidth, fun w v => .base (.bvzeroExtend w v)), (``BitVec.signExtend, fun w v => .base (.bvsignExtend w v)) ] -def reifMapBVConst3 : HashMap Name (Nat → Nat → Nat → LamTerm) := - HashMap.ofList [ +def reifMapBVConst3 : Std.HashMap Name (Nat → Nat → Nat → LamTerm) := + Std.HashMap.ofList [ (``BitVec.extractLsb, fun n h l => .base (.bvextract n h l)) ] -def reifMapAttributesProp : HashMap Name String := - HashMap.ofList [ +def reifMapAttributesProp : Std.HashMap Name String := + Std.HashMap.ofList [ (``SMT.Attribute.trigger, "pattern") ] @@ -1126,7 +1127,7 @@ def processSimpleLit (l : Literal) : LamTerm := | .strVal s => .base (.strVal s) def processSimpleConst (name : Name) (lvls : List Level) : ReifM (Option LamTerm) := do - if let .some t := reifMapConstNilLvl.find? name then + if let .some t := reifMapConstNilLvl.get? name then if lvls.length != 0 then throwError "processSimpleConst :: ConstNilLvl constants should have nil level list" return t @@ -1144,20 +1145,20 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do let fn := fn.getAppFn let .const name lvls := fn | return .none - match args.data with + match args.toList with | [] => throwError "processSimpleApp :: Unexpected error" | [arg] => - if let .some tcon := reifMapBVConst1.find? name then + if let .some tcon := reifMapBVConst1.get? name then if lvls.length != 0 then throwError "processSimpleApp :: BVConst1 should have nil level list" if let .some n ← @id (MetaM _) (Meta.evalNat arg) then return .some (tcon n) return .none - if let .some attrName := reifMapAttributesProp.find? name then + if let .some attrName := reifMapAttributesProp.get? name then if lvls.length != 1 then throwError "processSimpleApp :: Attribute should have one level" return .some (.base (.ocst (.smtAttr1T attrName (← reifType arg) (.base .prop)))) - if let .some tcon := reifMapIL.find? name then + if let .some tcon := reifMapIL.get? name then if name == ``Embedding.forallF then let [lvl₁, lvl₂] := lvls | throwError "processSimpleApp :: Auto.Embedding.forallF should have two levels" @@ -1166,7 +1167,7 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do return .some (.base (tcon (← reifType arg))) return .none | [arg₁, arg₂] => - if let .some tcon := reifMapBVConst2.find? name then + if let .some tcon := reifMapBVConst2.get? name then if lvls.length != 0 then throwError "processSimpleApp :: BVConst2 should have nil level list" match ← @id (MetaM _) (Meta.evalNat arg₁), @@ -1175,7 +1176,7 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do | _, _ => return .none return .none | [arg₁, arg₂, arg₃] => - if let .some tcon := reifMapBVConst3.find? name then + if let .some tcon := reifMapBVConst3.get? name then if lvls.length != 0 then throwError "processSimpleApp :: BVConst2 should have nil level list" match ← @id (MetaM _) (Meta.evalNat arg₁), @@ -1191,8 +1192,8 @@ open LamCstrD in fn : .const _ _ arg₁ : .const _ _ -/ -def reifMapLam0Arg2NoLit : HashMap (Name × Name) (Expr × LamTerm) := - HashMap.ofList [ +def reifMapLam0Arg2NoLit : Std.HashMap (Name × Name) (Expr × LamTerm) := + Std.HashMap.ofList [ ((``NatCast.natCast, ``Int), (.const ``Int.ofNat [], .base .iofNat)), ((``Neg.neg, ``Int), (.const ``Int.neg [], .base .ineg)), ((`Abs.abs, ``Int), (.const ``Int.abs [], .base .iabs)), @@ -1219,8 +1220,8 @@ open LamCstrD in fn : .const _ _ arg₁ : .app (.const _ _) natlit -/ -def reifMapLam0Arg2Natlit : HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := - HashMap.ofList [ +def reifMapLam0Arg2Natlit : Std.HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := + Std.HashMap.ofList [ ((``Neg.neg, ``BitVec), #[(fun n => .app (.const ``BitVec.neg []) (.lit (.natVal n)), fun n => .base (.bvneg n))]), ((`Abs.abs, ``BitVec), @@ -1246,8 +1247,8 @@ def reifMapLam0Arg2Natlit : HashMap (Name × Name) (Array ((Nat → Expr) × (Na arg₁ : .const _ _ arg₂ : .const _ _ -/ -def reifMapLam0Arg4NoLit : HashMap (Name × Name × Name) (Expr × LamTerm) := - HashMap.ofList [ +def reifMapLam0Arg4NoLit : Std.HashMap (Name × Name × Name) (Expr × LamTerm) := + Std.HashMap.ofList [ ((``HAdd.hAdd, ``Nat, ``Nat), (.const ``Nat.add [], .base .nadd)), ((``HAdd.hAdd, ``Int, ``Int), (.const ``Int.add [], .base .iadd)), ((``HSub.hSub, ``Nat, ``Nat), (.const ``Nat.sub [], .base .nsub)), @@ -1255,7 +1256,7 @@ def reifMapLam0Arg4NoLit : HashMap (Name × Name × Name) (Expr × LamTerm) := ((``HMul.hMul, ``Nat, ``Nat), (.const ``Nat.mul [], .base .nmul)), ((``HMul.hMul, ``Int, ``Int), (.const ``Int.mul [], .base .imul)), ((``HDiv.hDiv, ``Nat, ``Nat), (.const ``Nat.div [], .base .ndiv)), - ((``HDiv.hDiv, ``Int, ``Int), (.const ``Int.div [], .base .idiv)), + ((``HDiv.hDiv, ``Int, ``Int), (.const ``Int.tdiv [], .base .idiv)), ((``HAppend.hAppend, ``String, ``String), (.const ``String.append [], .base .sapp)) ] @@ -1265,8 +1266,8 @@ def reifMapLam0Arg4NoLit : HashMap (Name × Name × Name) (Expr × LamTerm) := arg₂ : .app (.const _ _) natlit₂ |- natlit₁ = natlit₂ -/ -def reifMapLam0Arg4NatLitNatLitEq : HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := - HashMap.ofList [ +def reifMapLam0Arg4NatLitNatLitEq : Std.HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := + Std.HashMap.ofList [ ((``HAdd.hAdd, ``BitVec), #[(fun n => .app (.const ``BitVec.add []) (.lit (.natVal n)), fun n => .base (.bvadd n))]), ((``HSub.hSub, ``BitVec), @@ -1291,8 +1292,8 @@ open LamCstrD in arg₂ : .app (.const _ _) natlit₂ |- natlit₁ ?? natlit₂ -/ -def reifMapLam0Arg4NatLitNatLitH : HashMap (Name × Name) (Array ((Nat → Nat → Expr) × (Nat → Nat → LamTerm))) := - HashMap.ofList [ +def reifMapLam0Arg4NatLitNatLitH : Std.HashMap (Name × Name) (Array ((Nat → Nat → Expr) × (Nat → Nat → LamTerm))) := + Std.HashMap.ofList [ ((``HAppend.hAppend, ``BitVec), #[(fun n m => mkApp2 (.const ``BitVec.append []) (.lit (.natVal n)) (.lit (.natVal m)), fun n m => .base (.bvappend n m))]), ((``HShiftLeft.hShiftLeft, ``BitVec), @@ -1307,8 +1308,8 @@ def reifMapLam0Arg4NatLitNatLitH : HashMap (Name × Name) (Array ((Nat → Nat arg₁ : .app (.const _ _) natlit arg₂ : .const _ _ -/ -def reifMapLam0Arg4NatLit : HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := - HashMap.ofList [ +def reifMapLam0Arg4NatLit : Std.HashMap (Name × Name) (Array ((Nat → Expr) × (Nat → LamTerm))) := + Std.HashMap.ofList [ ((``HShiftLeft.hShiftLeft, ``BitVec), #[(fun n => .app (.const ``BitVec.shiftLeft []) (.lit (.natVal n)), fun n => .base (.bvshl n))]), ((``HShiftRight.hShiftRight, ``BitVec), @@ -1322,14 +1323,14 @@ def processLam0Arg2 (e fn arg₁ arg₂ : Expr) : MetaM (Option LamTerm) := do if arg₁.isConst then let .const arg₁Name _ := arg₁ | throwError "processLam0Arg2 :: Unexpected error" - if let .some (e', t) := reifMapLam0Arg2NoLit.find? (fnName, arg₁Name) then + if let .some (e', t) := reifMapLam0Arg2NoLit.get? (fnName, arg₁Name) then if (← Meta.isDefEqD e e') then return .some t if arg₁.isApp then let .app arg₁fn arg₁arg := arg₁ | throwError "processLam0Arg2 :: Unexpected error" if let .const arg₁FnName _ := arg₁fn then - if let .some candidates := reifMapLam0Arg2Natlit.find? (fnName, arg₁FnName) then + if let .some candidates := reifMapLam0Arg2Natlit.get? (fnName, arg₁FnName) then if let .some n ← Meta.evalNat arg₁arg then for (e'con, tcon) in candidates do if (← Meta.isDefEqD e (e'con n)) then @@ -1370,7 +1371,7 @@ def processLam0Arg4 (e fn arg₁ arg₂ arg₃ arg₄ : Expr) : MetaM (Option La | throwError "processLam0Arg4 :: Unexpected error" let .const arg₂name _ := arg₂ | throwError "processLam0Arg4 :: Unexpected error" - if let .some (e', t) := reifMapLam0Arg4NoLit.find? (fnName, arg₁name, arg₂name) then + if let .some (e', t) := reifMapLam0Arg4NoLit.get? (fnName, arg₁name, arg₂name) then if (← Meta.isDefEqD e e') then return .some t return .none @@ -1380,7 +1381,7 @@ def processLam0Arg4 (e fn arg₁ arg₂ arg₃ arg₄ : Expr) : MetaM (Option La if arg₁fn.isConst then let .const arg₁fnName _ := arg₁fn | throwError "processLam0Arg4 :: Unexpected error {arg₁fn}" - if let .some candidates := reifMapLam0Arg4NatLit.find? (fnName, arg₁fnName) then + if let .some candidates := reifMapLam0Arg4NatLit.get? (fnName, arg₁fnName) then for (e'con, tcon) in candidates do if let .some n ← Meta.evalNat arg₁arg then if (← Meta.isDefEqD e (e'con n)) then @@ -1393,12 +1394,12 @@ def processLam0Arg4 (e fn arg₁ arg₂ arg₃ arg₄ : Expr) : MetaM (Option La if arg₁fn.isConst && arg₂fn.isConst then let .const arg₁fnName _ := arg₁fn | throwError "processLam0Arg4 :: Unexpected error" - if let .some candidates := reifMapLam0Arg4NatLitNatLitEq.find? (fnName, arg₁fnName) then + if let .some candidates := reifMapLam0Arg4NatLitNatLitEq.get? (fnName, arg₁fnName) then for (e'con, tcon) in candidates do if let .some n ← Meta.evalNat arg₁arg then if (← Meta.isDefEqD e (e'con n)) then return tcon n - if let .some candidates := reifMapLam0Arg4NatLitNatLitH.find? (fnName, arg₁fnName) then + if let .some candidates := reifMapLam0Arg4NatLitNatLitH.get? (fnName, arg₁fnName) then for (e'con, tcon) in candidates do match ← Meta.evalNat arg₁arg, ← Meta.evalNat arg₂arg with | .some n , .some m => @@ -1415,7 +1416,7 @@ def processComplexTermExpr (e : Expr) : MetaM (Option LamTerm) := do | 0 => let fn := e.getAppFn let args := e.getAppArgs - match args.data with + match args.toList with | [] => return .none | [_] => return .none | [arg₁, arg₂] => processLam0Arg2 e fn arg₁ arg₂ @@ -1444,24 +1445,24 @@ where return res newTermExpr e -def processTermExpr (lctx : HashMap FVarId Nat) (e : Expr) : ReifM LamTerm := do +def processTermExpr (lctx : Std.HashMap FVarId Nat) (e : Expr) : ReifM LamTerm := do if let .fvar fid := e then if let .some n := deBruijn? lctx fid then return .bvar n let e ← Reif.resolveVal e let varMap ← getVarMap -- If the expression has already been processed - if let .some id := varMap.find? e then + if let .some id := varMap.get? e then return .atom id -- If the expression has not been processed processNewTermExpr e where - deBruijn? (lctx : HashMap FVarId Nat) (id : FVarId) : Option Nat := - match lctx.find? id with + deBruijn? (lctx : Std.HashMap FVarId Nat) (id : FVarId) : Option Nat := + match lctx.get? id with | .some n => lctx.size - 1 - n | .none => none -partial def reifTerm (lctx : HashMap FVarId Nat) : Expr → ReifM LamTerm +partial def reifTerm (lctx : Std.HashMap FVarId Nat) : Expr → ReifM LamTerm | .app fn arg => do let lamFn ← reifTerm lctx fn let lamArg ← reifTerm lctx arg @@ -1512,7 +1513,7 @@ def reifInd (ind : SimpleIndVal) : ReifM (Option IndInfo) := do let rty ← reifType type let rctors ← ctors.mapM (fun (e, _) => reifTermCheckType e) let rprojs ← projs.mapM (fun ps => ps.mapM reifTermCheckType) - let ret := ⟨rty, rctors.data, rprojs.bind (·.data)⟩ + let ret := ⟨rty, rctors.toList, rprojs.bind (·.toList)⟩ trace[auto.lamReif.printResult] "Successfully reified inductive info {← ind.zetaReduce} to {ret}" return .some ret @@ -1522,7 +1523,7 @@ def reifMutInd (mind : Array SimpleIndVal) : ReifM (Option MutualIndInfo) := do let .some ii ← reifInd ind | return .none ret := ret.push ii - return ret.data + return ret.toList def reifMutInds (minds : Array (Array SimpleIndVal)) : ReifM (Array MutualIndInfo) := do let mut ret := #[] @@ -1575,12 +1576,12 @@ section BuildChecker let rst ← getRst let mut chkSteps := #[] for re in (← getRTable) do - if let .some cs := rst.chkMap.find? re then + if let .some cs := rst.chkMap.get? re then let .some n := rst.findPos? re | throwError "buildChkStepsExpr :: Unexpected error" chkSteps := chkSteps.push (cs, n) -- `ChkMap` are run using `foldl`, so we use `BinTree.ofListFoldl` - let e := Lean.toExpr (BinTree.ofListFoldl chkSteps.data) + let e := Lean.toExpr (BinTree.ofListFoldl chkSteps.toList) return e -- Given a list of expression of type `ty`, construct the corresponding `BinTree` @@ -1594,7 +1595,7 @@ section BuildChecker def buildTyVal : ReifM Expr := do let u ← getU -- `tyVal : List (Type u)` - let tyVal : List Expr := (← getTyVal).data.map (fun (e, lvl) => + let tyVal : List Expr := (← getTyVal).toList.map (fun (e, lvl) => Expr.app (.const ``Embedding.GLift [lvl, u]) e) -- `tyValExpr : Nat → Type u` let tyValExpr := exprListToLCtx tyVal (.succ u) (.sort (.succ u)) (.sort u) @@ -1604,7 +1605,7 @@ section BuildChecker def buildVarExpr (tyValExpr : Expr) : ReifM Expr := do let u ← getU let lamSortExpr := Expr.const ``LamSort [] - let varValPair := (← getVarVal).data + let varValPair := (← getVarVal).toList let vars ← varValPair.mapM (fun (e, s) => do let sExpr := toExpr s return Lean.mkApp3 (.const ``varSigmaMk [u]) tyValExpr sExpr (← varValInterp e s)) @@ -1616,7 +1617,7 @@ section BuildChecker def buildILExpr (tyValExpr : Expr) : ReifM Expr := do let u ← getU let lamSortExpr := Expr.const ``LamSort [] - let lamILTy := (← getLamILTy).data + let lamILTy := (← getLamILTy).toList -- `ils : List ((s : LamSort) × ILLift.{u} (s.interp tyVal))` let ils ← lamILTy.mapM (fun s => do let sExpr := toExpr s @@ -1704,8 +1705,8 @@ section BuildChecker printCheckerStats let startTime ← IO.monoMsNow let u ← getU - let lvtExpr := Lean.toExpr (BinTree.ofListGet ((← getVarVal).map Prod.snd).data) - let litExpr := Lean.toExpr (BinTree.ofListGet (← getLamILTy).data) + let lvtExpr := Lean.toExpr (BinTree.ofListGet ((← getVarVal).map Prod.snd).toList) + let litExpr := Lean.toExpr (BinTree.ofListGet (← getLamILTy).toList) let cpvExpr ← buildCPValExpr let cpvTy := Expr.const ``CPVal [u] let checker ← Meta.withLetDecl `cpval cpvTy cpvExpr fun cpvFVarExpr => do @@ -1757,9 +1758,9 @@ section BuildChecker printCheckerStats let startTime ← IO.monoMsNow let u ← getU - let lvtExpr := Lean.toExpr (BinTree.ofListGet ((← getVarVal).map Prod.snd).data) + let lvtExpr := Lean.toExpr (BinTree.ofListGet ((← getVarVal).map Prod.snd).toList) let lvtNativeName ← mkNativeAuxDecl `lam_ssrefl_lvt (Expr.app (.const ``BinTree [.zero]) (Lean.mkConst ``LamSort)) lvtExpr - let litExpr := Lean.toExpr (BinTree.ofListGet (← getLamILTy).data) + let litExpr := Lean.toExpr (BinTree.ofListGet (← getLamILTy).toList) let litNativeName ← mkNativeAuxDecl `lam_ssrefl_lit (Expr.app (.const ``BinTree [.zero]) (Lean.mkConst ``LamSort)) litExpr let cpvExpr ← buildCPValExpr let cpvTy := Expr.const ``CPVal [u] @@ -1804,16 +1805,16 @@ namespace Lam2Lam open Embedding.Lam LamReif structure TState where - typeH2lMap : HashMap Nat Nat := {} + typeH2lMap : Std.HashMap Nat Nat := {} -- This field should be in sync with `LamReif.State.tyVal` typeL2hMap : Array Nat := #[] - termH2lMap : HashMap Nat Nat := {} + termH2lMap : Std.HashMap Nat Nat := {} -- This field should be in sync with `LamReif.State.varVal` termL2hMap : Array Nat := #[] - etomH2lMap : HashMap Nat Nat := {} - etomL2hMap : HashMap Nat Nat := {} + etomH2lMap : Std.HashMap Nat Nat := {} + etomL2hMap : Std.HashMap Nat Nat := {} -- Maps from high-level chkstep to low-level chkstep - csH2lMap : HashMap ChkStep ChkStep := {} + csH2lMap : Std.HashMap ChkStep ChkStep := {} abbrev TransM := StateRefT TState ReifM @@ -1833,7 +1834,7 @@ open Embedding.Lam LamReif def transTypeAtom (a : Nat) (val : Expr × Level) : TransM Nat := do let typeH2lMap ← getTypeH2lMap - match typeH2lMap.find? a with + match typeH2lMap.get? a with | .some n => return n | .none => let idx := typeH2lMap.size @@ -1849,7 +1850,7 @@ open Embedding.Lam LamReif -/ def transTermAtom (a : Nat) (val : Expr × LamSort) : TransM Nat := do let termH2lMap ← getTermH2lMap - match termH2lMap.find? a with + match termH2lMap.get? a with | .some n => return n | .none => let idx := termH2lMap.size @@ -1893,7 +1894,7 @@ open Embedding.Lam LamReif let (cs, _) ← (lookupEtomChkStep! e).run ref trace[auto.buildChecker] "Collecting for etom {e} by ChkStep {cs}" let _ ← processChkStep ref cs - let .some n := (← getEtomH2lMap).find? e + let .some n := (← getEtomH2lMap).get? e | throwError "transEtom :: Cannot find translation of etom {e}" return n @@ -1993,7 +1994,7 @@ open Embedding.Lam LamReif | .wfOfBetaBounded pos bound => return .wfOfBetaBounded (← transPos ref pos) bound partial def processChkStep (ref : State) (cs : ChkStep) : TransM EvalResult := do - if let .some cs' := (← getCsH2lMap).find? cs then + if let .some cs' := (← getCsH2lMap).get? cs then return ← LamReif.lookupChkStepResult! cs' let cs' ← transChkStep ref cs setCsH2lMap ((← getCsH2lMap).insert cs cs') @@ -2012,7 +2013,7 @@ open Embedding.Lam LamReif -- Collect essential chksteps and assertions from the high-level `lam` -- into the low-level `lam` such that the low-level `lam` proves `re` partial def collectProofFor (ref : State) (hre : REntry) : TransM Unit := do - if let .some _ := (← getChkMap).find? hre then + if let .some _ := (← getChkMap).get? hre then return let (highLvlProof, _) ← (lookupREntryProof! hre).run ref match highLvlProof with diff --git a/Auto/Translation/LamUtils.lean b/Auto/Translation/LamUtils.lean index 5885c2e..e71b816 100644 --- a/Auto/Translation/LamUtils.lean +++ b/Auto/Translation/LamUtils.lean @@ -48,24 +48,24 @@ namespace LamExportUtils "Import versions of polymorphic logical " ++ "constants should have been eliminated" - def collectLamSortAtoms : LamSort → HashSet Nat - | .atom n => HashSet.empty.insert n - | .base _ => HashSet.empty + def collectLamSortAtoms : LamSort → Std.HashSet Nat + | .atom n => Std.HashSet.empty.insert n + | .base _ => Std.HashSet.empty | .func a b => (collectLamSortAtoms a).insertMany (collectLamSortAtoms b) - def collectLamSortsAtoms (ss : Array LamSort) : HashSet Nat := - ss.foldl (fun hs s => hs.insertMany (collectLamSortAtoms s)) HashSet.empty + def collectLamSortsAtoms (ss : Array LamSort) : Std.HashSet Nat := + ss.foldl (fun hs s => hs.insertMany (collectLamSortAtoms s)) Std.HashSet.empty - def collectLamSortBitVecLengths : LamSort → HashSet Nat - | .base (.bv n) => HashSet.empty.insert n + def collectLamSortBitVecLengths : LamSort → Std.HashSet Nat + | .base (.bv n) => Std.HashSet.empty.insert n | .func a b => (collectLamSortBitVecLengths a).insertMany (collectLamSortBitVecLengths b) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamSortsBitVecLengths (ss : Array LamSort) : HashSet Nat := - ss.foldl (fun hs s => hs.insertMany (collectLamSortBitVecLengths s)) HashSet.empty + def collectLamSortsBitVecLengths (ss : Array LamSort) : Std.HashSet Nat := + ss.foldl (fun hs s => hs.insertMany (collectLamSortBitVecLengths s)) Std.HashSet.empty /-- Collect type atoms in a LamBaseTerm -/ - def collectLamBaseTermAtoms (b : LamBaseTerm) : CoreM (HashSet Nat) := do + def collectLamBaseTermAtoms (b : LamBaseTerm) : CoreM (Std.HashSet Nat) := do let s? : Option LamSort ← (do match b with | .eqI _ => throwError ("collectAtoms :: " ++ exportError.ImpPolyLog) @@ -80,7 +80,7 @@ namespace LamExportUtils if let .some s := s? then return collectLamSortAtoms s else - return HashSet.empty + return Std.HashSet.empty /-- The first hashset is the type atoms @@ -92,18 +92,18 @@ namespace LamExportUtils does not occur in the `LamTerm` -/ def collectLamTermAtoms (lamVarTy : Array LamSort) (lamEVarTy : Array LamSort) : - LamTerm → CoreM (HashSet Nat × HashSet Nat × HashSet Nat) + LamTerm → CoreM (Std.HashSet Nat × Std.HashSet Nat × Std.HashSet Nat) | .atom n => do let .some s := lamVarTy[n]? | throwError "collectAtoms :: Unknown term atom {n}" - return (collectLamSortAtoms s, HashSet.empty.insert n, HashSet.empty) + return (collectLamSortAtoms s, Std.HashSet.empty.insert n, Std.HashSet.empty) | .etom n => do let .some s := lamEVarTy[n]? | throwError "collectAtoms :: Unknown etom {n}" - return (collectLamSortAtoms s, HashSet.empty, HashSet.empty.insert n) + return (collectLamSortAtoms s, Std.HashSet.empty, Std.HashSet.empty.insert n) | .base b => do - return (← collectLamBaseTermAtoms b, HashSet.empty, HashSet.empty) - | .bvar _ => pure (HashSet.empty, HashSet.empty, HashSet.empty) + return (← collectLamBaseTermAtoms b, Std.HashSet.empty, Std.HashSet.empty) + | .bvar _ => pure (Std.HashSet.empty, Std.HashSet.empty, Std.HashSet.empty) | .lam s t => do let (typeHs, termHs, etomHs) ← collectLamTermAtoms lamVarTy lamEVarTy t let sHs := collectLamSortAtoms s @@ -116,56 +116,56 @@ namespace LamExportUtils mergeHashSet etomHs₁ etomHs₂) def collectLamTermsAtoms (lamVarTy : Array LamSort) (lamEVarTy : Array LamSort) - (ts : Array LamTerm) : CoreM (HashSet Nat × HashSet Nat × HashSet Nat) := + (ts : Array LamTerm) : CoreM (Std.HashSet Nat × Std.HashSet Nat × Std.HashSet Nat) := ts.foldlM (fun (tyHs, aHs, eHs) t => do let (tyHs', aHs', eHs') ← collectLamTermAtoms lamVarTy lamEVarTy t return (mergeHashSet tyHs tyHs', mergeHashSet aHs aHs', mergeHashSet eHs eHs')) - (HashSet.empty, HashSet.empty, HashSet.empty) + (Std.HashSet.empty, Std.HashSet.empty, Std.HashSet.empty) - def collectLamTermNatConsts : LamTerm → HashSet NatConst - | .base (.ncst nc) => HashSet.empty.insert nc + def collectLamTermNatConsts : LamTerm → Std.HashSet NatConst + | .base (.ncst nc) => Std.HashSet.empty.insert nc | .lam _ body => collectLamTermNatConsts body | .app _ fn arg => mergeHashSet (collectLamTermNatConsts fn) (collectLamTermNatConsts arg) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamTermsNatConsts (ts : Array LamTerm) : HashSet NatConst := - ts.foldl (fun hs t => mergeHashSet hs (collectLamTermNatConsts t)) HashSet.empty + def collectLamTermsNatConsts (ts : Array LamTerm) : Std.HashSet NatConst := + ts.foldl (fun hs t => mergeHashSet hs (collectLamTermNatConsts t)) Std.HashSet.empty - def collectLamTermIntConsts : LamTerm → HashSet IntConst - | .base (.icst ic) => HashSet.empty.insert ic + def collectLamTermIntConsts : LamTerm → Std.HashSet IntConst + | .base (.icst ic) => Std.HashSet.empty.insert ic | .lam _ body => collectLamTermIntConsts body | .app _ fn arg => mergeHashSet (collectLamTermIntConsts fn) (collectLamTermIntConsts arg) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamTermsIntConsts (ts : Array LamTerm) : HashSet IntConst := - ts.foldl (fun hs t => mergeHashSet hs (collectLamTermIntConsts t)) HashSet.empty + def collectLamTermsIntConsts (ts : Array LamTerm) : Std.HashSet IntConst := + ts.foldl (fun hs t => mergeHashSet hs (collectLamTermIntConsts t)) Std.HashSet.empty - def collectLamTermStringConsts : LamTerm → HashSet StringConst - | .base (.scst sc) => HashSet.empty.insert sc + def collectLamTermStringConsts : LamTerm → Std.HashSet StringConst + | .base (.scst sc) => Std.HashSet.empty.insert sc | .lam _ body => collectLamTermStringConsts body | .app _ fn arg => mergeHashSet (collectLamTermStringConsts fn) (collectLamTermStringConsts arg) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamTermsStringConsts (ts : Array LamTerm) : HashSet StringConst := - ts.foldl (fun hs t => mergeHashSet hs (collectLamTermStringConsts t)) HashSet.empty + def collectLamTermsStringConsts (ts : Array LamTerm) : Std.HashSet StringConst := + ts.foldl (fun hs t => mergeHashSet hs (collectLamTermStringConsts t)) Std.HashSet.empty - def collectLamTermBitvecs : LamTerm → HashSet BitVecConst - | .base (.bvcst bvc) => HashSet.empty.insert bvc + def collectLamTermBitvecs : LamTerm → Std.HashSet BitVecConst + | .base (.bvcst bvc) => Std.HashSet.empty.insert bvc | .lam _ body => collectLamTermBitvecs body | .app _ fn arg => mergeHashSet (collectLamTermBitvecs fn) (collectLamTermBitvecs arg) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamTermsBitvecs (ts : Array LamTerm) : HashSet BitVecConst := - ts.foldl (fun hs t => mergeHashSet hs (collectLamTermBitvecs t)) HashSet.empty + def collectLamTermsBitvecs (ts : Array LamTerm) : Std.HashSet BitVecConst := + ts.foldl (fun hs t => mergeHashSet hs (collectLamTermBitvecs t)) Std.HashSet.empty - def collectLamTermIteSorts : LamTerm → HashSet LamSort - | .base (.ite s) => HashSet.empty.insert s + def collectLamTermIteSorts : LamTerm → Std.HashSet LamSort + | .base (.ite s) => Std.HashSet.empty.insert s | .lam _ body => collectLamTermIteSorts body | .app _ fn arg => mergeHashSet (collectLamTermIteSorts fn) (collectLamTermIteSorts arg) - | _ => HashSet.empty + | _ => Std.HashSet.empty - def collectLamTermsIteSorts (ts : Array LamTerm) : HashSet LamSort := - ts.foldl (fun hs t => mergeHashSet hs (collectLamTermIteSorts t)) HashSet.empty + def collectLamTermsIteSorts (ts : Array LamTerm) : Std.HashSet LamSort := + ts.foldl (fun hs t => mergeHashSet hs (collectLamTermIteSorts t)) Std.HashSet.empty end LamExportUtils @@ -226,8 +226,8 @@ namespace Lam2D | .iadd => return .const ``Int.add [] | .isub => return .const ``Int.sub [] | .imul => return .const ``Int.mul [] - | .idiv => return .const ``Int.div [] - | .imod => return .const ``Int.mod [] + | .idiv => return .const ``Int.tdiv [] + | .imod => return .const ``Int.tmod [] | .iediv => return .const ``Int.ediv [] | .iemod => return .const ``Int.emod [] | .ile => return .const ``Int.le [] @@ -303,16 +303,16 @@ namespace Lam2D Takes a `s : LamSort` and produces the `un-lifted` version of `s.interp` (note that `s.interp` is lifted) -/ - def interpLamSortAsUnlifted (tyVal : HashMap Nat Expr) : LamSort → CoreM Expr + def interpLamSortAsUnlifted (tyVal : Std.HashMap Nat Expr) : LamSort → CoreM Expr | .atom n => do - let .some e := tyVal.find? n + let .some e := tyVal.get? n | throwError "interpLamSortAsUnlifted :: Cannot find fvarId assigned to type atom {n}" return e | .base b => return Lam2D.interpLamBaseSortAsUnlifted b | .func s₁ s₂ => do return .forallE `_ (← interpLamSortAsUnlifted tyVal s₁) (← interpLamSortAsUnlifted tyVal s₂) .default - def interpOtherConstAsUnlifted (tyVal : HashMap Nat Expr) (oc : OtherConst) : MetaM Expr := do + def interpOtherConstAsUnlifted (tyVal : Std.HashMap Nat Expr) (oc : OtherConst) : MetaM Expr := do let .some (.defnInfo constIdVal) := (← getEnv).find? ``constId | throwError "interpOtherConstAsUnlifted :: Unexpected error" let constIdExpr := fun params => constIdVal.value.instantiateLevelParams constIdVal.levelParams params @@ -328,7 +328,7 @@ namespace Lam2D | throwError "interpOtherConstAsUnlifted :: Unexpected sort {sortterm}" return Lean.mkApp2 (constIdExpr [lvlattr, lvlterm]) tyattr tyterm - def interpLamBaseTermAsUnlifted (tyVal : HashMap Nat Expr) : LamBaseTerm → MetaM Expr + def interpLamBaseTermAsUnlifted (tyVal : Std.HashMap Nat Expr) : LamBaseTerm → MetaM Expr | .pcst pc => Lam2D.interpPropConstAsUnlifted pc | .bcst bc => Lam2D.interpBoolConstAsUnlifted bc | .ncst nc => Lam2D.interpNatConstAsUnlifted nc @@ -365,14 +365,14 @@ namespace Lam2D represent `etom`s. -/ def interpLamTermAsUnlifted - (tyVal : HashMap Nat Expr) (varVal : HashMap Nat Expr) (etomVal : HashMap Nat Expr) + (tyVal : Std.HashMap Nat Expr) (varVal : Std.HashMap Nat Expr) (etomVal : Std.HashMap Nat Expr) (lctx : Nat) : LamTerm → MetaM Expr | .atom n => do - let .some e := varVal.find? n + let .some e := varVal.get? n | throwError "interpLamTermAsUnlifted :: Cannot find fvarId assigned to term atom {n}" return e | .etom n => do - let .some efvar := etomVal.find? n + let .some efvar := etomVal.get? n | throwError "interpLamSortAsUnlifted :: Cannot find fvarId assigned to etom {n}" return efvar | .base b => interpLamBaseTermAsUnlifted tyVal b @@ -520,7 +520,7 @@ namespace Lam2D def lamBaseTermSimpNFList : List (Name × Expr) := natConstSimpNFList ++ intConstSimpNFList ++ stringConstSimpNFList ++ bitVecConstSimpNFList - def lamBaseTermSimpNFMap : HashMap Name Expr := HashMap.ofList lamBaseTermSimpNFList + def lamBaseTermSimpNFMap : Std.HashMap Name Expr := Std.HashMap.ofList lamBaseTermSimpNFList section CheckDefEq @@ -539,7 +539,7 @@ namespace Lam2D def approxSimpNF (e : Expr) : CoreM Expr := do let eRep := e.replace (fun sub => match sub with - | .const name _ => lamBaseTermSimpNFMap.find? name + | .const name _ => lamBaseTermSimpNFMap.get? name | _ => .none) let eRep ← Core.betaReduce eRep let replaceNatCast (e : Expr) : Option Expr := diff --git a/Auto/Translation/Monomorphization.lean b/Auto/Translation/Monomorphization.lean index b204b50..12b3f27 100644 --- a/Auto/Translation/Monomorphization.lean +++ b/Auto/Translation/Monomorphization.lean @@ -86,7 +86,7 @@ def CiHead.ofExpr? : Expr → Option CiHead def CiHead.toExpr : CiHead → Expr | .fvar id => .fvar id | .mvar id => .mvar id -| .const name lvls => .const name lvls.data +| .const name lvls => .const name lvls.toList /-- Ignore constant's levels -/ def CiHead.fingerPrint : CiHead → Expr @@ -169,7 +169,7 @@ private def ConstInst.toMessageDataAux (ci : ConstInst) : MessageData := let arr : Array (Option Expr) := Array.mk ((List.range narg).map (fun _ => .none)) let arr := (ci.argsInst.zip ci.argsIdx).foldl (fun acc (arg, idx) => acc.setD idx (.some arg)) arr let arr := arr.map (fun e? => match e? with | .some e => m!" ({e})" | .none => m!" _") - MessageData.intercalate "" arr.data + MessageData.intercalate "" arr.toList instance : ToMessageData ConstInst where toMessageData ci := m!"ConstInst ⦗⦗ {ci.head}{ci.toMessageDataAux} ⦘⦘" @@ -229,8 +229,8 @@ def ConstInst.matchExpr (e : Expr) (ci : ConstInst) : MetaM Bool := do · The expression does not contain level parameters in `params` -/ def ConstInst.ofExpr? (params : Array Name) (bvars : Array Expr) (e : Expr) : MetaM (Option ConstInst) := do - let paramSet := HashSet.empty.insertMany params - let bvarSet := HashSet.empty.insertMany bvars + let paramSet := Std.HashSet.empty.insertMany params + let bvarSet := Std.HashSet.empty.insertMany bvars let fn := e.getAppFn -- If the head contains bound variable, then this is not -- a valid instance @@ -307,7 +307,7 @@ def ConstInst.toExpr (ci : ConstInst) : MetaM Expr := do let mut args : Array (Option Expr) := (Array.mk (List.range nargs)).map (fun n => .none) for (arg, idx) in ci.argsInst.zip ci.argsIdx do args := args.setD idx (.some arg) - let .some ret := ConstInst.toExprAux args.data [] ci.head.toExpr type + let .some ret := ConstInst.toExprAux args.toList [] ci.head.toExpr type | throwError "ConstInst.toExpr :: Unexpected error" return ret @@ -379,11 +379,11 @@ def ConstInsts.canonicalize? (cis : ConstInsts) (ci : ConstInst) : MetaM (Option try to match `e` and the subexpressions of `e` against `ci`. This function is used by `LemmaInst.matchConstInst` only -/ -private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) : Expr → MetaM (HashSet LemmaInst) +private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) : Expr → MetaM (Std.HashSet LemmaInst) | .bvar _ => throwError "MLemmaInst.matchConstInst :: Loose bound variable" | e@(.app ..) => do let args := e.getAppArgs - let mut ret := HashSet.empty + let mut ret := Std.HashSet.empty for arg in args do ret := mergeHashSet ret (← MLemmaInst.matchConstInst ci mi arg) let s ← saveState @@ -406,14 +406,14 @@ private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) | .letE .. => throwError "MLemmaInst.matchConstInst :: Let-expressions should have been reduced" | .mdata .. => throwError "MLemmaInst.matchConstInst :: mdata should have been consumed" | .proj .. => throwError "MLemmaInst.matchConstInst :: Projections should have been turned into ordinary expressions" -| _ => return HashSet.empty +| _ => return Std.HashSet.empty /-- Given a LemmaInst `li` and a ConstInst `ci`, try to match all subexpressions of `li` against `ci` -/ -def LemmaInst.matchConstInst (ci : ConstInst) (li : LemmaInst) : MetaM (HashSet LemmaInst) := +def LemmaInst.matchConstInst (ci : ConstInst) (li : LemmaInst) : MetaM (Std.HashSet LemmaInst) := Meta.withNewMCtxDepth do let (lmvars, mvars, mi) ← MLemmaInst.ofLemmaInst li if lmvars.size == 0 && mvars.size == 0 then - return HashSet.empty + return Std.HashSet.empty MLemmaInst.matchConstInst ci mi mi.type /-- @@ -442,7 +442,7 @@ where | _ => false if hol && (← getMode) == .fol then return false - let fvarSet := HashSet.empty.insertMany fvars + let fvarSet := Std.HashSet.empty.insertMany fvars if ty.hasAnyFVar fvarSet.contains then return false leadingForallQuasiMonomorphicAux (fvars.push xid) bodyi @@ -500,7 +500,7 @@ def LemmaInst.monomorphic? (li : LemmaInst) : MetaM (Option LemmaInst) := do -/ structure State where -- The `Expr` is the fingerprint of the `ConstInst` - ciMap : HashMap Expr ConstInsts := {} + ciMap : Std.HashMap Expr ConstInsts := {} -- The `Expr` is the fingerprint of the `ConstInst` activeCi : Std.Queue (Expr × Nat) := Std.Queue.empty -- During initialization, we supply an array `lemmas` of lemmas @@ -517,9 +517,9 @@ abbrev MonoM := StateRefT State MetaM 2. `(ciMap.find? ci.name).getD #[]` 3. Canonicalized ConstInst -/ -def CiMap.canonicalize? (ciMap : HashMap Expr ConstInsts) (ci : ConstInst) : +def CiMap.canonicalize? (ciMap : Std.HashMap Expr ConstInsts) (ci : ConstInst) : MetaM (Bool × ConstInsts × ConstInst) := do - match ciMap.find? ci.fingerPrint with + match ciMap.get? ci.fingerPrint with | .some insts => match ← insts.canonicalize? ci with | .some ci' => return (true, insts, ci') @@ -570,7 +570,7 @@ def dequeueActiveCi? : MonoM (Option (Expr × Nat)) := do | .none => return .none def lookupActiveCi! (fgp : Expr) (idx : Nat) : MonoM ConstInst := do - let .some cis := (← getCiMap).find? fgp + let .some cis := (← getCiMap).get? fgp | throwError "lookupActiveCi :: Unknown CiHead {fgp}" let .some ci := cis[idx]? | throwError "lookupActiveCi :: Index {idx} out of bound" @@ -656,23 +656,23 @@ namespace FVarRep bfvars : Array FVarId := #[] -- Free variables representing abstracted expressions ffvars : Array FVarId := #[] - exprMap : HashMap Expr FVarId := {} - ciMap : HashMap Expr ConstInsts - ciIdMap : HashMap ConstInst FVarId := {} + exprMap : Std.HashMap Expr FVarId := {} + ciMap : Std.HashMap Expr ConstInsts + ciIdMap : Std.HashMap ConstInst FVarId := {} -- Canonicalization map for types - tyCanMap : HashMap Expr Expr := {} + tyCanMap : Std.HashMap Expr Expr := {} abbrev FVarRepM := StateRefT State MetaState.MetaStateM #genMonadState FVarRepM - def getBfvarSet : FVarRepM (HashSet FVarId) := do + def getBfvarSet : FVarRepM (Std.HashSet FVarId) := do let bfvars ← getBfvars - return HashSet.empty.insertMany bfvars + return Std.HashSet.empty.insertMany bfvars - def getFfvarSet : FVarRepM (HashSet FVarId) := do + def getFfvarSet : FVarRepM (Std.HashSet FVarId) := do let ffvars ← getFfvars - return HashSet.empty.insertMany ffvars + return Std.HashSet.empty.insertMany ffvars /-- Similar to `Monomorphization.processConstInst` -/ def processConstInst (ci : ConstInst) : FVarRepM Unit := do @@ -720,7 +720,7 @@ namespace FVarRep | (true, _, ci') => return ci' | _ => throwError "constInst2FVarId :: Cannot find canonicalized instance of {ci}") let ciIdMap ← FVarRep.getCiIdMap - match ciIdMap.find? ci with + match ciIdMap.get? ci with | .some fid => return fid | .none => do let userName := (`cifvar).appendIndexAfter (← getCiIdMap).size @@ -942,7 +942,7 @@ where let (inductiveVals, s) ← (fvarRepMInductAction inductiveVals).run s let exlis := s.exprMap.toList.map (fun (e, id) => (id, e)) let cilis ← s.ciIdMap.toList.mapM (fun (ci, id) => do return (id, ← MetaState.runMetaM ci.toExpr)) - let polyVal := HashMap.ofList (exlis ++ cilis) + let polyVal := Std.HashMap.ofList (exlis ++ cilis) return (s.ffvars, Reif.State.mk uvalids polyVal s.tyCanMap inhs inductiveVals none) fvarRepMFactAction (lis : Array LemmaInst) : FVarRep.FVarRepM (Array UMonoFact) := lis.filterMapM (fun li => do let liTypeRep? ← FVarRep.replacePolyWithFVar li.type diff --git a/Auto/Translation/Preprocessing.lean b/Auto/Translation/Preprocessing.lean index d257e75..b2430e9 100644 --- a/Auto/Translation/Preprocessing.lean +++ b/Auto/Translation/Preprocessing.lean @@ -60,7 +60,7 @@ def elabDefEq (name : Name) : TacticM (Array Lemma) := do addRecAsLemma val | some (.defnInfo _) => -- Generate definitional equation for (possibly recursive) declaration - match ← getEqnsFor? name (nonRec := true) with + match ← getEqnsFor? name with | some eqns => eqns.mapIdxM fun i eq => do elabLemma (← `($(mkIdent eq))) (.leaf s!"defeq {i.val} {name}") | none => return #[] @@ -99,28 +99,28 @@ def getConstUnfoldInfo (name : Name) : MetaM ConstUnfoldInfo := do If there is cyclic dependency, `topoSortUnfolds` throws an error -/ partial def topoSortUnfolds (unfolds : Array ConstUnfoldInfo) : MetaM (Array ConstUnfoldInfo) := do - let mut depMap : HashMap Name (HashSet Name) := {} + let mut depMap : Std.HashMap Name (Std.HashSet Name) := {} for ⟨i, val, _⟩ in unfolds do for ⟨j, _, _⟩ in unfolds do if (val.find? (fun e => e.constName? == .some j)).isSome then - let deps := (depMap.find? i).getD {} + let deps := (depMap.get? i).getD {} depMap := depMap.insert i (deps.insert j) let (_, _, nameArr) ← (unfolds.mapM (fun n => go depMap {} n.name)).run ({}, #[]) - let nameMap : HashMap Name _ := HashMap.ofList (unfolds.data.map (fun ui => (ui.name, ui))) + let nameMap : Std.HashMap Name _ := Std.HashMap.ofList (unfolds.toList.map (fun ui => (ui.name, ui))) let mut ret := #[] for name in nameArr do - let .some ui := nameMap.find? name + let .some ui := nameMap.get? name | throwError "topoSortUnfolds :: Unexpected error" ret := ret.push ui return ret.reverse where - go (depMap : HashMap Name (HashSet Name)) (stack : HashSet Name) (n : Name) : StateRefT (HashSet Name × Array Name) MetaM Unit := do + go (depMap : Std.HashMap Name (Std.HashSet Name)) (stack : Std.HashSet Name) (n : Name) : StateRefT (Std.HashSet Name × Array Name) MetaM Unit := do if stack.contains n then throwError "topoSortUnfolds :: Cyclic dependency" let (done, ret) ← get if done.contains n then return - let deps := (depMap.find? n).getD {} + let deps := (depMap.get? n).getD {} set (done.insert n, ret) for dep in deps do go depMap (stack.insert n) dep @@ -139,7 +139,7 @@ def unfoldConsts (unfolds : Array Prep.ConstUnfoldInfo) (e : Expr) : Expr := Id. match e with | .const name lvls => if name == uiname then - val.instantiateLevelParams params.data lvls + val.instantiateLevelParams params.toList lvls else .none | _ => .none) diff --git a/Auto/Translation/Reduction.lean b/Auto/Translation/Reduction.lean index c7e6869..7bc758d 100644 --- a/Auto/Translation/Reduction.lean +++ b/Auto/Translation/Reduction.lean @@ -32,9 +32,9 @@ def unfoldProj (e : Expr) : MetaM Expr := | throwError s!"unfoldProj :: {name} is not a inductive type" let some structInfo := getStructureInfo? (← getEnv) name | throwError s!"unfoldProj :: {name} is not a structure" - let nameMap : HashMap Name StructureFieldInfo := HashMap.ofList - (structInfo.fieldInfo.map (fun sfi => (sfi.fieldName, sfi))).data - let sorted := structInfo.fieldNames.map (fun name => nameMap.find? name) + let nameMap : Std.HashMap Name StructureFieldInfo := Std.HashMap.ofList + (structInfo.fieldInfo.map (fun sfi => (sfi.fieldName, sfi))).toList + let sorted := structInfo.fieldNames.map (fun name => nameMap.get? name) let .some (.some sfi) := sorted[idx]? | throwError s!"unfoldProj :: Projection index out of bound" let nones : List (Option Expr) := (List.range indi.numParams).map (fun _ => .none) @@ -78,4 +78,4 @@ def prepReduceDefeq (type : Expr) : MetaM (Option Expr) := do let eq' := Lean.mkApp3 (.const ``Eq lvls) ty lhs rhs return .some (← mkForallFVars xs eq') -end Auto \ No newline at end of file +end Auto diff --git a/Auto/Translation/ReifM.lean b/Auto/Translation/ReifM.lean index 1ed5796..e5256ce 100644 --- a/Auto/Translation/ReifM.lean +++ b/Auto/Translation/ReifM.lean @@ -16,9 +16,9 @@ structure State where -- During monomorphization, constants will be turned -- into free variables. This map records the original expression -- corresponding to these free variables. - exprFVarVal : HashMap FVarId Expr := {} + exprFVarVal : Std.HashMap FVarId Expr := {} -- Canonicalization map for types - tyCanMap : HashMap Expr Expr := {} + tyCanMap : Std.HashMap Expr Expr := {} -- Inhabited canonicalized types -- The second `Expr` should be of the form `ty₁ → ty₂ → ⋯ → tyₙ`, -- where `ty₁, ty₂, ⋯, tyₙ` are canonicalized types within `tyCanMap` @@ -36,11 +36,11 @@ abbrev ReifM := StateT State MetaM -/ @[inline] def resolveVal (e : Expr) : ReifM Expr := match e with - | .fvar id => do return ((← getExprFVarVal).find? id).getD e + | .fvar id => do return ((← getExprFVarVal).get? id).getD e | _ => return e @[inline] def resolveTy (e : Expr) : ReifM Expr := do - match (← getTyCanMap).find? (Expr.eraseMData e) with + match (← getTyCanMap).get? (Expr.eraseMData e) with | .some id => return id | .none => throwError "resolveTy :: Unable to resolve {e}" diff --git a/README.md b/README.md index 89e851f..342598c 100644 --- a/README.md +++ b/README.md @@ -6,6 +6,9 @@ Lean-auto is an interface between Lean and automated theorem provers. Up to now, Lean-auto is based on a monomorphization procedure from dependent type theory to higher-order logic and a deep embedding of higher-order logic into dependent type theory. It is capable of handling dependently-typed and/or universe-polymorphic input terms. Currently, proof reconstruction can be handled by [Duper](https://github.com/leanprover-community/duper), a higher-order superposition prover written in Lean. To enable Duper, please import the duper repo in your project, and set the following options: ```lean +import Auto.Tactic +import Duper.Tactic + open Lean Auto in def Auto.duperRaw (lemmas : Array Lemma) (_ : Array Lemma) : MetaM Expr := do let lemmas : Array (Expr × Expr × Array Name × Bool) ← lemmas.mapM diff --git a/TODO.md b/TODO.md index 4a7f378..c1c7a5c 100644 --- a/TODO.md +++ b/TODO.md @@ -1,7 +1,6 @@ __Auto Issues:__ * HOL to FOL. Do this in the verified checker. Introduce new etoms to represent instancecs of higher-order functions. * When we matched against ``Eq`` in monomorphization, we found that some type arguments of other constants got unified with ``Prop``, which created a bunch of junk lemma. Find out whether something similar happened in Duper. -* Matcher Equational Theorems: Please use ``Lean.Meta.Match.getEquationsFor``. Maybe ``getMatcherInfo?`` is also useful. * Implement native interpretation * Improve portfolio mode script. Download E and zipperposition from the web. * Floating point numbers diff --git a/Test/SmtTranslation/Inductive.lean b/Test/SmtTranslation/Inductive.lean index 52d2a64..0375759 100644 --- a/Test/SmtTranslation/Inductive.lean +++ b/Test/SmtTranslation/Inductive.lean @@ -42,7 +42,6 @@ section Enum | .Z3 => 4000 | .Z4 => 3000 - -- **TODO**: Better support for translation of functions with `match` set_option trace.auto.printLemmas true in example (x : Zone) : x.MinArea1 <= x.MinArea2 := by cases x <;> auto d[Zone.MinArea1, Zone.MinArea2] @@ -116,7 +115,6 @@ section Mixed have h₂ : ∀ (x : α) (ys : _), List.head? (x :: ys) = .some x := fun _ _ => rfl auto - -- **TODO**: Did not get desired definitional equation example (x : α) : List.head? [x] = .some x := by auto d[List.head?] diff --git a/Test/SmtTranslation/String.lean b/Test/SmtTranslation/String.lean index 2686a3b..cc41256 100644 --- a/Test/SmtTranslation/String.lean +++ b/Test/SmtTranslation/String.lean @@ -21,4 +21,5 @@ example : String.length "abc" = 3 := by auto example : String.isPrefixOf "ab" "abcd" := by auto +set_option auto.smt.solver.name "cvc5" example : String.replace "aaaaa" "aa" "ba" = "babaa" := by auto diff --git a/Test/Test_Regression.lean b/Test/Test_Regression.lean index d5a1a39..b2c4839 100644 --- a/Test/Test_Regression.lean +++ b/Test/Test_Regression.lean @@ -543,7 +543,7 @@ section Adhoc example (a b : Int) - : a * b - a % (Int.mod b a) = a * b - a % (Int.mod b a) := by auto + : a * b - a % (Int.tmod b a) = a * b - a % (Int.tmod b a) := by auto example (a : Int) (h₁ : a ≥ 0) (h₂ : -a ≤ 0) (h₃ : 0 < 1) (h₄ : 2 > 0) diff --git a/lean-toolchain b/lean-toolchain index e98229d..e0676e0 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0 \ No newline at end of file +leanprover/lean4:v4.13.0 \ No newline at end of file From 5367bcd32133a50ee5c4c39fb6cfa345893387bf Mon Sep 17 00:00:00 2001 From: PratherConid Date: Sat, 23 Nov 2024 19:58:13 -0800 Subject: [PATCH 06/25] fix small issues --- Auto/Parser/NDFA.lean | 8 ++++---- Auto/Tactic.lean | 2 +- Auto/Translation/Assumptions.lean | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Auto/Parser/NDFA.lean b/Auto/Parser/NDFA.lean index 0ea5eb9..82b4ae5 100644 --- a/Auto/Parser/NDFA.lean +++ b/Auto/Parser/NDFA.lean @@ -336,8 +336,8 @@ section NFA hash c := hash c.val def test₁ : NFA String := ⟨#[ - HashMap.ofList [(.inr "a", #[5]), (.inr "b", #[1, 0])], - HashMap.ofList [(.inl .unit, #[1]), (.inr "c", #[2, 4]), (.inr "a", #[6,1,2])] + Std.HashMap.ofList [(.inr "a", #[5]), (.inr "b", #[1, 0])], + Std.HashMap.ofList [(.inl .unit, #[1]), (.inr "c", #[2, 4]), (.inr "a", #[6,1,2])] ], #[]⟩ def test₂ : NFA String := test₁.normalize @@ -474,8 +474,8 @@ section DFA /- def test₄ : DFA String := ⟨HashSet.empty.insert 3, #[ - HashMap.ofList [("a", 5), ("b", 0)], - HashMap.ofList [("q", 1), ("c", 4), ("a", 2)]], #[.empty, .empty]⟩ + Std.HashMap.ofList [("a", 5), ("b", 0)], + Std.HashMap.ofList [("q", 1), ("c", 4), ("a", 2)]], #[.empty, .empty]⟩ local instance : Hashable Char where hash c := hash c.val diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 2c4159d..3b096aa 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -160,7 +160,7 @@ def collectUserLemmas (terms : Array Term) : TacticM (Array Lemma) := return lemmas def collectHintDBLemmas (names : Array Name) : TacticM (Array Lemma) := do - let mut hs : HashSet Name := HashSet.empty + let mut hs : Std.HashSet Name := Std.HashSet.empty let mut ret : Array Lemma := #[] for name in names do let .some db ← findLemDB name diff --git a/Auto/Translation/Assumptions.lean b/Auto/Translation/Assumptions.lean index b338b07..17ff339 100644 --- a/Auto/Translation/Assumptions.lean +++ b/Auto/Translation/Assumptions.lean @@ -101,7 +101,7 @@ def Lemma.equivQuick (lem₁ lem₂ : Lemma) : MetaM Bool := do /- Reorder top-level `∀` so that (non-prop / dependent) ones precede other ones -/ /- def Lemma.reorderForallInstDep (lem : Lemma) : MetaM Lemma := do - let depargs := HashSet.empty.insertMany (Expr.depArgs lem.type) + let depargs := Std.HashSet.empty.insertMany (Expr.depArgs lem.type) Meta.forallTelescope lem.type fun xs body => do let mut prec := #[] let mut trail := #[] From a50480e17aef0d7110a51b41a42e304dd965c0be Mon Sep 17 00:00:00 2001 From: PratherConid Date: Sat, 23 Nov 2024 23:26:33 -0800 Subject: [PATCH 07/25] Fixes #38 --- Auto/Embedding/LamConv.lean | 87 ++++++++++++++++++++----------------- Test/Test_Regression.lean | 25 +++++++++-- 2 files changed, 70 insertions(+), 42 deletions(-) diff --git a/Auto/Embedding/LamConv.lean b/Auto/Embedding/LamConv.lean index 950a022..2c680ee 100644 --- a/Auto/Embedding/LamConv.lean +++ b/Auto/Embedding/LamConv.lean @@ -1351,11 +1351,13 @@ def LamTerm.betaBounded (n : Nat) (t : LamTerm) := | .bvar _ => t | .lam s t => .lam s (t.betaBounded n') | .app .. => - let tb := t.headBetaBounded n' - let fn := tb.getAppFn - let args := tb.getAppArgs - let argsb := args.map (fun ((s, arg) : LamSort × _) => (s, betaBounded n' arg)) - LamTerm.mkAppN fn argsb + match t.isHeadBetaTarget with + | true => LamTerm.betaBounded n' (t.headBetaBounded n') + | false => + let fn := t.getAppFn + let args := t.getAppArgs + let argsb := args.map (fun ((s, arg) : LamSort × _) => (s, betaBounded n' arg)) + LamTerm.mkAppN fn argsb theorem LamTerm.maxEVarSucc_betaBounded : (LamTerm.betaBounded n t).maxEVarSucc ≤ t.maxEVarSucc := by @@ -1366,16 +1368,20 @@ theorem LamTerm.maxEVarSucc_betaBounded : case lam s t => apply IH case app s fn arg => dsimp [betaBounded, maxEVarSucc] - apply LamTerm.maxEVarSucc_mkAppN - case hs => - apply HList.toMapTy; dsimp [Function.comp] - apply HList.map _ LamTerm.maxEVarSucc_getAppArgs - intro a; cases a; dsimp; intro h - apply Nat.le_trans _ (Nat.le_trans h _) - apply IH; apply maxEVarSucc_headBetaBounded - case ht => - apply Nat.le_trans maxEVarSucc_getAppFn - apply maxEVarSucc_headBetaBounded + cases (app s fn arg).isHeadBetaTarget + case true => + apply Nat.le_trans IH + apply Nat.le_trans LamTerm.maxEVarSucc_headBetaBounded (Nat.le_refl _) + case false => + apply LamTerm.maxEVarSucc_mkAppN + case hs => + apply HList.toMapTy; dsimp [Function.comp] + apply HList.map _ LamTerm.maxEVarSucc_getAppArgs + intro a; cases a; dsimp; intro h + apply Nat.le_trans _ (Nat.le_trans h _) + apply IH; exact Nat.le_refl _ + case ht => + apply Nat.le_trans maxEVarSucc_getAppFn (Nat.le_refl _) def LamTerm.betaReduced (t : LamTerm) := match t with @@ -1395,30 +1401,33 @@ theorem LamEquiv.ofBetaBounded match wf with | .ofLam _ wf => apply LamEquiv.ofLam; apply IH wf case app s fn arg => - dsimp; - have ⟨_, ⟨wfhbb, _⟩⟩ := LamEquiv.ofHeadBetaBounded (n:=n) wf - apply LamEquiv.trans (LamEquiv.ofHeadBetaBounded (n:=n) wf) - apply LamEquiv.trans (LamEquiv.eq wfhbb (LamTerm.appFn_appArg_eq _)) - let masterArr := (LamTerm.getAppArgs (LamTerm.headBetaBounded n (.app s fn arg))).map (fun (s, arg) => (s, arg, arg.betaBounded n)) - have eq₁ : (LamTerm.getAppArgs (LamTerm.headBetaBounded n (.app s fn arg))) = masterArr.map (fun (s, arg₁, _) => (s, arg₁)) := by - dsimp; rw [List.map_map]; rw [List.map_equiv _ id, List.map_id] - intro x; cases x; rfl - have eq₂ : List.map - (fun x => (x.fst, LamTerm.betaBounded n x.snd)) - (LamTerm.getAppArgs (LamTerm.headBetaBounded n (.app s fn arg))) = masterArr.map (fun (s, _, arg₂) => (s, arg₂)) := by - dsimp; rw [List.map_map]; apply List.map_equiv; - intro x; cases x; rfl - rw [eq₂, eq₁]; have ⟨fnTy, wfFn⟩ := wfhbb.getAppFn - apply LamEquiv.congrs (fnTy:=fnTy) - case wfApp => rw [← eq₁, ← LamTerm.appFn_appArg_eq]; exact wfhbb - case hFn => apply LamEquiv.refl wfFn - case hArgs => - dsimp; - apply HList.toMapTy; dsimp [Function.comp] - apply HList.map - (β:=fun (s, t) => LamWF lval.toLamTyVal ⟨lctx, t, s⟩) - (fun (s, t) => @IH lctx t s) - apply LamWF.getAppArgs wfhbb + cases (LamTerm.app s fn arg).isHeadBetaTarget <;> dsimp + case true => + apply LamEquiv.trans (LamEquiv.ofHeadBetaBounded (n:=n) wf) + have ⟨_, ⟨wfhbb, _⟩⟩ := LamEquiv.ofHeadBetaBounded (n:=n) wf + apply IH wfhbb + case false => + apply LamEquiv.trans (LamEquiv.eq wf (LamTerm.appFn_appArg_eq _)) + let masterArr := (LamTerm.getAppArgs (.app s fn arg)).map (fun (s, arg) => (s, arg, arg.betaBounded n)) + have eq₁ : (LamTerm.getAppArgs (.app s fn arg)) = masterArr.map (fun (s, arg₁, _) => (s, arg₁)) := by + dsimp; rw [List.map_map]; rw [List.map_equiv _ id, List.map_id] + intro x; cases x; rfl + have eq₂ : List.map + (fun x => (x.fst, LamTerm.betaBounded n x.snd)) + (LamTerm.getAppArgs (.app s fn arg)) = masterArr.map (fun (s, _, arg₂) => (s, arg₂)) := by + dsimp; rw [List.map_map]; apply List.map_equiv; + intro x; cases x; rfl + rw [eq₂, eq₁]; have ⟨fnTy, wfFn⟩ := wf.getAppFn + apply LamEquiv.congrs (fnTy:=fnTy) + case wfApp => rw [← eq₁, ← LamTerm.appFn_appArg_eq]; exact wf + case hFn => apply LamEquiv.refl wfFn + case hArgs => + dsimp; + apply HList.toMapTy; dsimp [Function.comp] + apply HList.map + (β:=fun (s, t) => LamWF lval.toLamTyVal ⟨lctx, t, s⟩) + (fun (s, t) => @IH lctx t s) + apply LamWF.getAppArgs wf theorem LamThmEquiv.ofBetaBounded (wf : LamThmWF lval lctx rty t) : LamThmEquiv lval lctx rty t (t.betaBounded n) := fun lctx => LamEquiv.ofBetaBounded (wf lctx) diff --git a/Test/Test_Regression.lean b/Test/Test_Regression.lean index b2c4839..586b2a4 100644 --- a/Test/Test_Regression.lean +++ b/Test/Test_Regression.lean @@ -334,12 +334,31 @@ example (H : (fun (x y z t : Nat) => x) = (fun x y z t => x)) : True := by example {α : Sort u} (add : ((α → α) → (α → α)) → ((α → α) → (α → α)) → ((α → α) → (α → α))) - (Hadd : ∀ x y f n, add x y f n = (x f) ((y f) n)) + (hadd : ∀ x y f n, add x y f n = (x f) ((y f) n)) (mul : ((α → α) → (α → α)) → ((α → α) → (α → α)) → ((α → α) → (α → α))) - (Hmul : ∀ x y f, mul x y f = x (y f)) + (hmul : ∀ x y f, mul x y f = x (y f)) (w₁ w₂ : ((α → α) → (α → α)) → ((α → α) → (α → α)) → ((α → α) → (α → α))) (Hw₁₂ : (w₁ = w₂) = (w₂ = w₁)) : True := by - auto [Hadd, Hmul, Hw₁₂] + auto [hadd, hmul, Hw₁₂] + +example + (P : (α → γ) → Prop) (f : α → β) (g : β → γ) (h : β → β) + : P ((g ∘ h) ∘ f) = P (fun x => g (h (f x))) := by + auto [Function.comp_def] + +example + (A : Sort u) + (add : ∀ {α}, ((α → α) → (α → α)) → ((α → α) → (α → α)) → ((α → α) → (α → α))) + (hadd : ∀ {α} x y f n, @add α x y f n = (x f) ((y f) n)) + (mul : ∀ {α}, ((α → α) → (α → α)) → ((α → α) → (α → α)) → ((α → α) → (α → α))) + (hmul : ∀ {α} x y f, @mul α x y f = x (y f)) + (two : (A → A) → (A → A)) + (htwo : ∀ f x, two f x = f (f x)) + (three : (A → A) → (A → A)) + (hthree : ∀ f x, three f x = f (f (f x))) : + mul three (add two (add three three)) = + mul three (mul two (add two two)) := by + auto [hadd, hmul, htwo, hthree] -- Polymorphic Constant From 2511ca29f39488a89eae6d7cdc0b995dd170304d Mon Sep 17 00:00:00 2001 From: PratherConid Date: Sun, 24 Nov 2024 00:36:29 -0800 Subject: [PATCH 08/25] fix duper prelude in readme --- README.md | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/README.md b/README.md index 342598c..210f5d8 100644 --- a/README.md +++ b/README.md @@ -10,10 +10,10 @@ import Auto.Tactic import Duper.Tactic open Lean Auto in -def Auto.duperRaw (lemmas : Array Lemma) (_ : Array Lemma) : MetaM Expr := do +def Auto.duperRaw (lemmas : Array Lemma) (inhs : Array Lemma) : MetaM Expr := do let lemmas : Array (Expr × Expr × Array Name × Bool) ← lemmas.mapM (fun ⟨⟨proof, ty, _⟩, _⟩ => do return (ty, ← Meta.mkAppM ``eq_true #[proof], #[], true)) - runDuper lemmas.data 0 + Duper.runDuper lemmas.toList 0 attribute [rebind Auto.Native.solverFunc] Auto.duperRaw set_option auto.native true From fb61533bab6f494f8732022b2237319a79e23804 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Mon, 25 Nov 2024 20:37:29 -0800 Subject: [PATCH 09/25] Support for cond --- Auto/Lib/BoolExtra.lean | 8 ++++++++ Auto/Tactic.lean | 3 +++ Test/Test_Regression.lean | 10 ++++++++++ 3 files changed, 21 insertions(+) diff --git a/Auto/Lib/BoolExtra.lean b/Auto/Lib/BoolExtra.lean index e85c5b1..e299d1c 100644 --- a/Auto/Lib/BoolExtra.lean +++ b/Auto/Lib/BoolExtra.lean @@ -86,6 +86,14 @@ theorem Bool.ite'_eq_true (p : Prop) (a b : α) : p → ite' p a b = a := by theorem Bool.ite'_eq_false (p : Prop) (a b : α) : ¬ p → ite' p a b = b := by intro hp; apply Bool.ite_simp ▸ @Bool.ite_eq_false _ p (Classical.propDecidable _) a b hp +/-- + Invoke `rw` at the beginning of `auto` +-/ +theorem Bool.cond_simp.{u} : @cond = fun (α : Type u) b => @ite' α (b = true) := by + funext α b x y; cases b + case false => rw [cond_false, Bool.ite'_eq_false]; exact Bool.false_ne_true + case true => rw [cond_true, Bool.ite'_eq_true]; rfl + /-- Invoke `rw` at the beginning of `auto` -/ diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 3b096aa..d50352f 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -434,6 +434,9 @@ def runAuto -- Simplify `ite` let ite_simp_lem ← Lemma.ofConst ``Auto.Bool.ite_simp (.leaf "hw Auto.Bool.ite_simp") let lemmas ← lemmas.mapM (fun lem => Lemma.rewriteUPolyRigid lem ite_simp_lem) + -- Simplify `cond` + let cond_simp_lem ← Lemma.ofConst ``Auto.Bool.cond_simp (.leaf "hw Auto.Bool.cond_simp") + let lemmas ← lemmas.mapM (fun lem => Lemma.rewriteUPolyRigid lem cond_simp_lem) -- Simplify `decide` let decide_simp_lem ← Lemma.ofConst ``Auto.Bool.decide_simp (.leaf "hw Auto.Bool.decide_simp") let lemmas ← lemmas.mapM (fun lem => Lemma.rewriteUPolyRigid lem decide_simp_lem) diff --git a/Test/Test_Regression.lean b/Test/Test_Regression.lean index 586b2a4..34b6e69 100644 --- a/Test/Test_Regression.lean +++ b/Test/Test_Regression.lean @@ -532,6 +532,16 @@ section Adhoc (h : if (a = b) then True else a = b) : a = b := by auto + -- Cond + example : cond true a b = a ∧ cond false a b = b := by + auto + + example (h : p = true) : cond p a b = a := by + auto + + example (h : p = false) : cond p a b = b := by + auto + -- Decide example : ∀ b, !(b = true) ↔ b = false := by auto From 2078181d6ff8adc160455b0571dfb0ca5a42ec9a Mon Sep 17 00:00:00 2001 From: PratherConid Date: Tue, 26 Nov 2024 01:41:38 -0800 Subject: [PATCH 10/25] clean up corecheck code --- Auto/Lib/MetaExtra.lean | 52 ++++++++++++----------------------- Auto/Tactic.lean | 2 +- Auto/Translation/LamReif.lean | 4 --- 3 files changed, 19 insertions(+), 39 deletions(-) diff --git a/Auto/Lib/MetaExtra.lean b/Auto/Lib/MetaExtra.lean index 06f4c9b..25e0f56 100644 --- a/Auto/Lib/MetaExtra.lean +++ b/Auto/Lib/MetaExtra.lean @@ -65,22 +65,9 @@ unsafe def evalFromMetaTactic : Tactic.Tactic Tactic.liftMetaTactic mt | _ => Elab.throwUnsupportedSyntax -/-- We assume that `value` contains no free variables or metavariables -/ -def Meta.exprAddAndCompile (value : Expr) (declName : Name) : MetaM Unit := do - let type ← inferType value - let us := collectLevelParams {} value |>.params - let decl := Declaration.defnDecl { - name := declName - levelParams := us.toList - type := type - value := value - hints := ReducibilityHints.opaque - safety := DefinitionSafety.unsafe - } - addDecl decl - -def Meta.coreCheck (e : Expr) : MetaM Unit := do - let startTime ← IO.monoMsNow +-- **-TODO: Deal with mutual dependencies between mvar and fvar** +-- **- Maybe inspect the LocalContext in the metavariable declaration** +def Meta.abstractAllMVarFVar (e : Expr) : MetaM Expr := do let mut curr := e -- **TODO: Fix** while true do @@ -90,8 +77,8 @@ def Meta.coreCheck (e : Expr) : MetaM Unit := do curr := next let e := curr -- Now `(e == (← instantiateMVars) e) && (e == (← zetaReduce e))` - let res ← Meta.abstractMVars e - -- Now metavariables in `e` are abstracted + let res ← Meta.abstractMVars (levels := false) e + -- Now expr mvars in `e` are abstracted let mut e := res.expr -- **TODO: Fix** while true do @@ -100,23 +87,20 @@ def Meta.coreCheck (e : Expr) : MetaM Unit := do e ← mkLambdaFVars (collectFVarsState.fvarIds.map Expr.fvar) e if !e.hasFVar then break - -- Use `Core.addAndCompile` to typecheck `e` - let coreChkStart ← IO.monoMsNow - trace[auto.metaExtra] "Meta.coreCheck :: Preprocessing done in {coreChkStart - startTime}" - let env ← getEnv - try - Meta.exprAddAndCompile e `_coreCheck - finally - trace[auto.metaExtra] "Meta.coreCheck :: Core check done in {(← IO.monoMsNow) - coreChkStart}" - setEnv env + let res ← Meta.abstractMVars (levels := true) e + -- Now level mvars in `e` are abstracted + -- Level mvars must be abstracted after fvars are abstracted, + -- because they may occur in fvars + return res.expr + +def Meta.coreCheck (e : Expr) : MetaM (Option Expr) := do + let e ← Meta.abstractAllMVarFVar e + match Kernel.check (← getEnv) {} e with + | Except.ok type => return .some type + | Except.error _ => return .none def Meta.isTypeCorrectCore (e : Expr) : MetaM Bool := do - try - Meta.coreCheck e - pure true - catch e => - let msg := MessageData.compose "Meta.isTypeCorrectCore failed with message : \n" e.toMessageData - trace[auto.metaExtra] msg - pure false + let type? ← Meta.coreCheck e + return type?.isSome end Auto diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index d50352f..20060ff 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -131,7 +131,7 @@ def parseUOrDs (stxs : Array (TSyntax ``uord)) : TacticM (Array Prep.ConstUnfold defeqs := defeqs.append d return (unfolds, defeqs) -def collectLctxLemmas (lctxhyps : Bool) (ngoalAndBinders : Array FVarId) : TacticM (Array Lemma) := +def collectLctxLemmas (lctxhyps : Bool) (ngoalAndBinders : Array FVarId) : MetaM (Array Lemma) := Meta.withNewMCtxDepth do let fVarIds ← if lctxhyps then pure (← getLCtx).getFVarIds else pure ngoalAndBinders let mut lemmas := #[] diff --git a/Auto/Translation/LamReif.lean b/Auto/Translation/LamReif.lean index aabe7a5..49ddc87 100644 --- a/Auto/Translation/LamReif.lean +++ b/Auto/Translation/LamReif.lean @@ -1628,7 +1628,6 @@ section BuildChecker (.app (.const ``ilβ [u]) tyValExpr)) def buildCPValExpr : ReifM Expr := do - -- let startTime ← IO.monoMsNow let u ← getU let tyValExpr ← buildTyVal let tyValTy := Expr.forallE `_ (.const ``Nat []) (.sort (.succ u)) .default @@ -1637,9 +1636,6 @@ section BuildChecker let ilExpr ← buildILExpr tyValFVarExpr let checkerValuationExpr := Lean.mkApp3 (.const ``CPVal.mk [u]) tyValExpr varExpr ilExpr Meta.mkLetFVars #[tyValFVarExpr] checkerValuationExpr - -- if !(← Meta.isTypeCorrectCore lamValuationExpr) then - -- throwError "buildLamValuation :: Malformed LamValuation" - -- trace[auto.buildChecker] "LamValuation typechecked in time {(← IO.monoMsNow) - startTime}" return lamValuationExpr /-- `lvalExpr` is the `LamValuation` -/ From 3e894ca719771ec180ed8c75f0cf0345794a5623 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Tue, 26 Nov 2024 19:30:35 -0800 Subject: [PATCH 11/25] use decl_name% --- Auto/Embedding/LamTermInterp.lean | 10 +- Auto/IR/SMT.lean | 2 +- Auto/LemDB.lean | 2 +- Auto/Lib/DeCompile.lean | 8 +- Auto/Lib/ExprExtra.lean | 14 +- Auto/Lib/MetaExtra.lean | 6 +- Auto/Lib/MetaState.lean | 2 +- Auto/Lib/MonadUtils.lean | 18 +- Auto/Parser/TPTP.lean | 10 +- Auto/Solver/Native.lean | 2 +- Auto/Solver/SMT.lean | 16 +- Auto/Solver/TPTP.lean | 4 +- Auto/Tactic.lean | 42 ++--- Auto/Translation/Assumptions.lean | 26 +-- Auto/Translation/Inductive.lean | 20 +-- Auto/Translation/Inhabitation.lean | 4 +- Auto/Translation/Lam2DAtomAsFVar.lean | 16 +- Auto/Translation/Lam2TH0.lean | 4 +- Auto/Translation/LamFOL2SMT.lean | 54 +++--- Auto/Translation/LamReif.lean | 230 ++++++++++++------------- Auto/Translation/LamUtils.lean | 38 ++-- Auto/Translation/Monomorphization.lean | 46 ++--- Auto/Translation/Preprocessing.lean | 22 +-- Auto/Translation/Reduction.lean | 6 +- Auto/Translation/ReifM.lean | 4 +- Doc/LamPULift.lean | 33 ++-- 26 files changed, 319 insertions(+), 320 deletions(-) diff --git a/Auto/Embedding/LamTermInterp.lean b/Auto/Embedding/LamTermInterp.lean index 69de7de..ccc69dd 100644 --- a/Auto/Embedding/LamTermInterp.lean +++ b/Auto/Embedding/LamTermInterp.lean @@ -268,10 +268,10 @@ abbrev InterpM := StateRefT State MetaState.MetaStateM def getLCtxTy! (idx : Nat) : InterpM LamSort := do let lctxTyRev ← getLctxTyRev if idx ≥ lctxTyRev.size then - throwError "getLCtxTy! :: Index out of bound" + throwError "{decl_name%} :: Index out of bound" match lctxTyRev[idx]? with | .some s => return s - | .none => throwError "getLCtxTy! :: Unexpected error" + | .none => throwError "{decl_name%} :: Unexpected error" /-- Turning a sort into `fvar` in a hash-consing manner @@ -303,7 +303,7 @@ def collectSortFor (ltv : LamTyVal) : LamTerm → InterpM LamSort | .atom n => do let _ ← sort2FVarId (ltv.lamVarTy n) return ltv.lamVarTy n -| .etom _ => throwError "collectSortFor :: etoms should not occur here" +| .etom _ => throwError "{decl_name%} :: etoms should not occur here" | .base b => do let s := b.lamCheck ltv let _ ← sort2FVarId s @@ -325,8 +325,8 @@ def collectSortFor (ltv : LamTyVal) : LamTerm → InterpM LamSort if argTy' == argTy && argTy' == s then return resTy else - throwError "collectSortFor :: Application type mismatch" - | _ => throwError "collectSortFor :: Malformed application" + throwError "{decl_name%} :: Application type mismatch" + | _ => throwError "{decl_name%} :: Malformed application" where withLCtxTy {α : Type} (s : LamSort) (k : InterpM α) : InterpM α := do let lctxTyRev ← getLctxTyRev setLctxTyRev (lctxTyRev.push s) diff --git a/Auto/IR/SMT.lean b/Auto/IR/SMT.lean index 1384e37..32c70eb 100644 --- a/Auto/IR/SMT.lean +++ b/Auto/IR/SMT.lean @@ -472,7 +472,7 @@ section if let .some name := h2lMap.get? cstr then return name let .some nameSuggestion := nameSuggestion - | throwError "IR.SMT.h2Symb :: Fresh high-level constraint {cstr} without name suggestion" + | throwError "{decl_name%} :: Fresh high-level constraint {cstr} without name suggestion" let name ← processSuggestedName nameSuggestion setL2hMap (l2hMap.insert name cstr) setH2lMap (h2lMap.insert cstr name) diff --git a/Auto/LemDB.lean b/Auto/LemDB.lean index 9ab8ad6..f7f2a36 100644 --- a/Auto/LemDB.lean +++ b/Auto/LemDB.lean @@ -42,7 +42,7 @@ partial def LemDB.toHashSet : LemDB → AttrM (Std.HashSet Name) let mut ret := Std.HashSet.empty for hdb in hdbs do let some hdb := state.get? hdb - | throwError "LemDB.toHashSet :: Unknown lemma database {hdb}" + | throwError "{decl_name%} :: Unknown lemma database {hdb}" let hset ← hdb.toHashSet ret := ret.insertMany hset return ret diff --git a/Auto/Lib/DeCompile.lean b/Auto/Lib/DeCompile.lean index 4d79c70..f9f2f3a 100644 --- a/Auto/Lib/DeCompile.lean +++ b/Auto/Lib/DeCompile.lean @@ -86,7 +86,7 @@ private partial def exprDeCompileAux (final : Bool) (e : Expr) : ExprDeCompM Str | .natVal l => return toString l | .strVal s => return toString s | Expr.sort l => return s!"Sort ({l})" - | Expr.mvar .. => throwError "exprDeCompile :: Metavariable is not supported" + | Expr.mvar .. => throwError "{decl_name%} :: Metavariable is not supported" | Expr.fvar f => let some decl := (← getLCtx).find? f | throwError "Unknown free variable {e}" @@ -103,9 +103,9 @@ where String.intercalate ", " (us.map toString) ++ "}" etaExpandConst (cst : Expr) (appliedArgs : Nat) := do if !cst.isConst then - throwError s!"exprDeCompile :: etaExpandConst received non-const {cst}" + throwError s!"{decl_name%} :: etaExpandConst received non-const {cst}" let some val := (← getEnv).find? cst.constName! - | throwError s!"exprDeCompile :: Unknown identifier {cst.constName!}" + | throwError s!"{decl_name%} :: Unknown identifier {cst.constName!}" let ty := val.type let lparams := val.levelParams let bs := Expr.forallBinders ty @@ -145,7 +145,7 @@ def exprCollectNames (e : Expr) : StateT (Std.HashSet String) MetaM Unit := do | Expr.mvar .. => return | Expr.fvar f => let some decl := (← getLCtx).find? f - | throwError "Unknown free variable {e}" + | throwError "{decl_name%} :: Unknown free variable {e}" let name := decl.userName.toString modify (fun s => s.insert name) | Expr.bvar .. => return diff --git a/Auto/Lib/ExprExtra.lean b/Auto/Lib/ExprExtra.lean index a83f10b..336d554 100644 --- a/Auto/Lib/ExprExtra.lean +++ b/Auto/Lib/ExprExtra.lean @@ -203,7 +203,7 @@ def Expr.depArgs (e : Expr) : Array Nat := ⟨Expr.depArgsIdx e 0⟩ -/ def Expr.constDepArgs (c : Name) : CoreM (Array Nat) := do let .some decl := (← getEnv).find? c - | throwError "Expr.constDepArgs :: Unknown constant {c}" + | throwError "{decl_name%} :: Unknown constant {c}" return Expr.depArgs decl.type def Expr.numLeadingDepArgs : Expr → Nat @@ -234,14 +234,14 @@ def Expr.formatWithUsername (e : Expr) : MetaM Format := do let names ← fvarIds.mapM (fun fid => do match ← fid.findDecl? with | .some decl => return decl.userName - | .none => throwError "Expr.formatWithUsername :: Unknown free variable {(Expr.fvar fid).dbgToString}") + | .none => throwError "{decl_name%}e :: Unknown free variable {(Expr.fvar fid).dbgToString}") let e := e.replaceFVars (fvarIds.map Expr.fvar) (names.map (Expr.const · [])) let e ← instantiateMVars e let mvarIds := (Expr.collectMVars {} e).result let names ← mvarIds.mapM (fun mid => do match ← mid.findDecl? with | .some decl => return decl.userName - | .none => throwError "xpr.formatWithUsername :: Unknown metavariable {(Expr.mvar mid).dbgToString}") + | .none => throwError "{decl_name%} :: Unknown metavariable {(Expr.mvar mid).dbgToString}") let e := (e.abstract (mvarIds.map Expr.mvar)).instantiateRev (names.map (Expr.const · [])) return f!"{e}" @@ -257,14 +257,14 @@ unsafe def elabGetExprAndApply : CommandElab := fun stx => match stx with | `(command | #getExprAndApply[ $t:term | $i:ident ]) => withRef stx <| do let some iexpr ← Term.resolveId? i - | throwError "elabGetExprAndApply :: Unknown identifier {i}" + | throwError "{decl_name%} :: Unknown identifier {i}" let e ← Term.elabTerm t none Term.synthesizeSyntheticMVarsNoPostponing (ignoreStuckTC := true) let e ← Term.levelMVarToParam (← instantiateMVars e) let fname := iexpr.constName! match (← getEnv).evalConst (Expr → TermElabM Unit) (← getOptions) fname with | Except.ok f => f e - | Except.error err => throwError "elabGetExprAndApply :: Failed to evaluate {fname} to a term of type (Expr → TermElabM Unit), error : {err}" + | Except.error err => throwError "{decl_name%} :: Failed to evaluate {fname} to a term of type (Expr → TermElabM Unit), error : {err}" | _ => throwUnsupportedSyntax /-- @@ -274,7 +274,7 @@ unsafe def elabGetExprAndApply : CommandElab := fun stx => unsafe def exprFromExpr (eToExpr : Expr) : TermElabM Expr := do let ty ← Meta.inferType eToExpr if ! (← Meta.isDefEq ty (.const ``Expr [])) then - throwError "exprFromExpr :: Type `{ty}` of input is not definitionally equal to `Expr`" + throwError "{decl_name%} :: Type `{ty}` of input is not definitionally equal to `Expr`" let declName := `_exprFromExpr let addAndCompile (value : Expr) : TermElabM Unit := do let value ← Term.levelMVarToParam (← instantiateMVars value) @@ -359,7 +359,7 @@ section EvalAtTermElabM let α ← reduce (skipTypes := false) α synthInstance (mkApp (Lean.mkConst evalClassName [u]) α) catch _ => - throwError "expression{indentExpr e}\nhas type{indentExpr α}\nbut instance{indentExpr inst}\nfailed to be synthesized, this instance instructs Lean on how to display the resulting value, recall that any type implementing the `Repr` class also implements the `{evalClassName}` class" + throwError "{decl_name%} :: expression{indentExpr e}\nhas type{indentExpr α}\nbut instance{indentExpr inst}\nfailed to be synthesized, this instance instructs Lean on how to display the resulting value, recall that any type implementing the `Repr` class also implements the `{evalClassName}` class" private def mkRunMetaEval (e : Expr) : MetaM Expr := withLocalDeclD `env (mkConst ``Lean.Environment) fun env => diff --git a/Auto/Lib/MetaExtra.lean b/Auto/Lib/MetaExtra.lean index 25e0f56..b16d25d 100644 --- a/Auto/Lib/MetaExtra.lean +++ b/Auto/Lib/MetaExtra.lean @@ -40,7 +40,7 @@ def Meta.inspectMVarAssignments : MetaM Unit := do def Meta.trySynthInhabited (e : Expr) : MetaM (Option Expr) := do let eSort ← Expr.normalizeType (← instantiateMVars (← Meta.inferType e)) let .sort lvl := eSort - | throwError "trySynthInhabited :: {e} is not a type" + | throwError "{decl_name%} :: {e} is not a type" try if let .some inh ← Meta.synthInstance? e then return .some inh @@ -58,10 +58,10 @@ syntax (name := fromMetaTactic) "fromMetaTactic" "[" ident "]" : tactic unsafe def evalFromMetaTactic : Tactic.Tactic | `(fromMetaTactic | fromMetaTactic [ $i ]) => do let some iexpr ← Term.resolveId? i - | throwError "evalFromMetaTactic :: Unknown identifier {i}" + | throwError "{decl_name%} :: Unknown identifier {i}" let mtname := iexpr.constName! let Except.ok mt := (← getEnv).evalConst (MVarId → MetaM (List MVarId)) (← getOptions) mtname - | throwError "evalFromMetaTactic :: Failed to evaluate {mtname} to a term of type (Expr → TermElabM Unit)" + | throwError "{decl_name%} :: Failed to evaluate {mtname} to a term of type (Expr → TermElabM Unit)" Tactic.liftMetaTactic mt | _ => Elab.throwUnsupportedSyntax diff --git a/Auto/Lib/MetaState.lean b/Auto/Lib/MetaState.lean index 458d95d..abfe707 100644 --- a/Auto/Lib/MetaState.lean +++ b/Auto/Lib/MetaState.lean @@ -75,7 +75,7 @@ private def runWithFVars (lctx : LocalContext) (fvarids : Array FVarId) (k : Met newlctx := newlctx.mkLocalDecl fvarId userName type bi kind | .ldecl _ fvarId userName type value nonDep kind => newlctx := newlctx.mkLetDecl fvarId userName type value nonDep kind - | .none => throwError "runWithFVars :: Unknown free variable {Expr.fvar fid}" + | .none => throwError "{decl_name%} :: Unknown free variable {Expr.fvar fid}" withReader (fun ctx => {ctx with lctx := newlctx}) k private def runWithIntroducedFVarsImp (m : MetaStateM (Array FVarId × α)) (k : α → MetaM β) : MetaM β := do diff --git a/Auto/Lib/MonadUtils.lean b/Auto/Lib/MonadUtils.lean index ba33e99..57bbda7 100644 --- a/Auto/Lib/MonadUtils.lean +++ b/Auto/Lib/MonadUtils.lean @@ -53,13 +53,13 @@ syntax (name := genMonadState) "#genMonadState" term : command -/ def structureProjs (structTy : Expr) : CoreM (Name × Expr × Array (Name × Expr)) := do let .const structName lvls := structTy.getAppFn - | throwError s!"structureProjs :: Head symbol of {structTy} is not a constant" + | throwError s!"{decl_name%} :: Head symbol of {structTy} is not a constant" let some structInfo := getStructureInfo? (← getEnv) structName - | throwError s!"structureProjs :: {structName} is not a structure" + | throwError s!"{decl_name%} :: {structName} is not a structure" let .some (.inductInfo structi) := (← getEnv).find? structName - | throwError s!"structureProjs :: Unknown identifier {structName}" + | throwError s!"{decl_name%} :: Unknown identifier {structName}" let [structDotMk] := structi.ctors - | throwError s!"structureProjs :: {structName} is not a structure" + | throwError s!"{decl_name%} :: {structName} is not a structure" let structMkExpr := mkAppN (Expr.const structDotMk lvls) structTy.getAppArgs let tyArgs := structTy.getAppArgs let nameMap : Std.HashMap Name StructureFieldInfo := Std.HashMap.ofList @@ -75,7 +75,7 @@ private def elabStx (stx : Term) : TermElabM Expr := do Term.synthesizeSyntheticMVarsNoPostponing (ignoreStuckTC := true) let e ← instantiateMVars expr if e.hasMVar then - throwError "elabStx :: {e} contains metavariables" + throwError "{decl_name%} :: {e} contains metavariables" return e private def addDefnitionValFromExpr (e : Expr) (name : Name) : MetaM Unit := do @@ -91,9 +91,9 @@ private def addDefnitionValFromExpr (e : Expr) (name : Name) : MetaM Unit := do private def elabGenMonadContextAux (m : Term) (stx : Syntax) : CommandElabM Unit := runTermElabM <| fun xs => do let mexpr ← elabStx m let .const mname _ := mexpr.getAppFn - | throwError "elabGenMonadContext :: Unknown monad" + | throwError "{decl_name%} :: Unknown monad" let .str pre _ := mname - | throwError "elabGenMonadContext :: {mname} is not a valid constant name" + | throwError "{decl_name%} :: {mname} is not a valid constant name" let pureInst ← elabStx (← `(inferInstanceAs (Pure $m))) let mctxInst ← elabStx (← `(inferInstanceAs (MonadReader _ $m))) let mctxInstTy ← Meta.inferType mctxInst @@ -122,9 +122,9 @@ where private def elabGenMonadStateAux (m : Term) (stx : Syntax) : CommandElabM Unit := runTermElabM <| fun xs => do let mexpr ← elabStx m let .const mname _ := mexpr.getAppFn - | throwError "elabGenMonadState :: Unknown monad" + | throwError "{decl_name%} :: Unknown monad" let .str pre _ := mname - | throwError "elabGenMonadState :: {mname} is not a valid constant name" + | throwError "{decl_name%} :: {mname} is not a valid constant name" let pureInst ← elabStx (← `(inferInstanceAs (Pure $m))) let mstateInst ← elabStx (← `(inferInstanceAs (MonadState _ $m))) let mstateInstTy ← Meta.inferType mstateInst diff --git a/Auto/Parser/TPTP.lean b/Auto/Parser/TPTP.lean index 15351c7..fd66be9 100644 --- a/Auto/Parser/TPTP.lean +++ b/Auto/Parser/TPTP.lean @@ -854,8 +854,8 @@ def getProof (lamVarTy lamEVarTy : Array LamSort) (cmds : Array Command) : MetaM match term2LamTerm pi val with | .term (.base .prop) t => ret := ret.push t - | .error e => throwError ("getProof :: " ++ e) - | lc => throwError "getProof :: Unexpected LamConstr {lc}, expected term" + | .error e => throwError (decl_name% ++ " :: " ++ e) + | lc => throwError "{decl_name%} :: Unexpected LamConstr {lc}, expected term" else match val with | ⟨.ident name, [ty]⟩ => @@ -863,13 +863,13 @@ def getProof (lamVarTy lamEVarTy : Array LamSort) (cmds : Array Command) : MetaM if name.take 3 == "sk_" then match term2LamTerm pi ty with | .sort s => pi := pi.addSkolem name s - | lc => throwError "getProof :: Unexpected LamConstr {lc}, expected sort" + | lc => throwError "{decl_name%} :: Unexpected LamConstr {lc}, expected sort" else continue | _ => continue | _ => continue - | "include" => throwError "getProof :: `include` should not occur in proof output of TPTP solvers" - | _ => throwError "getProof :: Unknown command {cmd}" + | "include" => throwError "{decl_name%} :: `include` should not occur in proof output of TPTP solvers" + | _ => throwError "{decl_name%} :: Unknown command {cmd}" return ret /-- Returns the unsat core of an array of TSTP steps -/ diff --git a/Auto/Solver/Native.lean b/Auto/Solver/Native.lean index ebfa75d..61181ae 100644 --- a/Auto/Solver/Native.lean +++ b/Auto/Solver/Native.lean @@ -42,7 +42,7 @@ opaque queryNative : Array Lemma → Array Lemma → MetaM Expr def emulateNative (lemmas : Array Lemma) (_ : Array Lemma) : MetaM Expr := do let _ ← lemmas.mapM (fun lem => do if lem.params.size != 0 then - throwError "Solver.emulateNative :: Universe levels parameters are not supported") + throwError "{decl_name%} :: Universe levels parameters are not supported") let descrs := lemmas.zipWithIndex.map (fun (lem, i) => (Name.mkSimple s!"lem{i}", lem.type, .default)) let sty := Expr.mkForallFromBinderDescrs descrs (.const ``False []) let sorryExpr := Lean.mkApp2 (.const ``sorryAx [.zero]) sty (.const ``Bool.false []) diff --git a/Auto/Solver/SMT.lean b/Auto/Solver/SMT.lean index 06687c5..67fca0d 100644 --- a/Auto/Solver/SMT.lean +++ b/Auto/Solver/SMT.lean @@ -97,7 +97,7 @@ private def emitCommands (p : SolverProc) (c : Array IR.SMT.Command) : IO Unit : def createSolver (name : SolverName) : MetaM SolverProc := do let tlim := auto.smt.timeout.get (← getOptions) match name with - | .none => throwError "createSolver :: Unexpected error" + | .none => throwError "{decl_name%} :: Unexpected error" | .z3 => createAux "z3" #["-in", "-smt2", s!"-T:{tlim}"] | .cvc4 => throwError "cvc4 is not supported" | .cvc5 => createAux "cvc5" #[s!"--tlimit={tlim * 1000}", "--produce-models"] @@ -111,31 +111,31 @@ axiom autoSMTSorry.{u} (α : Sort u) : α def getSexp (s : String) : MetaM (Sexp × String) := match parseSexp s ⟨0⟩ {} with | .complete se p => return (se, Substring.toString ⟨s, p, s.endPos⟩) - | .incomplete _ _ => throwError s!"getSexp :: Incomplete input {s}" - | .malformed => throwError s!"getSexp :: Malformed (prefix of) input {s}" + | .incomplete _ _ => throwError s!"{decl_name%} :: Incomplete input {s}" + | .malformed => throwError s!"{decl_name%} :: Malformed (prefix of) input {s}" /-- Recover id of valid facts from unsat core. Refer to `lamFOL2SMT` -/ def validFactOfUnsatCore (unsatCore : Sexp) : MetaM (Array Nat) := do let .app unsatCore := unsatCore - | throwError "validFactOfUnsatCore :: Malformed unsat core `{unsatCore}`" + | throwError "{decl_name%} :: Malformed unsat core `{unsatCore}`" let mut ret := #[] for sexp in unsatCore do let .atom (.symb name) := sexp | continue if name.take 11 == "valid_fact_" then let .some n := (name.drop 11).toNat? - | throwError "validFactOfUnsatCore :: The id {name.drop 11} of {name} is invalid" + | throwError "{decl_name%} :: The id {name.drop 11} of {name} is invalid" ret := ret.push n return ret /-- Only put declarations in the query -/ def querySolver (query : Array IR.SMT.Command) : MetaM (Option (Sexp × String)) := do if !(auto.smt.get (← getOptions)) then - throwError "querySolver :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" if (auto.smt.solver.name.get (← getOptions) == .none) then - logInfo "querySolver :: Skipped" + logInfo s!"{decl_name%} :: Skipped" return .none let name := auto.smt.solver.name.get (← getOptions) let solver ← createSolver name @@ -175,7 +175,7 @@ def querySolver (query : Array IR.SMT.Command) : MetaM (Option (Sexp × String)) def saveQuery (query : Array IR.SMT.Command) : MetaM Unit := do if !(auto.smt.save.get (← getOptions)) then - throwError "querySolver :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" let path := auto.smt.savepath.get (← getOptions) IO.FS.withFile path IO.FS.Mode.write fun fd => for cmd in query do diff --git a/Auto/Solver/TPTP.lean b/Auto/Solver/TPTP.lean index f5c4f70..121b236 100644 --- a/Auto/Solver/TPTP.lean +++ b/Auto/Solver/TPTP.lean @@ -114,7 +114,7 @@ def queryZEPort (zept : ZEPortType) (query : String) : MetaM String := do catch e => let estr := toString (← (Exception.toMessageData e).format) if estr.extract ⟨0⟩ ⟨44⟩ != "already exists (error code: 17, file exists)" then - throwError "queryZEPort :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" idx := idx + (← IO.rand 0 100) IO.FS.withFile s!"./.zeport_ignore/problem{idx}.p" .writeNew (fun stream => stream.putStr query) let solver ← createSolver path idx @@ -149,7 +149,7 @@ def queryE (query : String) : MetaM String := do def querySolver (query : String) : MetaM (Array Parser.TPTP.Command) := do if !(auto.tptp.get (← getOptions)) then - throwError "querySolver :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" let stdout ← (do match auto.tptp.solver.name.get (← getOptions) with | .zipperposition => queryZipperposition query diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 20060ff..ddee40c 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -87,11 +87,11 @@ def parseUnfolds : TSyntax ``unfolds → TacticM (Array Prep.ConstUnfoldInfo) | `(unfolds| u[ $[$hs],* ]) => do let exprs ← hs.mapM (fun i => do let some expr ← Term.resolveId? i - | throwError "parseUnfolds :: Unknown identifier {i}. {defeqUnfoldErrHint}" + | throwError "{decl_name%} :: Unknown identifier {i}. {defeqUnfoldErrHint}" return expr) exprs.mapM (fun expr => do let some name := expr.constName? - | throwError "parseUnfolds :: Unknown declaration {expr}. {defeqUnfoldErrHint}" + | throwError "{decl_name%} :: Unknown declaration {expr}. {defeqUnfoldErrHint}" Prep.getConstUnfoldInfo name) | _ => throwUnsupportedSyntax @@ -99,11 +99,11 @@ def parseDefeqs : TSyntax ``defeqs → TacticM (Array Name) | `(defeqs| d[ $[$hs],* ]) => do let exprs ← hs.mapM (fun i => do let some expr ← Term.resolveId? i - | throwError "parseDefeqs :: Unknown identifier {i}. {defeqUnfoldErrHint}" + | throwError "{decl_name%} :: Unknown identifier {i}. {defeqUnfoldErrHint}" return expr) exprs.mapM (fun expr => do let some name := expr.constName? - | throwError "parseDefeqs :: Unknown declaration {expr}. {defeqUnfoldErrHint}" + | throwError "{decl_name%} :: Unknown declaration {expr}. {defeqUnfoldErrHint}" return name) | _ => throwUnsupportedSyntax @@ -121,12 +121,12 @@ def parseUOrDs (stxs : Array (TSyntax ``uord)) : TacticM (Array Prep.ConstUnfold match ← parseUOrD stx with | .inl u => if hasUnfolds then - throwError "Auto :: Duplicated unfold hint" + throwError "{decl_name%} :: Duplicated unfold hint" hasUnfolds := true unfolds := u | .inr d => if hasDefeqs then - throwError "Auto :: Duplicated defeq hint" + throwError "{decl_name%} :: Duplicated defeq hint" hasDefeqs := true defeqs := defeqs.append d return (unfolds, defeqs) @@ -262,7 +262,7 @@ def queryTPTP (exportFacts : Array REntry) : LamReif.ReifM (Array Embedding.Lam. let exportLamTerms ← exportFacts.mapM (fun re => do match re with | .valid [] t => return t - | _ => throwError "runAuto :: Unexpected error") + | _ => throwError "{decl_name%} :: Unexpected error") let query ← lam2TH0 lamVarTy lamEVarTy exportLamTerms trace[auto.tptp.printQuery] "\n{query}" let tptpProof ← Solver.TPTP.querySolver query @@ -276,7 +276,7 @@ def queryTPTP (exportFacts : Array REntry) : LamReif.ReifM (Array Embedding.Lam. let mut ret := #[] for n in unsatCore do let .some re := exportFacts[n]? - | throwError "queryTPTP :: Index {n} out of range" + | throwError "{decl_name%} :: Index {n} out of range" ret := ret.push re return ret @@ -287,7 +287,7 @@ def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : L let exportLamTerms ← exportFacts.mapM (fun re => do match re with | .valid [] t => return t - | _ => throwError "runAuto :: Unexpected error") + | _ => throwError "{decl_name%} :: Unexpected error") let sni : SMT.SMTNamingInfo := {tyVal := (← LamReif.getTyVal), varVal := (← LamReif.getVarVal), lamEVarTy := (← LamReif.getLamEVarTy)} let ((commands, validFacts), state) ← (lamFOL2SMT sni lamVarTy lamEVarTy exportLamTerms exportInds).run @@ -307,20 +307,20 @@ def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : L -- **Print STerms corresponding to `validFacts` in unsatCore** for id in unsatCoreIds do let .some sterm := validFacts[id]? - | throwError "runAuto :: Index {id} of `validFacts` out of range" + | throwError "{decl_name%} :: Index {id} of `validFacts` out of range" trace[auto.smt.unsatCore.smtTerms] "|valid_fact_{id}| : {sterm}" -- **Print Lean expressions correesponding to `validFacts` in unsatCore** SMT.withExprValuation sni state.h2lMap (fun tyValMap varValMap etomValMap => do for id in unsatCoreIds do let .some t := exportLamTerms[id]? - | throwError "runAuto :: Index {id} of `exportLamTerms` out of range" + | throwError "{decl_name%} :: Index {id} of `exportLamTerms` out of range" let e ← Lam2D.interpLamTermAsUnlifted tyValMap varValMap etomValMap 0 t trace[auto.smt.unsatCore.leanExprs] "|valid_fact_{id}| : {← Core.betaReduce e}" ) -- **Print derivation of unsatCore** for id in unsatCoreIds do let .some t := exportLamTerms[id]? - | throwError "runAuto :: Index {id} of `exportLamTerm` out of range" + | throwError "{decl_name%} :: Index {id} of `exportLamTerm` out of range" let vderiv ← LamReif.collectDerivFor (.valid [] t) trace[auto.smt.unsatCore.deriv] "|valid_fact_{id}| : {vderiv}" if auto.smt.rconsProof.get (← getOptions) then @@ -392,18 +392,18 @@ def callNative_direct (Lam2DAAF.callNativeWithAtomAsFVar nonemptiesWithDTr validsWithDTr prover).run' { tyVal := tyVal, varVal := varVal, lamEVarTy := lamEVarTy } if usedEtoms.size != 0 then - throwError "callNative_direct :: etoms should not occur here" + throwError "{decl_name%} :: etoms should not occur here" let ss ← usedInhs.mapM (fun re => do match ← lookupREntryProof! re with | .inhabitation e _ _ => return e | .chkStep (.n (.nonemptyOfAtom n)) => match varVal[n]? with | .some (e, _) => return e - | .none => throwError "callNative_direct :: Unexpected error" - | _ => throwError "callNative_direct :: Cannot find external proof of {re}") + | .none => throwError "{decl_name%} :: Unexpected error" + | _ => throwError "{decl_name%} :: Cannot find external proof of {re}") let ts ← usedHyps.mapM (fun re => do let .assertion e _ _ ← lookupREntryProof! re - | throwError "callNative_direct :: Cannot find external proof of {re}" + | throwError "{decl_name%} :: Cannot find external proof of {re}" return e) return mkAppN proof (ss ++ ts) @@ -484,7 +484,7 @@ def evalAuto : Tactic -- now the goal is just `G` let (goalBinders, newGoal) ← (← getMainGoal).intros let [nngoal] ← newGoal.apply (.const ``Classical.byContradiction []) - | throwError "evalAuto :: Unexpected result after applying Classical.byContradiction" + | throwError "{decl_name%} :: Unexpected result after applying Classical.byContradiction" let (ngoal, absurd) ← MVarId.intro1 nngoal replaceMainGoal [absurd] withMainContext do @@ -505,7 +505,7 @@ def evalIntromono : Tactic | `(intromono | intromono $hints $[$uords]*) => withMainContext do let (goalBinders, newGoal) ← (← getMainGoal).intros let [nngoal] ← newGoal.apply (.const ``Classical.byContradiction []) - | throwError "evalAuto :: Unexpected result after applying Classical.byContradiction" + | throwError "{decl_name%} :: Unexpected result after applying Classical.byContradiction" let (ngoal, absurd) ← MVarId.intro1 nngoal replaceMainGoal [absurd] withMainContext do @@ -552,12 +552,12 @@ def runNativeProverWithAuto if let .some expr ← queryNative declName? exportFacts exportInhs prover then return expr else - throwError "runNativeProverWithAuto :: Failed to find proof") + throwError "{decl_name%} :: Failed to find proof") let (proof, _) ← Monomorphization.monomorphize lemmas inhFacts (@id (Reif.ReifM Expr) do let s ← get let u ← computeMaxLevel s.facts (afterReify s.facts s.inhTys).run' {u := u}) - trace[auto.tactic] "runNativeProverWithAuto :: Found proof of {← Meta.inferType proof}" + trace[auto.tactic] "{decl_name%} :: Found proof of {← Meta.inferType proof}" return proof @[tactic mononative] @@ -568,7 +568,7 @@ def evalMonoNative : Tactic -- now the goal is just `G` let (goalBinders, newGoal) ← (← getMainGoal).intros let [nngoal] ← newGoal.apply (.const ``Classical.byContradiction []) - | throwError "evalAuto :: Unexpected result after applying Classical.byContradiction" + | throwError "{decl_name%} :: Unexpected result after applying Classical.byContradiction" let (ngoal, absurd) ← MVarId.intro1 nngoal replaceMainGoal [absurd] withMainContext do diff --git a/Auto/Translation/Assumptions.lean b/Auto/Translation/Assumptions.lean index 17ff339..c332a03 100644 --- a/Auto/Translation/Assumptions.lean +++ b/Auto/Translation/Assumptions.lean @@ -73,7 +73,7 @@ def Lemma.betaReduceType (lem : Lemma) : CoreM Lemma := do /-- Create a `Lemma` out of a constant, given the name of the constant -/ def Lemma.ofConst (name : Name) (deriv : DTr) : CoreM Lemma := do let .some decl := (← getEnv).find? name - | throwError "Lemma.ofConst :: Unknown constant {name}" + | throwError "{decl_name%} :: Unknown constant {name}" let type := decl.type let params := decl.levelParams return ⟨⟨.const name (params.map Level.param), type, deriv⟩, ⟨params⟩⟩ @@ -125,7 +125,7 @@ def Lemma.reorderForallInstDep (lem : Lemma) : MetaM Lemma := do def Lemma.rewriteUMonoRigid? (lem : Lemma) (rw : UMonoFact) : MetaM (Option Lemma) := do let ⟨rwproof, rwtype, rwDeriv⟩ := rw let .some (α, lhs, rhs) ← Meta.matchEq? rwtype - | throwError "Lemma.rewriteUMonoRigid :: {rwtype} is not an equality" + | throwError "{decl_name%} :: {rwtype} is not an equality" let ⟨⟨proof, e, lemDeriv⟩, params⟩ := lem let eAbst ← Meta.kabstract e lhs unless eAbst.hasLooseBVars do @@ -133,7 +133,7 @@ def Lemma.rewriteUMonoRigid? (lem : Lemma) (rw : UMonoFact) : MetaM (Option Lemm let eNew := eAbst.instantiate1 rhs let motive := mkLambda `_a BinderInfo.default α eAbst unless (← Meta.isTypeCorrect motive) do - throwError "Lemma.rewriteUMonoRigid :: Motive {motive} is not type correct" + throwError "{decl_name%} :: Motive {motive} is not type correct" let eqPrf ← Meta.mkEqNDRec motive proof rwproof return .some ⟨⟨eqPrf, eNew, .node "rw" #[lemDeriv, rwDeriv]⟩, params⟩ @@ -149,16 +149,16 @@ def Lemma.rewriteUPolyRigid (lem : Lemma) (rw : Lemma) : MetaM Lemma := do let s ← saveState -- Test whether `rhs` contains `lhs let .some (_, lhs, rhs) ← Meta.matchEq? rw.type - | throwError "Lemma.rewriteUMonoRigid :: {rw.type} is not an equality" + | throwError "{decl_name%} :: {rw.type} is not an equality" let lhs' := lhs.instantiateLevelParamsArray rw.params (← rw.params.mapM (fun _ => Meta.mkFreshLevelMVar)) let rhs' := rhs.instantiateLevelParamsArray rw.params (← rw.params.mapM (fun _ => Meta.mkFreshLevelMVar)) if (← Meta.kabstract rhs' lhs').hasLooseBVars then - throwError "Lemma.rewriteUPolyRigid :: Right-hand side {rhs} of equality contains left-hand side {lhs}" + throwError "{decl_name%} :: Right-hand side {rhs} of equality contains left-hand side {lhs}" restoreState s while true do let umvars ← rw.params.mapM (fun _ => Meta.mkFreshLevelMVar) let .some urw := (rw.instantiateLevelParamsArray umvars).toUMonoFact? - | throwError "Lemma.rewriteUPolyRigid :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let .some lem' ← Lemma.rewriteUMonoRigid? lem urw | break let restmvars := (← umvars.mapM Level.collectLevelMVars).concatMap id @@ -219,7 +219,7 @@ def LemmaInst.ofLemmaLeadingDepOnly (lem : Lemma) : MetaM LemmaInst := do let nld := Expr.numLeadingDepArgs type Meta.forallBoundedTelescope type nld fun xs _ => do if xs.size != nld then - throwError "LemmaInst.ofLemmaLeadingDepOnly :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" let proof ← Meta.mkLambdaFVars xs (mkAppN proof xs) let lem' : Lemma := ⟨⟨proof, type, deriv⟩, params⟩ return ⟨lem', xs.size, xs.size⟩ @@ -235,19 +235,19 @@ instance : ToMessageData LemmaInst where def LemmaInst.subsumeQuick (li₁ li₂ : LemmaInst) : MetaM Bool := Meta.withNewMCtxDepth <| do if li₁.nargs != li₂.nargs then - throwError "LemmaInst.subsumeQuick :: {li₁} and {li₂} are not instance of the same lemma" + throwError "{decl_name%} :: {li₁} and {li₂} are not instance of the same lemma" let lem₁ := li₁.toLemma let lem₂ := li₂.toLemma let (_, _, body₂) ← Meta.lambdaMetaTelescope lem₂.proof li₂.nbinders let args₂ := Expr.getAppBoundedArgs li₂.nargs body₂ if args₂.size != li₂.nargs then - throwError "LemmaInst.subsumeQuick :: {li₂} is expected to have {li₂.nargs} args, but actually has {args₂.size}" + throwError "{decl_name%} :: {li₂} is expected to have {li₂.nargs} args, but actually has {args₂.size}" Meta.withNewMCtxDepth do let paramInst₁ ← lem₁.params.mapM (fun _ => Meta.mkFreshLevelMVar) let (_, _, body₁) ← Meta.lambdaMetaTelescope lem₁.proof li₁.nbinders let args₁ := Expr.getAppBoundedArgs li₁.nargs body₁ if args₁.size != li₁.nargs then - throwError "LemmaInst.subsumeQuick :: {li₁} is expected to have {li₁.nargs} args, but actually has {args₁.size}" + throwError "{decl_name%} :: {li₁} is expected to have {li₁.nargs} args, but actually has {args₁.size}" let args₁ := args₁.map (fun e => e.instantiateLevelParamsArray lem₁.params paramInst₁) for (arg₁, arg₂) in args₁.zip args₂ do if !(← Meta.isDefEq arg₁ arg₂) then @@ -288,10 +288,10 @@ def MLemmaInst.ofLemmaInst (li : LemmaInst) : MetaM (Array Level × Array Expr let type := type.instantiateLevelParamsArray params lvls let (mvars, _, proof) ← Meta.lambdaMetaTelescope proof li.nbinders let .some origProof := Expr.getAppFnN li.nargs proof - | throwError "MLemmaInst.ofLemmaInst :: Insufficient number of arguments" + | throwError "{decl_name%} :: Insufficient number of arguments" let args := Expr.getAppBoundedArgs li.nargs proof if args.size != li.nargs then - throwError "MLemmaInst.ofLemmaInst :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" let type ← Meta.instantiateForall type mvars return (lvls, mvars, ⟨origProof, args, type, deriv⟩) @@ -340,7 +340,7 @@ partial def collectUniverseLevels : Expr → MetaM (Std.HashSet Level) return mergeHashSet (mergeHashSet tys vs) bodys | .lit _ => return Std.HashSet.empty.insert (.succ .zero) | .mdata _ e' => collectUniverseLevels e' -| .proj .. => throwError "Please unfold projections before collecting universe levels" +| .proj .. => throwError "{decl_name%} :: Please unfold projections before collecting universe levels" def computeMaxLevel (facts : Array UMonoFact) : MetaM Level := do let levels ← facts.foldlM (fun hs ⟨_, ty, _⟩ => do diff --git a/Auto/Translation/Inductive.lean b/Auto/Translation/Inductive.lean index 52b6083..fff8a96 100644 --- a/Auto/Translation/Inductive.lean +++ b/Auto/Translation/Inductive.lean @@ -18,7 +18,7 @@ namespace Auto -/ def isFamily (tyctorname : Name) : CoreM Bool := do let .some (.inductInfo val) := (← getEnv).find? tyctorname - | throwError "isFamily :: {tyctorname} is not a type constructor" + | throwError "{decl_name%} :: {tyctorname} is not a type constructor" return (Expr.forallBinders val.type).size != val.numParams /-- @@ -26,7 +26,7 @@ def isFamily (tyctorname : Name) : CoreM Bool := do -/ def isIndProp (tyctorname : Name) : CoreM Bool := do let .some (.inductInfo val) := (← getEnv).find? tyctorname - | throwError "isIndProp :: {tyctorname} is not a type constructor" + | throwError "{decl_name%} :: {tyctorname} is not a type constructor" (Meta.withTransparency (n := MetaM) .all <| Meta.forallTelescopeReducing val.type fun _ body => Meta.isDefEq body (.sort .zero)).run' {} @@ -36,7 +36,7 @@ def isIndProp (tyctorname : Name) : CoreM Bool := do -/ def isSimpleCtor (ctorname : Name) : CoreM Bool := do let .some (.ctorInfo val) := (← getEnv).find? ctorname - | throwError "isSimpleCtor :: {ctorname} is not a type constructor" + | throwError "{decl_name%} :: {ctorname} is not a type constructor" Meta.MetaM.run' <| Meta.forallBoundedTelescope val.type val.numParams fun _ body => pure ((Expr.depArgs body).size == 0) @@ -46,7 +46,7 @@ def isSimpleCtor (ctorname : Name) : CoreM Bool := do -/ def isSimpleInductive (tyctorname : Name) : CoreM Bool := do let .some (.inductInfo val) := (← getEnv).find? tyctorname - | throwError "isSimple :: {tyctorname} is not a type constructor" + | throwError "{decl_name%} :: {tyctorname} is not a type constructor" return (← val.ctors.allM isSimpleCtor) && !(← isFamily tyctorname) structure SimpleIndVal where @@ -90,7 +90,7 @@ abbrev IndCollectM := StateRefT CollectInduct.State MetaM private def collectSimpleInduct (tyctor : Name) (lvls : List Level) (args : Array Expr) : MetaM SimpleIndVal := do let .some (.inductInfo val) := (← getEnv).find? tyctor - | throwError "collectSimpleInduct :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let ctors ← (Array.mk val.ctors).mapM (fun ctorname => do let instctor := mkAppN (Expr.const ctorname lvls) args let type ← Meta.inferType instctor @@ -100,7 +100,7 @@ private def collectSimpleInduct let projs ← (getStructureInfo? env tyctor).mapM (fun si => do si.fieldNames.mapM (fun fieldName => do let .some projFn := getProjFnForField? env tyctor fieldName - | throwError "collectSimpleInduct :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return mkAppN (Expr.const projFn lvls) args)) return ⟨tyctor, mkAppN (Expr.const tyctor lvls) args, ctors, projs⟩ @@ -128,7 +128,7 @@ mutual if !(← getRecorded).contains tyctor then setRecorded ((← getRecorded).insert tyctor #[]) let .some arr := (← getRecorded).get? tyctor - | throwError "collectAppInstSimpleInduct :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" for e' in arr do if ← Meta.isDefEq e e' then return @@ -151,9 +151,9 @@ mutual return collectExprSimpleInduct ty collectExprSimpleInduct body - | .letE .. => throwError "collectExprSimpleInduct :: Let-expressions should have been reduced" - | .mdata .. => throwError "collectExprSimpleInduct :: mdata should have been consumed" - | .proj .. => throwError "collectExprSimpleInduct :: Projections should have been turned into ordinary expressions" + | .letE .. => throwError "{decl_name%} :: Let-expressions should have been reduced" + | .mdata .. => throwError "{decl_name%} :: mdata should have been consumed" + | .proj .. => throwError "{decl_name%} :: Projections should have been turned into ordinary expressions" | e => collectAppInstSimpleInduct e end diff --git a/Auto/Translation/Inhabitation.lean b/Auto/Translation/Inhabitation.lean index 14cf56e..ff21a13 100644 --- a/Auto/Translation/Inhabitation.lean +++ b/Auto/Translation/Inhabitation.lean @@ -96,7 +96,7 @@ private def inhFactMatchAtomTysAux (inhTy : Lemma) (atomTys : Array Expr) : Meta let s ← saveState let (_, _, mi) ← MLemmaInst.ofLemmaInst li let .some e := getExprNonImpPosition mip mi.type - | throwError "inhFactMatchAtomTys :: Unexpected error, cannot get position {mip} from {mi.type}" + | throwError "{decl_name%} :: Unexpected error, cannot get position {mip} from {mi.type}" for a in atomTys do let s' ← saveState if (← Meta.isDefEq e a) then @@ -134,6 +134,6 @@ where for a in atomTys do if (← Meta.withNewMCtxDepth (Meta.isDefEq inhTy a)) then return inhTy - throwError "inhFactMatchAtomTys :: Unable to canonicalize {inhTy}" + throwError "{decl_name%} :: Unable to canonicalize {inhTy}" end Auto.Inhabitation diff --git a/Auto/Translation/Lam2DAtomAsFVar.lean b/Auto/Translation/Lam2DAtomAsFVar.lean index ce5f137..e9c0bfa 100644 --- a/Auto/Translation/Lam2DAtomAsFVar.lean +++ b/Auto/Translation/Lam2DAtomAsFVar.lean @@ -71,7 +71,7 @@ def withTypeAtomsAsFVar (atoms : Array Nat) : ExternM Unit := if (← getTypeAtomFVars).contains atom then continue let .some (e, lvl) := (← getTyVal)[atom]? - | throwError "withTypeAtomAsFVar :: Unknown type atom {atom}" + | throwError "{decl_name%} :: Unknown type atom {atom}" let name := (`_exTy).appendIndexAfter (← getTypeAtomFVars).size let newFVarId ← withLocalDecl name .default (.sort lvl) .default setAtomsToAbstract ((← getAtomsToAbstract).push (newFVarId, e)) @@ -82,7 +82,7 @@ def withTermAtomsAsFVar (atoms : Array Nat) : ExternM Unit := if (← getTermAtomFVars).contains atom then continue let .some (e, s) := (← getVarVal)[atom]? - | throwError "withTermAtomAsFVar :: Unknown term atom {atom}" + | throwError "{decl_name%} :: Unknown term atom {atom}" let sinterp ← Lam2D.interpLamSortAsUnlifted (← getTypeAtomFVars) s let name := (`e!).appendIndexAfter (← getTermAtomFVars).size let newFVarId ← withLocalDecl name .default sinterp .default @@ -94,7 +94,7 @@ def withEtomsAsFVar (etoms : Array Nat) : ExternM Unit := if (← getEtomFVars).contains etom then return let .some s := (← getLamEVarTy)[etom]? - | throwError "withEtomAsFVar :: Unknown etom {etom}" + | throwError "{decl_name%} :: Unknown etom {etom}" let sinterp ← Lam2D.interpLamSortAsUnlifted (← getTypeAtomFVars) s let name := (`e?).appendIndexAfter (← getEtomFVars).size let newFVarId ← withLocalDecl name .default sinterp .default @@ -147,7 +147,7 @@ def callNativeWithAtomAsFVar let ss ← nonemptiesWithDTr.mapM (fun (re, _) => do match re with | .nonempty s => return s - | _ => throwError "callNativeWithAtomAsFVar :: {re} is not a `nonempty` entry") + | _ => throwError "{decl_name%} :: {re} is not a `nonempty` entry") let inhs ← withTranslatedLamSorts ss for inh in inhs do trace[auto.lam2D.printInhs] "{inh}" @@ -158,13 +158,13 @@ def callNativeWithAtomAsFVar let ts ← valids.mapM (fun re => do match re with | .valid [] t => return t - | _ => throwError "callNativeWithAtomAsFVar :: {re} is not a `valid` entry") + | _ => throwError "{decl_name%} :: {re} is not a `valid` entry") let hyps ← withTranslatedLamTerms ts for hyp in hyps do if !(← runMetaM <| Meta.isTypeCorrect hyp) then - throwError "callNative :: Malformed hypothesis {hyp}" + throwError "{decl_name%} :: Malformed hypothesis {hyp}" if !(← runMetaM <| Meta.isProp hyp) then - throwError "callNative :: Hypothesis {hyp} is not a proposition" + throwError "{decl_name%} :: Hypothesis {hyp} is not a proposition" trace[auto.lam2D.printHyps] "{hyp}" let hyps ← runMetaM <| hyps.mapM (fun e => Core.betaReduce e) let hypFvars ← withHyps hyps @@ -207,7 +207,7 @@ def callNativeWithAtomAsFVar let proofLamTermPre := proofLamTermPre.abstractsRevImp ((Array.mk usedEtoms).map LamTerm.etom) let usedEtomTys ← usedEtoms.mapM (fun etom => do let .some ty := lamEVarTy[etom]? - | throwError "callNative :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return ty) let proof := mkAppN expr ⟨usedVals⟩ let proofLamTerm := usedEtomTys.foldr (fun s cur => LamTerm.mkForallEF s cur) proofLamTermPre diff --git a/Auto/Translation/Lam2TH0.lean b/Auto/Translation/Lam2TH0.lean index b1996fa..48eb783 100644 --- a/Auto/Translation/Lam2TH0.lean +++ b/Auto/Translation/Lam2TH0.lean @@ -26,11 +26,11 @@ def lam2TH0 (lamVarTy : Array LamSort) (lamEVarTy : Array LamSort) (facts : Arra bvHs.toList.map (fun bvc => s!"thf(typedecl_bv{transBitVecConst bvc}, type, t_bv{transBitVecConst bvc}: {transBitVecConstSort bvc}).") ++ (← termHs.toList.mapM (fun i => do let .some s := lamVarTy.get? i - | throwError "lam2TH0 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return s!"thf(typedecl_t_a{i}, type, t_a{i}: {transLamSort s}).")) ++ (← etomHs.toList.mapM (fun i => do let .some s := lamEVarTy.get? i - | throwError "lam2TH0 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return s!"thf(typedecl_t_e{i}, type, t_e{i}: {transLamSort s}).")) -- Empty type is not inhabited if (lamVarTy ++ lamEVarTy).any (fun s => s == .base .empty) then diff --git a/Auto/Translation/LamFOL2SMT.lean b/Auto/Translation/LamFOL2SMT.lean index 29f5090..912c5b3 100644 --- a/Auto/Translation/LamFOL2SMT.lean +++ b/Auto/Translation/LamFOL2SMT.lean @@ -55,15 +55,15 @@ def LamAtomic.toLeanExpr match atomic with | .sort n => do let .some e := tyValMap.get? n - | throwError "SMT.printValuation :: Unknown sort atom {n}" + | throwError "{decl_name%} :: Unknown sort atom {n}" return e | .term n => do let .some e := varValMap.get? n - | throwError "SMT.printValuation :: Unknown term atom {n}" + | throwError "{decl_name%} :: Unknown term atom {n}" return e | .etom n => do let .some e := etomValMap.get? n - | throwError "SMT.printValuation :: Unknown etom {n}" + | throwError "{decl_name%} :: Unknown etom {n}" return e | .bvOfNat n => do let e := Expr.app (.const ``BitVec.ofNat []) (.lit (.natVal n)) @@ -95,7 +95,7 @@ where go : LamSort → MetaM String | .atom n => do let .some (e, _) := sni.tyVal[n]? - | throwError "lamSortAtom2String :: Unexpected sort atom {repr (LamSort.atom n)}" + | throwError "{decl_name%} :: Unexpected sort atom {repr (LamSort.atom n)}" exprToSuggestion e | .base b => return b.toString -- Suggest name based on return type @@ -109,11 +109,11 @@ where go : LamAtomic → MetaM String | .sort n => do let .some (e, _) := sni.tyVal[n]? - | throwError "lamSortAtom2String :: Unexpected sort atom {repr (LamSort.atom n)}" + | throwError "{decl_name%} :: Unexpected sort atom {repr (LamSort.atom n)}" exprToSuggestion e | .term n => do let .some (e, _) := sni.varVal[n]? - | throwError "lamTermAtom2String :: Unexpected term atom {repr (LamTerm.atom n)}" + | throwError "{decl_name%} :: Unexpected term atom {repr (LamTerm.atom n)}" exprToSuggestion e | .etom n => return s!"sk{n}" | .bvOfNat n => exprToSuggestion (Expr.app (.const ``BitVec.ofNat []) (.lit (.natVal n))) @@ -145,7 +145,7 @@ private def lamSortAtom2String (sni : SMTNamingInfo) (n : Nat) : TransM LamAtomi private def lamSort2SSortAux (sni : SMTNamingInfo) : LamSort → TransM LamAtomic SSort | .atom n => do return .app (.symb (← lamSortAtom2String sni n)) #[] | .base b => return lamBaseSort2SSort b -| .func _ _ => throwError "lamSort2STermAux :: Unexpected error. Higher order input?" +| .func _ _ => throwError "{decl_name%} :: Unexpected error. Higher order input?" /-- Only translates first-order types -/ private def lamSort2SSort (sni : SMTNamingInfo) : LamSort → TransM LamAtomic (List SSort × SSort) @@ -193,7 +193,7 @@ private def bitVec2STerm (n i : Nat) : STerm := private def lamBaseTerm2STerm_Arity3 (arg1 arg2 arg3 : STerm) : LamBaseTerm → TransM LamAtomic STerm | .scst .srepall => return .qStrApp "str.replace_all" #[arg1, arg2, arg3] -| t => throwError "lamTerm2STerm :: The arity of {repr t} is not 3" +| t => throwError "{decl_name%} :: The arity of {repr t} is not 3" private def lamBaseTerm2STerm_Arity2 (arg1 arg2 : STerm) : LamBaseTerm → TransM LamAtomic STerm | .pcst .and => return .qStrApp "and" #[arg1, arg2] @@ -248,7 +248,7 @@ private def lamBaseTerm2STerm_Arity2 (arg1 arg2 : STerm) : LamBaseTerm → Trans | .bvcst (.bvshOp _ smt? op) => match smt? with | false => - throwError "lamTerm2STerm :: Non-smt shift operations should not occur here" + throwError "{decl_name%} :: Non-smt shift operations should not occur here" | true => match op with | .shl => return .qStrApp "bvshl" #[arg1, arg2] @@ -256,7 +256,7 @@ private def lamBaseTerm2STerm_Arity2 (arg1 arg2 : STerm) : LamBaseTerm → Trans | .ashr => return .qStrApp "bvashr" #[arg1, arg2] | .bvcst (.bvappend _ _) => return .qStrApp "concat" #[arg1, arg2] | .ocst (.smtAttr1T name _ _) => return .attrApp name arg1 arg2 -| t => throwError "lamTerm2STerm :: The arity of {repr t} is not 2" +| t => throwError "{decl_name%} :: The arity of {repr t} is not 2" private def lamBaseTerm2STerm_Arity1 (sni : SMTNamingInfo) (arg : STerm) : LamBaseTerm → TransM LamAtomic STerm | .pcst .not => return .qStrApp "not" #[arg] @@ -316,7 +316,7 @@ private def lamBaseTerm2STerm_Arity1 (sni : SMTNamingInfo) (arg : STerm) : LamBa return .qIdApp (.ident (.indexed "sign_extend" #[.inr (v - w)])) #[arg] else return .qIdApp (.ident (.indexed "extract" #[.inr (v - 1), .inr 0])) #[arg] -| t => throwError "lamTerm2STerm :: The arity of {repr t} is not 1" +| t => throwError "{decl_name%} :: The arity of {repr t} is not 1" where solverName : MetaM Solver.SMT.SolverName := do return Solver.SMT.auto.smt.solver.name.get (← getOptions) @@ -332,11 +332,11 @@ private def lamBaseTerm2STerm_Arity0 : LamBaseTerm → TransM LamAtomic STerm | .ncst (.natVal n) => return .sConst (.num n) | .scst (.strVal s) => return .sConst (.str s) | .bvcst (.bvVal n i) => return bitVec2STerm n i -| t => throwError "lamTerm2STerm :: The arity of {repr t} is not 0" +| t => throwError "{decl_name%} :: The arity of {repr t} is not 0" def lamTermAtom2String (sni : SMTNamingInfo) (lamVarTy : Array LamSort) (n : Nat) : TransM LamAtomic (LamSort × String) := do let .some s := lamVarTy[n]? - | throwError "lamTermAtom2String :: Unexpected term atom {repr (LamTerm.atom n)}" + | throwError "{decl_name%} :: Unexpected term atom {repr (LamTerm.atom n)}" -- Empty type is not inhabited if s == .base .empty then addCommand (.assert (.qStrApp "false" #[])) @@ -349,7 +349,7 @@ def lamTermAtom2String (sni : SMTNamingInfo) (lamVarTy : Array LamSort) (n : Nat def lamTermEtom2String (sni : SMTNamingInfo) (lamEVarTy : Array LamSort) (n : Nat) : TransM LamAtomic (LamSort × String) := do let .some s := lamEVarTy[n]? - | throwError "lamTerm2STerm :: Unexpected etom {repr (LamTerm.etom n)}" + | throwError "{decl_name%} :: Unexpected etom {repr (LamTerm.etom n)}" -- Empty type is not inhabited if s == .base .empty then addCommand (.assert (.qStrApp "false" #[])) @@ -365,12 +365,12 @@ private def lamTerm2STermAux (sni : SMTNamingInfo) (lamVarTy lamEVarTy : Array L | .atom n => do let (s, name) ← lamTermAtom2String sni lamVarTy n if args.size != s.getArgTys.length then - throwError "lamTerm2STerm :: Argument number mismatch. Higher order input?" + throwError "{decl_name%} :: Argument number mismatch. Higher order input?" return .qIdApp (QualIdent.ofString name) args | .etom n => do let (s, name) ← lamTermEtom2String sni lamEVarTy n if args.size != s.getArgTys.length then - throwError "lamTerm2STerm :: Argument number mismatch. Higher order input?" + throwError "{decl_name%} :: Argument number mismatch. Higher order input?" return .qIdApp (QualIdent.ofString name) args | .base b => match args with @@ -378,8 +378,8 @@ private def lamTerm2STermAux (sni : SMTNamingInfo) (lamVarTy lamEVarTy : Array L | #[u₁] => lamBaseTerm2STerm_Arity1 sni u₁ b | #[u₁, u₂] => lamBaseTerm2STerm_Arity2 u₁ u₂ b | #[u₁, u₂, u₃] => lamBaseTerm2STerm_Arity3 u₁ u₂ u₃ b - | _ => throwError "lamTerm2STerm :: Argument number mismatch. Higher order input?" -| t => throwError "lamTerm2STerm :: Unexpected head term {repr t}" + | _ => throwError "{decl_name%} :: Argument number mismatch. Higher order input?" +| t => throwError "{decl_name%} :: Unexpected head term {repr t}" def lamQuantified2STerm (sni : SMTNamingInfo) (forall? : Bool) (s : LamSort) (body : TransM LamAtomic STerm) : TransM LamAtomic STerm := do -- Empty type is not inhabited @@ -400,13 +400,13 @@ private partial def lamTerm2STerm (sni : SMTNamingInfo) (lamVarTy lamEVarTy : Ar | .base b => lamBaseTerm2STerm_Arity0 b | .bvar n => return .bvar n | .app _ (.app _ (.base (.eqI _)) _) _ => - throwError ("lamTerm2STerm :: " ++ LamExportUtils.exportError.ImpPolyLog) + throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) | .app _ (.base (.forallEI _)) (.lam _ _) => - throwError ("lamTerm2STerm :: " ++ LamExportUtils.exportError.ImpPolyLog) + throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) | .app _ (.base (.existEI _)) (.lam _ _) => - throwError ("lamTerm2STerm :: " ++ LamExportUtils.exportError.ImpPolyLog) + throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) | .app _ (.app _ (.app _ (.base (.iteI _)) _) _) _ => - throwError ("lamTerm2STerm :: " ++ LamExportUtils.exportError.ImpPolyLog) + throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) | .app _ (.app _ (.base (.eq _)) arg₁) arg₂ => do let arg₁' ← lamTerm2STerm sni lamVarTy lamEVarTy arg₁ let arg₂' ← lamTerm2STerm sni lamVarTy lamEVarTy arg₂ @@ -442,16 +442,16 @@ private def lamMutualIndInfo2STerm (sni : SMTNamingInfo) (mind : MutualIndInfo) -- be declared during the following `lamSort2SSort` for ⟨type, _, _⟩ in mind do let .atom sn := type - | throwError "lamMutualIndInfo2STerm :: Inductive type {type} is not a sort atom" + | throwError "{decl_name%} :: Inductive type {type} is not a sort atom" -- Do not use `lamSortAtom2String` because we don't want to `declare-sort` let _ ← h2Symb (.sort sn) (← sni.suggestNameForAtomic (.sort sn)) for ⟨type, ctors, projs⟩ in mind do let .atom sn := type - | throwError "lamMutualIndInfo2STerm :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let mut projInfos : Array (LamSort × String) := #[] if let .some projs := projs then if ctors.length != 1 then - throwError "lamMutualIndInfo2STerm :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" for (s, t) in projs do let mut projname := "" match t with @@ -477,7 +477,7 @@ private def lamMutualIndInfo2STerm (sni : SMTNamingInfo) (mind : MutualIndInfo) let mut selDecls := #[] if projs.isSome then if argTys.length != projInfos.size then - throwError "lamMutualIndInfo2STerm :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" selDecls := ((Array.mk argTys).zip projInfos).map (fun (argTy, _, name) => (name, argTy)) else selDecls := (Array.mk argTys).zipWithIndex.map (fun (argTy, idx) => @@ -534,7 +534,7 @@ def withExprValuation match atomic with | .etom n => .some (n, name) | _ => .none) let declInfos ← etomsWithName.mapM (fun (n, name) => do let .some s := sni.lamEVarTy[n]? - | throwError "SMT.printValuation :: Unknown etom {n}" + | throwError "{decl_name%} :: Unknown etom {n}" let type ← Lam2D.interpLamSortAsUnlifted tyValMap s return (Name.mkSimple name, .default, fun _ => pure type)) Meta.withLocalDecls declInfos (fun etomFVars => do diff --git a/Auto/Translation/LamReif.lean b/Auto/Translation/LamReif.lean index 49ddc87..6a1f1f4 100644 --- a/Auto/Translation/LamReif.lean +++ b/Auto/Translation/LamReif.lean @@ -141,7 +141,7 @@ def printProofs : ReifM Unit := do for re in (← getRTable) do if let .some cs := chkMap.get? re then let .some n := (← getRst).findPos? re - | throwError "printProofs :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let etoms := match (← getChkStepCache).get? cs with | .some (_, arr) => @@ -152,13 +152,13 @@ def printProofs : ReifM Unit := do match re with | .valid [] t => let .some (expr, _, n) := (← getAssertions).get? t - | throwError "printProofs :: Unable to find assertion associated with {t}" + | throwError "{decl_name%} :: Unable to find assertion associated with {t}" trace[auto.lamReif.printProofs] "{n} : External fact ⦗⦗{← Meta.inferType expr}⦘⦘ proves ⦗⦗{re}⦘⦘" | .nonempty s => let .some (expr, _, n) := (← getInhabitations).get? s - | throwError "printProofs :: Unable to find inhabitation fact associated with {s}" + | throwError "{decl_name%} :: Unable to find inhabitation fact associated with {s}" trace[auto.lamReif.printProofs] "{n} : Inhabitation fact ⦗⦗{← Meta.inferType expr}⦘⦘ proves ⦗⦗{re}⦘⦘" - | _ => throwError "printProofs :: Unexpected entry {re}" + | _ => throwError "{decl_name%} :: Unexpected entry {re}" def sort2LamILTyIdx (s : LamSort) : ReifM Nat := do let isomTyMap ← getIsomTyMap @@ -175,32 +175,32 @@ def lookupTyVal! (n : Nat) : ReifM (Expr × Level) := do if let .some r := (← getTyVal)[n]? then return r else - throwError "lookupTyVal! :: Unknown type atom {n}" + throwError "{decl_name%} :: Unknown type atom {n}" /-- Lookup valuation of term atom -/ def lookupVarVal! (n : Nat) : ReifM (Expr × LamSort) := do if let .some r := (← getVarVal)[n]? then return r else - throwError "lookupVarVal! :: Unknown term atom {n}" + throwError "{decl_name%} :: Unknown term atom {n}" def lookupLamILTy! (idx : Nat) : ReifM LamSort := do if let .some s := (← getLamILTy)[idx]? then return s else - throwError "lookupLamILTy! :: Unknown index {idx}" + throwError "{decl_name%} :: Unknown index {idx}" def lookupAssertion! (t : LamTerm) : ReifM (Expr × DTr × LamTerm × Nat) := do if let .some r := (← getAssertions).get? t then return r else - throwError "lookupAssertion! :: Unknown assertion {t}" + throwError "{decl_name%} :: Unknown assertion {t}" def lookupRTable! (pos : Nat) : ReifM REntry := do if let .some r := (← getRTable).get? pos then return r else - throwError "lookupRTable! :: Unknown REntry {pos}" + throwError "{decl_name%} :: Unknown REntry {pos}" def lookupREntryPos! (re : REntry) : ReifM Nat := do match (← getRst).findPos? re with @@ -210,12 +210,12 @@ def lookupREntryPos! (re : REntry) : ReifM Nat := do | .valid [] t => match (← getAssertions).get? t with | .some (_, _, _, n) => return n - | .none => throwError "lookupREntryPos! :: Unknown REntry {re}" + | .none => throwError "{decl_name%} :: Unknown REntry {re}" | .nonempty s => match (← getInhabitations).get? s with | .some (_, _, n) => return n - | .none => throwError "lookupREntryPos! :: Unknown REntry {re}" - | _ => throwError "lookupREntryPos! :: Unknown REntry {re}" + | .none => throwError "{decl_name%} :: Unknown REntry {re}" + | _ => throwError "{decl_name%} :: Unknown REntry {re}" inductive REntryProof where | chkStep : ChkStep → REntryProof @@ -240,31 +240,31 @@ def lookupREntryProof? (re : REntry) : ReifM (Option REntryProof) := do def lookupREntryProof! (re : REntry) : ReifM REntryProof := do match ← lookupREntryProof? re with | .some proof => return proof - | .none => throwError "lookupREntryProof! :: Unknown REntry {re}" + | .none => throwError "{decl_name%} :: Unknown REntry {re}" def lookupLamEVarTy! (idx : Nat) : ReifM LamSort := do if let .some s := (← getLamEVarTy)[idx]? then return s else - throwError "lookupLamEVarTy! :: Unknown etom {idx}" + throwError "{decl_name%} :: Unknown etom {idx}" def lookupChkStepEtom! (cs : ChkStep) : ReifM (Array Nat) := do if let .some (_, arr) := (← getChkStepCache).get? cs then return arr else - throwError "lookupChkStepEtom! :: ChkStep {cs} did not produce new etom" + throwError "{decl_name%} :: ChkStep {cs} did not produce new etom" def lookupChkStepResult! (cs : ChkStep) : ReifM EvalResult := do if let .some (er, _) := (← getChkStepCache).get? cs then return er else - throwError "lookupChkStepEtom! :: ChkStep {cs} did not produce new etom" + throwError "{decl_name%} :: ChkStep {cs} did not produce new etom" def lookupEtomChkStep! (eidx : Nat) : ReifM ChkStep := do if let .some c := (← getEtomChkStep).get? eidx then return c else - throwError "lookupEtomChkStep! :: Unknown etom {eidx}" + throwError "{decl_name%} :: Unknown etom {eidx}" /-- This should only be used at the meta level, i.e. in code that will @@ -289,7 +289,7 @@ def resolveLamBaseTermImport : LamBaseTerm → ReifM LamBaseTerm /-- Models `resolveImport` on the `meta` level -/ def resolveImport : LamTerm → ReifM LamTerm | .atom n => return .atom n -| .etom _ => throwError "resolveImport :: etom should not occur here" +| .etom _ => throwError "{decl_name%} :: etom should not occur here" | .base b => return .base (← resolveLamBaseTermImport b) | .bvar n => return .bvar n | .lam s t => return .lam s (← resolveImport t) @@ -311,7 +311,7 @@ def resolveImport : LamTerm → ReifM LamTerm -/ def mkImportVersion : LamTerm → ReifM LamTerm | .atom n => return (.atom n) -| .etom _ => throwError "mkImportVersion :: etom should not occur here" +| .etom _ => throwError "{decl_name%} :: etom should not occur here" | .base b => match b with | .eq s => return .base (.eqI (← sort2LamILTyIdx s)) @@ -343,9 +343,9 @@ def newChkStep (c : ChkStep) (res? : Option EvalResult) : ReifM (Bool × EvalRes let res := c.eval ltv.lamVarTy ltv.lamILTy ⟨← getRTableTree, ← getMaxEVarSucc, ← getLamEVarTyTree⟩ if let .some res' := res? then if res' != res then - throwError "newChkStep :: Result {res} of ChkStep {c} does not match with expected {res'}" + throwError "{decl_name%} :: Result {res} of ChkStep {c} does not match with expected {res'}" match res with - | .fail => throwError "newChkStep :: Evaluation of ChkStep {c} produces `fail`" + | .fail => throwError "{decl_name%} :: Evaluation of ChkStep {c} produces `fail`" | .addEntry re => -- If `re` is already provable, do nothing if let .some _ ← lookupREntryProof? re then @@ -478,7 +478,7 @@ section ILLifting let (upFunc, downFunc, ty, upTy) ← updownFunc s let sortOrig ← Expr.normalizeType (← Meta.inferType ty) let .sort uOrig := sortOrig - | throwError "mkImportAux :: Unexpected sort {sortOrig} when processing sort {s}" + | throwError "{decl_name%} :: Unexpected sort {sortOrig} when processing sort {s}" return (upFunc, downFunc, ty, upTy, uOrig) set_option pp.universes true @@ -503,20 +503,20 @@ section Checker def nonemptyOfAtom (n : Nat) : ReifM (Option REntry) := do let .some (_, s) := (← getVarVal).get? n - | throwError "nonemptyOfAtom :: Index {n} out of bound" + | throwError "{decl_name%} :: Index {n} out of bound" if !(← inhSubsumptionCheck s) then return .none let (_, .addEntry re) ← newChkStep (.n (.nonemptyOfAtom n)) (.some (.addEntry (.nonempty s))) - | throwError "nonemptyOfAtom :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return .some re def nonemptyOfEtom (n : Nat) : ReifM (Option REntry) := do let .some s := (← getLamEVarTy).get? n - | throwError "nonemptyOfEtom :: Index {n} out of bound" + | throwError "{decl_name%} :: Index {n} out of bound" if !(← inhSubsumptionCheck s) then return .none let (_, .addEntry re) ← newChkStep (.n (.nonemptyOfEtom n)) (.some (.addEntry (.nonempty s))) - | throwError "nonemptyOfEtom :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return .some re def validOfIntros (v : REntry) (idx : Nat) : ReifM REntry := do @@ -524,12 +524,12 @@ section Checker return v let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.l (.validOfIntros p idx)) .none - | throwError "validOfIntros :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfIntroMost (v : REntry) : ReifM REntry := do let .valid _ t := v - | throwError "validOfIntroMost :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" let mut idx := 0 let mut t := t while true do @@ -543,12 +543,12 @@ section Checker def validOfReverts (v : REntry) (idx : Nat) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.l (.validOfReverts p idx)) .none - | throwError "validOfReverts :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfRevertAll (v : REntry) : ReifM REntry := do let .valid lctx _ := v - | throwError "validOfRevertAll :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" if lctx.length == 0 then return v validOfReverts v lctx.length @@ -556,26 +556,26 @@ section Checker def validOfAppend (v : REntry) (ex : Array LamSort) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.l (.validOfAppend p ex.toList)) .none - | throwError "validOfAppend :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfPrepend (v : REntry) (ex : Array LamSort) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.l (.validOfPrepend p ex.toList)) .none - | throwError "validOfPrepend :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfHeadBeta (v : REntry) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.c (.validOfHeadBeta p)) .none - | throwError "validOfHeadBeta :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfHnf (v : REntry) : ReifM REntry := do let mut v := v while true do let .valid _ t := v - | throwError "validOfHnf :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" if !t.isHeadBetaTarget then break v ← validOfHeadBeta v @@ -584,50 +584,50 @@ section Checker def validOfBetaBounded (v : REntry) (bound : Nat) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.c (.validOfBetaBounded p bound)) .none - | throwError "validOfBetaBounded :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfBetaReduce (v : REntry) : ReifM REntry := do let .valid _ t := v - | throwError "validOfBetaReduce :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" validOfBetaBounded v t.betaReduceHackyIdx def validOfExtensionalize (v : REntry) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.c (.validOfExtensionalize p)) .none - | throwError "validOfExtensionalize :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEqSymm (v : REntry) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.c (.validOfEqSymm p)) .none - | throwError "validOfEqSymm :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfMp (vp : REntry) (vrw : REntry) : ReifM REntry := do let pp ← lookupREntryPos! vp let prw ← lookupREntryPos! vrw let (_, .addEntry re) ← newChkStep (.c (.validOfMp pp prw)) .none - | throwError "validOfMp :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfMpAll (vp : REntry) (vrw : REntry) : ReifM REntry := do let pp ← lookupREntryPos! vp let prw ← lookupREntryPos! vrw let (_, .addEntry re) ← newChkStep (.c (.validOfMpAll pp prw)) .none - | throwError "validOfMpAll :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEtaExpand1At (v : REntry) (occ : List Bool) : ReifM REntry := do let pv ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.ca (.validOfEtaExpand1At pv occ)) .none - | throwError "validOfEtaExpand1At :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEtaReduce1At (v : REntry) (occ : List Bool) : ReifM REntry := do let pv ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.ca (.validOfEtaReduce1At pv occ)) .none - | throwError "validOfEtaReduce1At :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEtaExpandNAt (v : REntry) (n : Nat) (occ : List Bool) : ReifM REntry := do @@ -637,7 +637,7 @@ section Checker if n == 1 then return ← validOfEtaExpand1At v occ let (_, .addEntry re) ← newChkStep (.ca (.validOfEtaExpandNAt pv n occ)) .none - | throwError "validOfEtaExpandNAt :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEtaReduceNAt (v : REntry) (n : Nat) (occ : List Bool) : ReifM REntry := do @@ -647,55 +647,55 @@ section Checker if n == 1 then return ← validOfEtaReduce1At v occ let (_, .addEntry re) ← newChkStep (.ca (.validOfEtaReduceNAt pv n occ)) .none - | throwError "validOfEtaReduceNAt :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfEtaExpandAt (v : REntry) (occ : List Bool) : ReifM REntry := do let .valid _ t := v - | throwError "validOfEtaExpandAt :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" let .some (rty, _) := LamTerm.getPosWith occ (.base .prop) t - | throwError "validOfEtaExpandAt :: {occ} is not a valid position of {t}" + | throwError "{decl_name%} :: {occ} is not a valid position of {t}" let n := rty.getArgTys.length validOfEtaExpandNAt v n occ def validOfEtaReduceAt (v : REntry) (occ : List Bool) : ReifM REntry := do let .valid _ t := v - | throwError "validOfEtaReduceAt :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" let .some tocc := LamTerm.getPos occ t - | throwError "validOfEtaReduceAt :: {occ} is not a valid position of {t}" + | throwError "{decl_name%} :: {occ} is not a valid position of {t}" let n := tocc.getLamTys.length validOfEtaReduceNAt v n occ def validOfExtensionalizeEqAt (v : REntry) (occ : List Bool) : ReifM REntry := do let pv ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.ca (.validOfExtensionalizeEqAt pv occ)) .none - | throwError "validOfExtensionalizeEqAt :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfExtensionalizeEqFNAt (v : REntry) (n : Nat) (occ : List Bool) : ReifM REntry := do let pv ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.ca (.validOfExtensionalizeEqFNAt pv n occ)) .none - | throwError "validOfExtensionalizeEqFNAt :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfIntensionalizeEqAt (v : REntry) (occ : List Bool) : ReifM REntry := do let pv ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.ca (.validOfIntensionalizeEqAt pv occ)) .none - | throwError "validOfIntensionalizeEqAt :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfBVarLower (v : REntry) (n : REntry) : ReifM REntry := do let pv ← lookupREntryPos! v let pn ← lookupREntryPos! n let (_, .addEntry re) ← newChkStep (.i (.validOfBVarLower pv pn)) .none - | throwError "validOfBVarLower :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfBVarLowers (v : REntry) (ns : Array REntry) : ReifM REntry := do let pv ← lookupREntryPos! v let pns ← ns.mapM lookupREntryPos! let (_, .addEntry re) ← newChkStep (.i (.validOfBVarLowers pv pns.toList)) .none - | throwError "validOfBVarLowers :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re /-- @@ -710,27 +710,27 @@ section Checker let p₁₂ ← lookupREntryPos! v₁₂ let p₁ ← lookupREntryPos! v₁ let (_, .addEntry re) ← newChkStep (.i (.validOfImp p₁₂ p₁)) .none - | throwError "validOfImp :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfImps (impV : REntry) (hypVs : Array REntry) : ReifM REntry := do let imp ← lookupREntryPos! impV let ps ← hypVs.mapM lookupREntryPos! let (_, .addEntry re) ← newChkStep (.i (.validOfImps imp ps.toList)) .none - | throwError "validOfImps :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re /-- Repeated instantiation -/ def validOfInstantiate (v : REntry) (args : Array LamTerm) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiate p args.toList)) .none - | throwError "validOfInstantiate :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfInstantiateRev (v : REntry) (args : Array LamTerm) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.i (.validOfInstantiateRev p args.toList)) .none - | throwError "validOfInstantiateRev :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re /-- @@ -746,22 +746,22 @@ section Checker def validOfAndLeft (v : REntry) (occ : List Bool) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.i (.validOfAndLeft p occ)) .none - | throwError "validOfAndLeft :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def validOfAndRight (v : REntry) (occ : List Bool) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.i (.validOfAndRight p occ)) .none - | throwError "validOfAndRight :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re /-- Exhaustively decompose `∧` at position `occ` -/ partial def decomposeAnd (v : REntry) (occ : List Bool) : ReifM (Array REntry) := do let .valid _ t := v - | throwError "decomposeAnd :: Unexpected entry" - if !(t.isSign true occ) then throwError "decomposeAnd :: {occ} is not a positive position of {t}" + | throwError "{decl_name%} :: Unexpected entry" + if !(t.isSign true occ) then throwError "{decl_name%} :: {occ} is not a positive position of {t}" let .some sub := t.getPos occ - | throwError "decomposeAnd :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if sub.getAppFn == .base .and && sub.getAppArgs.length == 2 then let left ← validOfAndLeft v occ let right ← validOfAndRight v occ @@ -771,19 +771,19 @@ section Checker def boolFacts : ReifM REntry := do let (_, .addEntry re) ← newChkStep (.f .boolFacts) .none - | throwError "boolFacts :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def iteSpec (s : LamSort) : ReifM REntry := do let (_, .addEntry re) ← newChkStep (.f (.iteSpec s)) .none - | throwError "iteSpec :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re def skolemize (exV : REntry) : ReifM REntry := do let eidx ← getMaxEVarSucc let ex ← lookupREntryPos! exV let (new?, .newEtomWithValid _ lctx t) ← newChkStep (.e (.skolemize ex)) .none - | throwError "skolemize :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" if new? then let _ ← nonemptyOfEtom eidx return .valid lctx t @@ -793,11 +793,11 @@ section Checker let mut exV := exV while true do let .valid _ t := exV - | throwError "skolemizeMostIntoForall :: Unexpected entry {exV}" + | throwError "{decl_name%} :: Unexpected entry {exV}" if t.isMkForallE then exV ← validOfIntroMost exV let .valid _ t := exV - | throwError "skolemizeMostIntoForall :: Unexpected entry {exV}" + | throwError "{decl_name%} :: Unexpected entry {exV}" if t.isMkExistE then exV ← skolemize exV exV ← validOfHnf exV @@ -808,7 +808,7 @@ section Checker def define (t : LamTerm) : ReifM REntry := do let eidx ← getMaxEVarSucc let (new?, .newEtomWithValid _ [] t) ← newChkStep (.e (.define t)) .none - | throwError "define :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" if new? then let _ ← nonemptyOfEtom eidx return .valid [] t @@ -816,7 +816,7 @@ section Checker def validOfPrepConv (pc : PrepConvStep) (v : REntry) (occ : List Bool) : ReifM REntry := do let p ← lookupREntryPos! v let (_, .addEntry re) ← newChkStep (.p pc p occ) .none - | throwError "validOfPrepConv :: Unexpected evaluation result" + | throwError "{decl_name%} :: Unexpected evaluation result" return re end Checker @@ -829,17 +829,17 @@ section CheckerUtils -/ def reorderLCtx (v : REntry) (rmap : Array Nat) : ReifM REntry := do let .valid lctx _ := v - | throwError "reorderLCtx :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" let lsize := lctx.length if rmap.size != lsize then - throwError "reorderLCtx :: Length of lctx does not equal size of reorder map" + throwError "{decl_name%} :: Length of lctx does not equal size of reorder map" let mut ex : Array LamSort := rmap.map (fun _ => .atom 0) let mut argBVarIdx : Array Nat := #[] for (s, i) in (Array.mk lctx).zipWithIndex do let .some i' := rmap[i]? - | throwError "reorderLCtx :: Does not know where does `.bvar i` map to" + | throwError "{decl_name%} :: Does not know where does `.bvar i` map to" if i' >= lsize then - throwError "reorderLCtx :: Index {i'} out of bound {lsize}" + throwError "{decl_name%} :: Index {i'} out of bound {lsize}" ex := ex.set! i' s argBVarIdx := argBVarIdx.push (lsize - i - 1 + i') let v ← validOfAppend v ex @@ -848,7 +848,7 @@ section CheckerUtils /-- Refer to docstring of `LamTerm.isGeneral` -/ def toDefinition? (v : REntry) : ReifM (Option REntry) := do let .valid lctx t := v - | throwError "toDefinition :: Unexpected entry {v}" + | throwError "{decl_name%} :: Unexpected entry {v}" let mut introed := t.betaReduceHacky let mut lctx' := lctx while true do @@ -863,7 +863,7 @@ section CheckerUtils let mr := rhs.getLamBody.isGeneral (lctx'.length + rhs.getLamTys.length) if ml.isNone && mr.isNone then return .none let .some m := ml <|> mr - | throwError "toDefinition? :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let v ← validOfBetaReduce v let v ← validOfIntroMost v let v ← (do if (head == .iff) then validOfPrepConv .validOfPropext v [] else pure v) @@ -897,7 +897,7 @@ section CheckerUtils | passive := passive.push back; continue trace[auto.lamReif.prep.def] "Entry {back} is def-like and is turned into {v'}" let .valid [] rw@(.app _ (.app _ (.base (.eq _)) lhs) rhs) := v' - | throwError "recognizeDefsAndUnfold :: Unexpected definition entry {v'}" + | throwError "{decl_name%} :: Unexpected definition entry {v'}" -- If the left-hand-side occurs inside the right-hand-side, -- then this definition is recursive and we will not unfold it if rhs.abstractsImp #[lhs] != rhs then @@ -906,7 +906,7 @@ section CheckerUtils active ← active.mapM (validOfMpAll · v') minds ← minds.mapM (fun mind => do let .some mind := mind.mpAll? rw - | throwError "recognizeDefsAndUnfold :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return mind) return (passive, minds) @@ -972,7 +972,7 @@ def newTypeExpr (e : Expr) : ReifM LamSort := do let origSort ← Meta.inferType e let origSort := (← instantiateMVars origSort).headBeta let .sort lvl := origSort - | throwError "newTypeExpr :: Unexpected sort {origSort}" + | throwError "{decl_name%} :: Unexpected sort {origSort}" setTyVal (tyVal.push (e, lvl)) setTyVarMap (tyVarMap.insert e idx) return .atom idx @@ -1129,15 +1129,15 @@ def processSimpleLit (l : Literal) : LamTerm := def processSimpleConst (name : Name) (lvls : List Level) : ReifM (Option LamTerm) := do if let .some t := reifMapConstNilLvl.get? name then if lvls.length != 0 then - throwError "processSimpleConst :: ConstNilLvl constants should have nil level list" + throwError "{decl_name%} :: ConstNilLvl constants should have nil level list" return t if name == ``Embedding.ImpF then let [u, v] := lvls - | throwError "processSimpleConst :: Invalid levels {lvls} for Auto.Embedding.ImpF" + | throwError "{decl_name%} :: Invalid levels {lvls} for Auto.Embedding.ImpF" if (← Meta.isLevelDefEq u .zero) ∧ (← Meta.isLevelDefEq v .zero) then return .some (.base .imp) else - throwError "processSimpleConst :: Non-propositional implication is not supported" + throwError "{decl_name%} :: Non-propositional implication is not supported" return .none def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do @@ -1146,22 +1146,22 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do let .const name lvls := fn | return .none match args.toList with - | [] => throwError "processSimpleApp :: Unexpected error" + | [] => throwError "{decl_name%} :: Unexpected error" | [arg] => if let .some tcon := reifMapBVConst1.get? name then if lvls.length != 0 then - throwError "processSimpleApp :: BVConst1 should have nil level list" + throwError "{decl_name%} :: BVConst1 should have nil level list" if let .some n ← @id (MetaM _) (Meta.evalNat arg) then return .some (tcon n) return .none if let .some attrName := reifMapAttributesProp.get? name then if lvls.length != 1 then - throwError "processSimpleApp :: Attribute should have one level" + throwError "{decl_name%} :: Attribute should have one level" return .some (.base (.ocst (.smtAttr1T attrName (← reifType arg) (.base .prop)))) if let .some tcon := reifMapIL.get? name then if name == ``Embedding.forallF then let [lvl₁, lvl₂] := lvls - | throwError "processSimpleApp :: Auto.Embedding.forallF should have two levels" + | throwError "{decl_name%} :: Auto.Embedding.forallF should have two levels" if !(← Meta.isLevelDefEq lvl₂ .zero) then return .none return .some (.base (tcon (← reifType arg))) @@ -1169,7 +1169,7 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do | [arg₁, arg₂] => if let .some tcon := reifMapBVConst2.get? name then if lvls.length != 0 then - throwError "processSimpleApp :: BVConst2 should have nil level list" + throwError "{decl_name%} :: BVConst2 should have nil level list" match ← @id (MetaM _) (Meta.evalNat arg₁), ← @id (MetaM _) (Meta.evalNat arg₂) with | .some n, .some m => return .some (tcon n m) @@ -1178,7 +1178,7 @@ def processSimpleApp (fn arg : Expr) : ReifM (Option LamTerm) := do | [arg₁, arg₂, arg₃] => if let .some tcon := reifMapBVConst3.get? name then if lvls.length != 0 then - throwError "processSimpleApp :: BVConst2 should have nil level list" + throwError "{decl_name%} :: BVConst2 should have nil level list" match ← @id (MetaM _) (Meta.evalNat arg₁), ← @id (MetaM _) (Meta.evalNat arg₂), ← @id (MetaM _) (Meta.evalNat arg₃) with @@ -1322,13 +1322,13 @@ def processLam0Arg2 (e fn arg₁ arg₂ : Expr) : MetaM (Option LamTerm) := do | return .none if arg₁.isConst then let .const arg₁Name _ := arg₁ - | throwError "processLam0Arg2 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if let .some (e', t) := reifMapLam0Arg2NoLit.get? (fnName, arg₁Name) then if (← Meta.isDefEqD e e') then return .some t if arg₁.isApp then let .app arg₁fn arg₁arg := arg₁ - | throwError "processLam0Arg2 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if let .const arg₁FnName _ := arg₁fn then if let .some candidates := reifMapLam0Arg2Natlit.get? (fnName, arg₁FnName) then if let .some n ← Meta.evalNat arg₁arg then @@ -1344,20 +1344,20 @@ def processLam0Arg3 (e fn arg₁ arg₂ arg₃ : Expr) : MetaM (Option LamTerm) | .const ``Nat _ => if (← Meta.isDefEqD e arg₂) then let .lit (.natVal nv) := arg₂ - | throwError "processLam0Arg3 :: OfNat.ofNat instance is not based on a nat literal" + | throwError "{decl_name%} :: OfNat.ofNat instance is not based on a nat literal" return .some (.base (.natVal nv)) return .none | .const ``Int _ => if (← Meta.isDefEqD e (.app (.const ``Int.ofNat []) arg₂)) then let .lit (.natVal nv) := arg₂ - | throwError "processLam0Arg3 :: OfNat.ofNat instance is not based on a nat literal" + | throwError "{decl_name%} :: OfNat.ofNat instance is not based on a nat literal" return .some (.mkIOfNat (.base (.natVal nv))) return .none | .app (.const ``BitVec []) nExpr => if let .some n ← Meta.evalNat nExpr then if (← Meta.isDefEqD e (mkApp2 (.const ``BitVec.ofNat []) (.lit (.natVal n)) arg₂)) then let .lit (.natVal nv) := arg₂ - | throwError "processLam0Arg3 :: OfNat.ofNat instance is not based on a nat literal" + | throwError "{decl_name%} :: OfNat.ofNat instance is not based on a nat literal" return .some (.base (.bvVal n nv)) return .none | _ => return .none @@ -1368,19 +1368,19 @@ def processLam0Arg4 (e fn arg₁ arg₂ arg₃ arg₄ : Expr) : MetaM (Option La | return .none if arg₁.isConst && arg₂.isConst then let .const arg₁name _ := arg₁ - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let .const arg₂name _ := arg₂ - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if let .some (e', t) := reifMapLam0Arg4NoLit.get? (fnName, arg₁name, arg₂name) then if (← Meta.isDefEqD e e') then return .some t return .none if arg₁.isApp && arg₂.isConst then let .app arg₁fn arg₁arg := arg₁ - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if arg₁fn.isConst then let .const arg₁fnName _ := arg₁fn - | throwError "processLam0Arg4 :: Unexpected error {arg₁fn}" + | throwError "{decl_name%} :: Unexpected error {arg₁fn}" if let .some candidates := reifMapLam0Arg4NatLit.get? (fnName, arg₁fnName) then for (e'con, tcon) in candidates do if let .some n ← Meta.evalNat arg₁arg then @@ -1388,12 +1388,12 @@ def processLam0Arg4 (e fn arg₁ arg₂ arg₃ arg₄ : Expr) : MetaM (Option La return tcon n if arg₁.isApp && arg₂.isApp then let .app arg₁fn arg₁arg := arg₁ - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let .app arg₂fn arg₂arg := arg₂ - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if arg₁fn.isConst && arg₂fn.isConst then let .const arg₁fnName _ := arg₁fn - | throwError "processLam0Arg4 :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" if let .some candidates := reifMapLam0Arg4NatLitNatLitEq.get? (fnName, arg₁fnName) then for (e'con, tcon) in candidates do if let .some n ← Meta.evalNat arg₁arg then @@ -1481,7 +1481,7 @@ def reifTermCheckType (e : Expr) : ReifM (LamSort × LamTerm) := do let t ← reifTerm .empty e let ltv ← getLamTyValAtMeta let .some s := t.lamCheck? ltv Embedding.Lam.dfLCtxTy - | throwError "reifTermCheckType :: LamTerm {t} is not type correct" + | throwError "{decl_name%} :: LamTerm {t} is not type correct" return (s, t) /-- Return the positions of the reified and `resolveImport`-ed facts within the `validTable` -/ @@ -1489,7 +1489,7 @@ def reifFacts (facts : Array UMonoFact) : ReifM (Array LamTerm) := facts.mapM (fun ⟨proof, ty, deriv⟩ => do let (s, lamty) ← reifTermCheckType ty if s != .base .prop then - throwError "reifFacts :: Fact {lamty} is not of type `prop`" + throwError "{decl_name%} :: Fact {lamty} is not of type `prop`" trace[auto.lamReif.printResult] "Successfully reified proof of {← Meta.zetaReduce ty} to λterm `{lamty}`" newAssertion proof deriv lamty return lamty) @@ -1578,7 +1578,7 @@ section BuildChecker for re in (← getRTable) do if let .some cs := rst.chkMap.get? re then let .some n := rst.findPos? re - | throwError "buildChkStepsExpr :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" chkSteps := chkSteps.push (cs, n) -- `ChkMap` are run using `foldl`, so we use `BinTree.ofListFoldl` let e := Lean.toExpr (BinTree.ofListFoldl chkSteps.toList) @@ -1650,7 +1650,7 @@ section BuildChecker let itEntry := Lean.mkApp3 (.const ``importTablePSigmaMk [u]) chkValExpr ieExpr e importTable := importTable.insert n itEntry if t.maxLooseBVarSucc != 0 || t.maxEVarSucc != 0 then - throwError "buildImportTableExpr :: Invalid imported fact {t}" + throwError "{decl_name%} :: Invalid imported fact {t}" let veEntry := REntry.valid [] t importedFactsTree := importedFactsTree.insert n veEntry for (s, (inh, _, n)) in (← getInhabitations).toList do @@ -1683,7 +1683,7 @@ section BuildChecker let (itExpr, _) ← buildImportTableExpr cpvFVarExpr let csExpr ← buildChkStepsExpr let .valid lctx t := re - | throwError "buildFullCheckerExprFor :: {re} is not a `valid` entry" + | throwError "{decl_name%} :: {re} is not a `valid` entry" let vExpr := Lean.toExpr (← lookupREntryPos! re) let eqExpr ← Meta.mkAppM ``Eq.refl #[← Meta.mkAppM ``Option.some #[Lean.toExpr (lctx, t)]] let getEntry := Lean.mkApp7 (.const ``Checker.getValidExport_directReduce [u]) @@ -1709,7 +1709,7 @@ section BuildChecker let (itExpr, ifExpr) ← buildImportTableExpr cpvFVarExpr let csExpr ← buildChkStepsExpr let .valid lctx t := re - | throwError "buildFullCheckerExprFor :: {re} is not a `valid` entry" + | throwError "{decl_name%} :: {re} is not a `valid` entry" let vExpr := Lean.toExpr (← lookupREntryPos! re) let hImportExpr ← Meta.mkAppM ``Eq.refl #[ifExpr] let hLvtExpr ← Meta.mkAppM ``Eq.refl #[lvtExpr] @@ -1766,7 +1766,7 @@ section BuildChecker let csExpr ← buildChkStepsExpr let csNativeName ← mkNativeAuxDecl `lam_ssrefl_cs (Lean.mkConst ``ChkSteps) csExpr let .valid lctx t := re - | throwError "buildFullCheckerExprFor :: {re} is not a `valid` entry" + | throwError "{decl_name%} :: {re} is not a `valid` entry" let vExpr := Lean.toExpr (← lookupREntryPos! re) let hImportExpr ← Meta.mkAppM ``Eq.refl #[Lean.mkConst ifNativeName] let hLvtExpr ← Meta.mkAppM ``Eq.refl #[Lean.mkConst lvtNativeName] @@ -1891,7 +1891,7 @@ open Embedding.Lam LamReif trace[auto.buildChecker] "Collecting for etom {e} by ChkStep {cs}" let _ ← processChkStep ref cs let .some n := (← getEtomH2lMap).get? e - | throwError "transEtom :: Cannot find translation of etom {e}" + | throwError "{decl_name%} :: Cannot find translation of etom {e}" return n partial def transLamTerm (ref : State) : LamTerm → TransM LamTerm @@ -1974,11 +1974,11 @@ open Embedding.Lam LamReif match cs with | .nonemptyOfAtom n => do let .atom n' ← transLamTerm ref (.atom n) - | throwError "transChkStep :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return .nonemptyOfAtom n' | .nonemptyOfEtom n => do let .etom n' ← transLamTerm ref (.etom n) - | throwError "transChkStep :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return .nonemptyOfEtom n' | .p cs pos occ => return ChkStep.p cs (← transPos ref pos) occ | .w cs => ChkStep.w <$> @@ -1995,7 +1995,7 @@ open Embedding.Lam LamReif let cs' ← transChkStep ref cs setCsH2lMap ((← getCsH2lMap).insert cs cs') let (true, er) ← newChkStep cs' .none - | throwError "processChkStep :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" trace[auto.buildChecker] "Checkstep {cs} translated to {cs'}, producing {er}" match er with | .newEtomWithValid .. => do @@ -2017,22 +2017,22 @@ open Embedding.Lam LamReif trace[auto.buildChecker] "Collecting for {hre} by ChkStep {cs}" let er ← processChkStep ref cs match er with - | .fail => throwError "collectProofFor :: Unexpected evaluation result" + | .fail => throwError "{decl_name%} :: Unexpected evaluation result" | .addEntry reNew => do let expectedEntry ← transREntry ref hre - if expectedEntry != reNew then throwError "collectProofFor :: Entry mismatch" + if expectedEntry != reNew then throwError "{decl_name%} :: Entry mismatch" | .newEtomWithValid _ lctx t => do let expectedEntry ← transREntry ref hre - if expectedEntry != .valid lctx t then throwError "collectProofFor :: Entry mismatch" + if expectedEntry != .valid lctx t then throwError "{decl_name%} :: Entry mismatch" | .inhabitation e deriv _ => let .nonempty hs := hre - | throwError "collectProofFor :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let s ← transLamSort ref hs newInhabitation e deriv s trace[auto.buildChecker] "Inhabitation fact {hs} translated to {s}" | .assertion e deriv _ => let .valid [] ht := hre - | throwError "collectProofFor :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let t ← transLamTerm ref ht newAssertion e deriv t trace[auto.buildChecker] "Import fact {ht} translated to {t}" diff --git a/Auto/Translation/LamUtils.lean b/Auto/Translation/LamUtils.lean index e71b816..ad8ccd7 100644 --- a/Auto/Translation/LamUtils.lean +++ b/Auto/Translation/LamUtils.lean @@ -68,10 +68,10 @@ namespace LamExportUtils def collectLamBaseTermAtoms (b : LamBaseTerm) : CoreM (Std.HashSet Nat) := do let s? : Option LamSort ← (do match b with - | .eqI _ => throwError ("collectAtoms :: " ++ exportError.ImpPolyLog) - | .forallEI _ => throwError ("collectAtoms :: " ++ exportError.ImpPolyLog) - | .existEI _ => throwError ("collectAtoms :: " ++ exportError.ImpPolyLog) - | .iteI _ => throwError ("collectAtoms :: " ++ exportError.ImpPolyLog) + | .eqI _ => throwError (decl_name% ++ " :: " ++ exportError.ImpPolyLog) + | .forallEI _ => throwError (decl_name% ++ " :: " ++ exportError.ImpPolyLog) + | .existEI _ => throwError (decl_name% ++ " :: " ++ exportError.ImpPolyLog) + | .iteI _ => throwError (decl_name% ++ " :: " ++ exportError.ImpPolyLog) | .eq s => return .some s | .forallE s => return .some s | .existE s => return .some s @@ -95,11 +95,11 @@ namespace LamExportUtils LamTerm → CoreM (Std.HashSet Nat × Std.HashSet Nat × Std.HashSet Nat) | .atom n => do let .some s := lamVarTy[n]? - | throwError "collectAtoms :: Unknown term atom {n}" + | throwError "{decl_name%} :: Unknown term atom {n}" return (collectLamSortAtoms s, Std.HashSet.empty.insert n, Std.HashSet.empty) | .etom n => do let .some s := lamEVarTy[n]? - | throwError "collectAtoms :: Unknown etom {n}" + | throwError "{decl_name%} :: Unknown etom {n}" return (collectLamSortAtoms s, Std.HashSet.empty, Std.HashSet.empty.insert n) | .base b => do return (← collectLamBaseTermAtoms b, Std.HashSet.empty, Std.HashSet.empty) @@ -194,7 +194,7 @@ namespace Lam2D | .or => return .const ``Or [] | .imp => do let .some (.defnInfo impVal) := (← getEnv).find? ``ImpF - | throwError "interpLamBaseTermAsUnlifted :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return impVal.value.instantiateLevelParams impVal.levelParams [.zero, .zero] | .iff => return .const ``Iff [] @@ -306,7 +306,7 @@ namespace Lam2D def interpLamSortAsUnlifted (tyVal : Std.HashMap Nat Expr) : LamSort → CoreM Expr | .atom n => do let .some e := tyVal.get? n - | throwError "interpLamSortAsUnlifted :: Cannot find fvarId assigned to type atom {n}" + | throwError "{decl_name%} :: Cannot find fvarId assigned to type atom {n}" return e | .base b => return Lam2D.interpLamBaseSortAsUnlifted b | .func s₁ s₂ => do @@ -314,18 +314,18 @@ namespace Lam2D def interpOtherConstAsUnlifted (tyVal : Std.HashMap Nat Expr) (oc : OtherConst) : MetaM Expr := do let .some (.defnInfo constIdVal) := (← getEnv).find? ``constId - | throwError "interpOtherConstAsUnlifted :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let constIdExpr := fun params => constIdVal.value.instantiateLevelParams constIdVal.levelParams params match oc with | .smtAttr1T _ sattr sterm => do let tyattr ← interpLamSortAsUnlifted tyVal sattr let sortattr ← Expr.normalizeType (← Meta.inferType tyattr) let Expr.sort lvlattr := sortattr - | throwError "interpOtherConstAsUnlifted :: Unexpected sort {sortattr}" + | throwError "{decl_name%} :: Unexpected sort {sortattr}" let tyterm ← interpLamSortAsUnlifted tyVal sterm let sortterm ← Expr.normalizeType (← Meta.inferType tyterm) let Expr.sort lvlterm := sortterm - | throwError "interpOtherConstAsUnlifted :: Unexpected sort {sortterm}" + | throwError "{decl_name%} :: Unexpected sort {sortterm}" return Lean.mkApp2 (constIdExpr [lvlattr, lvlterm]) tyattr tyterm def interpLamBaseTermAsUnlifted (tyVal : Std.HashMap Nat Expr) : LamBaseTerm → MetaM Expr @@ -336,19 +336,19 @@ namespace Lam2D | .scst sc => Lam2D.interpStringConstAsUnlifted sc | .bvcst bvc => Lam2D.interpBitVecConstAsUnlifted bvc | .ocst oc => interpOtherConstAsUnlifted tyVal oc - | .eqI _ => throwError ("interpLamTermAsUnlifted :: " ++ LamExportUtils.exportError.ImpPolyLog) - | .forallEI _ => throwError ("interpLamTermAsUnlifted :: " ++ LamExportUtils.exportError.ImpPolyLog) - | .existEI _ => throwError ("interpLamTermAsUnlifted :: " ++ LamExportUtils.exportError.ImpPolyLog) - | .iteI _ => throwError ("interpLamTermAsUnlifted :: " ++ LamExportUtils.exportError.ImpPolyLog) + | .eqI _ => throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) + | .forallEI _ => throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) + | .existEI _ => throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) + | .iteI _ => throwError (decl_name% ++ " :: " ++ LamExportUtils.exportError.ImpPolyLog) | .eq s => do return ← Meta.mkAppOptM ``Eq #[← interpLamSortAsUnlifted tyVal s] | .forallE s => do let ty ← interpLamSortAsUnlifted tyVal s let sort ← Expr.normalizeType (← Meta.inferType ty) let Expr.sort lvl := sort - | throwError "interpLamBaseTermAsUnlifted :: Unexpected sort {sort}" + | throwError "{decl_name%} :: Unexpected sort {sort}" let .some (.defnInfo forallVal) := (← getEnv).find? ``forallF - | throwError "interpLamBaseTermAsUnlifted :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let forallFExpr := forallVal.value.instantiateLevelParams forallVal.levelParams [lvl, .zero] return mkAppN forallFExpr #[← interpLamSortAsUnlifted tyVal s] | .existE s => do @@ -369,11 +369,11 @@ namespace Lam2D (lctx : Nat) : LamTerm → MetaM Expr | .atom n => do let .some e := varVal.get? n - | throwError "interpLamTermAsUnlifted :: Cannot find fvarId assigned to term atom {n}" + | throwError "{decl_name%} :: Cannot find fvarId assigned to term atom {n}" return e | .etom n => do let .some efvar := etomVal.get? n - | throwError "interpLamSortAsUnlifted :: Cannot find fvarId assigned to etom {n}" + | throwError "{decl_name%} :: Cannot find fvarId assigned to etom {n}" return efvar | .base b => interpLamBaseTermAsUnlifted tyVal b | .bvar n => return .bvar n diff --git a/Auto/Translation/Monomorphization.lean b/Auto/Translation/Monomorphization.lean index 12b3f27..bc41cdd 100644 --- a/Auto/Translation/Monomorphization.lean +++ b/Auto/Translation/Monomorphization.lean @@ -184,7 +184,7 @@ def ConstInst.equiv (ci₁ ci₂ : ConstInst) : MetaM Bool := do let ⟨head₁, argsInst₁, idx₁⟩ := ci₁ let ⟨head₂, argsInst₂, idx₂⟩ := ci₂ if head₁.fingerPrint != head₂.fingerPrint then - throwError "ConstInst.equiv :: {ci₁.head} and {ci₂.head} have different fingerprints" + throwError "{decl_name%} :: {ci₁.head} and {ci₂.head} have different fingerprints" if !(← head₁.equiv head₂) then return false if argsInst₁.size != argsInst₂.size || idx₁ != idx₂ then @@ -207,7 +207,7 @@ def ConstInst.matchExpr (e : Expr) (ci : ConstInst) : MetaM Bool := do return false let argsIdx := ci.argsIdx if argsIdx.size != ci.argsInst.size then - throwError "ConstInst.matchExpr :: Unexpected error" + throwError "{decl_name%} :: Unexpected error" let args := e.getAppArgs for (idx, ciarg) in argsIdx.zip ci.argsInst do let .some arg := args[idx]? @@ -252,7 +252,7 @@ def ConstInst.ofExpr? (params : Array Name) (bvars : Array Expr) (e : Expr) : Me for (arg, idx) in args.zipWithIndex do headType ← Core.betaReduce headType let .forallE _ ty body bi := headType - | throwError "ConstInst.ofExpr? :: {headType} is not a `∀`" + | throwError "{decl_name%} :: {headType} is not a `∀`" if let some _ := ty.find? (fun e => bvarSet.contains e) then return .none if ← shouldInstantiate fn ty body bi then @@ -308,7 +308,7 @@ def ConstInst.toExpr (ci : ConstInst) : MetaM Expr := do for (arg, idx) in ci.argsInst.zip ci.argsIdx do args := args.setD idx (.some arg) let .some ret := ConstInst.toExprAux args.toList [] ci.head.toExpr type - | throwError "ConstInst.toExpr :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" return ret /-- @@ -350,9 +350,9 @@ private partial def collectConstInsts (params : Array Name) (bvars : Array Expr) return insts else return insts ++ (← collectConstInsts params bvars ty) -| .letE .. => throwError "collectConstInsts :: Let-expressions should have been reduced" -| .mdata .. => throwError "collectConstInsts :: mdata should have been consumed" -| .proj .. => throwError "collectConstInsts :: Projections should have been turned into ordinary expressions" +| .letE .. => throwError "{decl_name%} :: Let-expressions should have been reduced" +| .mdata .. => throwError "{decl_name%} :: mdata should have been consumed" +| .proj .. => throwError "{decl_name%} :: Projections should have been turned into ordinary expressions" | _ => return #[] where processOther (params : Array Name) (e : Expr) : MetaM (Array ConstInst) := do match ← ConstInst.ofExpr? params bvars e with @@ -380,7 +380,7 @@ def ConstInsts.canonicalize? (cis : ConstInsts) (ci : ConstInst) : MetaM (Option This function is used by `LemmaInst.matchConstInst` only -/ private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) : Expr → MetaM (Std.HashSet LemmaInst) -| .bvar _ => throwError "MLemmaInst.matchConstInst :: Loose bound variable" +| .bvar _ => throwError "{decl_name%} :: Loose bound variable" | e@(.app ..) => do let args := e.getAppArgs let mut ret := Std.HashSet.empty @@ -395,7 +395,7 @@ private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) let mut ret ← MLemmaInst.matchConstInst ci mi body for x in xs do let .fvar id := x - | throwError "MLemmaInst.matchConstInst :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let type ← id.getType ret := mergeHashSet ret (← MLemmaInst.matchConstInst ci mi type) return ret @@ -403,9 +403,9 @@ private partial def MLemmaInst.matchConstInst (ci : ConstInst) (mi : MLemmaInst) let tyInst ← MLemmaInst.matchConstInst ci mi ty let bodyInst ← MLemmaInst.matchConstInst ci mi (body.instantiate1 x) return mergeHashSet tyInst bodyInst -| .letE .. => throwError "MLemmaInst.matchConstInst :: Let-expressions should have been reduced" -| .mdata .. => throwError "MLemmaInst.matchConstInst :: mdata should have been consumed" -| .proj .. => throwError "MLemmaInst.matchConstInst :: Projections should have been turned into ordinary expressions" +| .letE .. => throwError "{decl_name%} :: Let-expressions should have been reduced" +| .mdata .. => throwError "{decl_name%} :: mdata should have been consumed" +| .proj .. => throwError "{decl_name%} :: Projections should have been turned into ordinary expressions" | _ => return Std.HashSet.empty /-- Given a LemmaInst `li` and a ConstInst `ci`, try to match all subexpressions of `li` against `ci` -/ @@ -426,7 +426,7 @@ where match e with | .forallE name ty body bi => Meta.withLocalDecl name bi ty fun x => do let Expr.fvar xid := x - | throwError "Monomorphization.leadingForallQuasiMonomorphic :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let bodyi := body.instantiate1 x if ← Meta.isProp ty then if !(← Meta.isProp bodyi) then @@ -472,7 +472,7 @@ def LemmaInst.monomorphic? (li : LemmaInst) : MetaM (Option LemmaInst) := do | return .none match mvar with | .mvar id => id.assign inst - | _ => throwError "LemmaInst.monomorphic? :: Unexpected error" + | _ => throwError "{decl_name%} :: Unexpected error" LemmaInst.ofMLemmaInst mi /- @@ -571,9 +571,9 @@ def dequeueActiveCi? : MonoM (Option (Expr × Nat)) := do def lookupActiveCi! (fgp : Expr) (idx : Nat) : MonoM ConstInst := do let .some cis := (← getCiMap).get? fgp - | throwError "lookupActiveCi :: Unknown CiHead {fgp}" + | throwError "{decl_name%} :: Unknown CiHead {fgp}" let .some ci := cis[idx]? - | throwError "lookupActiveCi :: Index {idx} out of bound" + | throwError "{decl_name%} :: Index {idx} out of bound" return ci def saturationThresholdReached? (cnt : Nat) : CoreM Bool := do @@ -718,7 +718,7 @@ namespace FVarRep let ci ← MetaState.runMetaM (do match ← CiMap.canonicalize? ciMap ci with | (true, _, ci') => return ci' - | _ => throwError "constInst2FVarId :: Cannot find canonicalized instance of {ci}") + | _ => throwError "{decl_name%} :: Cannot find canonicalized instance of {ci}") let ciIdMap ← FVarRep.getCiIdMap match ciIdMap.get? ci with | .some fid => return fid @@ -778,7 +778,7 @@ namespace FVarRep -- Type of λ binder cannot depend on previous bound variables let (ty, hasBfvars) ← processType ty if hasBfvars then - return .inr m!"replacePolyWithFVar :: Type {ty} of λ binder contains bound variables" + return .inr m!"{decl_name%} :: Type {ty} of λ binder contains bound variables" let fvarId ← MetaState.withLocalDecl name binfo ty .default setBfvars ((← getBfvars).push fvarId) let b' ← replacePolyWithFVar (body.instantiate1 (.fvar fvarId)) @@ -789,14 +789,14 @@ namespace FVarRep | e@(.forallE name ty body binfo) => do let tysort ← MetaState.runMetaM (do Expr.normalizeType (← Meta.inferType ty)) let .sort tylvl := tysort - | throwError "replacePolyWithFVar :: Unexpected error, {tysort} is not a sort" + | throwError "{decl_name%} :: Unexpected error, {tysort} is not a sort" let (ty, tyHasBfvars) ← processType ty let fvarId ← MetaState.withLocalDecl name binfo ty .default setBfvars ((← getBfvars).push fvarId) let body' := body.instantiate1 (.fvar fvarId) let bodysort ← MetaState.runMetaM <| do Expr.normalizeType (← Meta.inferType body') let .sort bodylvl := bodysort - | throwError "replacePolyWithFVar :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" let bodyrep ← replacePolyWithFVar body' let .inl bodyrep := bodyrep | return bodyrep @@ -805,9 +805,9 @@ namespace FVarRep -- of this `∀` does not occur in the body if ← MetaState.isLevelDefEqRigid tylvl .zero then if !(← MetaState.isLevelDefEqRigid bodylvl .zero) then - return .inr m!"replacePolyWithFVar :: In {e}, type of ∀ bound variable is of sort `Prop`, but body isn't of sort `Prop`" + return .inr m!"{decl_name%} :: In {e}, type of ∀ bound variable is of sort `Prop`, but body isn't of sort `Prop`" if body.hasLooseBVar 0 then - return .inr m!"replacePolyWithFVar :: In {e}, type of dependent ∀ bound variable is of sort `Prop`" + return .inr m!"{decl_name%} :: In {e}, type of dependent ∀ bound variable is of sort `Prop`" let impFun := Expr.const ``ImpF [.zero, .zero] addForallImpFInst impFun let tyrep ← replacePolyWithFVar ty @@ -819,7 +819,7 @@ namespace FVarRep -- bound variables else if tyHasBfvars then - return .inr m!"replacePolyWithFVar :: In {e}, type of ∀ bound variable is not of sort `Prop`, and depends on bound variables" + return .inr m!"{decl_name%} :: In {e}, type of ∀ bound variable is not of sort `Prop`, and depends on bound variables" let forallFun := Expr.app (.const ``forallF [tylvl, bodylvl]) ty addForallImpFInst forallFun let forallFunId ← replacePolyWithFVar forallFun diff --git a/Auto/Translation/Preprocessing.lean b/Auto/Translation/Preprocessing.lean index b2430e9..dc2d127 100644 --- a/Auto/Translation/Preprocessing.lean +++ b/Auto/Translation/Preprocessing.lean @@ -28,7 +28,7 @@ def elabLemma (stx : Term) (deriv : DTr) : TacticM Lemma := def addRecAsLemma (recVal : RecursorVal) : MetaM (Array Lemma) := do let some (.inductInfo indVal) := (← getEnv).find? recVal.getInduct - | throwError "Expected inductive datatype: {recVal.getInduct}" + | throwError "{decl_name%} :: Expected inductive datatype: {recVal.getInduct}" let expr := mkConst recVal.name (recVal.levelParams.map Level.param) let res ← forallBoundedTelescope (← inferType expr) recVal.getMajorIdx fun xs _ => do let expr := mkAppN expr xs @@ -39,7 +39,7 @@ def addRecAsLemma (recVal : RecursorVal) : MetaM (Array Lemma) := do let ctor := mkAppN ctor ys let expr := mkApp expr ctor let some redExpr ← reduceRecMatcher? expr - | throwError "Could not reduce recursor application: {expr}" + | throwError "{decl_name%} :: Could not reduce recursor application: {expr}" let redExpr ← Core.betaReduce redExpr let eq ← mkEq expr redExpr let proof ← mkEqRefl expr @@ -50,7 +50,7 @@ def addRecAsLemma (recVal : RecursorVal) : MetaM (Array Lemma) := do for lem in res do let ty' ← Meta.inferType lem.proof if !(← Meta.isDefEq ty' lem.type) then - throwError "addRecAsLemma :: Application type mismatch" + throwError "{decl_name%} :: Application type mismatch" return Array.mk res def elabDefEq (name : Name) : TacticM (Array Lemma) := do @@ -69,10 +69,10 @@ def elabDefEq (name : Name) : TacticM (Array Lemma) := do -- If we have inductively defined propositions, we might -- need to add constructors as lemmas | some (.ctorInfo _) => return #[] - | some (.opaqueInfo _) => throwError "Opaque constants cannot be provided as lemmas" - | some (.quotInfo _) => throwError "Quotient constants cannot be provided as lemmas" - | some (.inductInfo _) => throwError "Inductive types cannot be provided as lemmas" - | none => throwError "Unknown constant {name}" + | some (.opaqueInfo _) => throwError "{decl_name%} :: Opaque constants cannot be provided as lemmas" + | some (.quotInfo _) => throwError "{decl_name%} :: Quotient constants cannot be provided as lemmas" + | some (.inductInfo _) => throwError "{decl_name%} :: Inductive types cannot be provided as lemmas" + | none => throwError "{decl_name%} :: Unknown constant {name}" def isNonemptyInhabited (ty : Expr) : MetaM Bool := do let .some name ← Meta.isClass? ty @@ -86,9 +86,9 @@ structure ConstUnfoldInfo where def getConstUnfoldInfo (name : Name) : MetaM ConstUnfoldInfo := do let .some ci := (← getEnv).find? name - | throwError "getConstUnfoldInfo :: Unknown declaration {name}" + | throwError "{decl_name%} :: Unknown declaration {name}" let .some val := ci.value? - | throwError "getConstUnfoldInfo :: {name} is not a definition, thus cannot be unfolded" + | throwError "{decl_name%} :: {name} is not a definition, thus cannot be unfolded" let val ← prepReduceExpr val let params := ci.levelParams return ⟨name, val, ⟨params⟩⟩ @@ -110,13 +110,13 @@ partial def topoSortUnfolds (unfolds : Array ConstUnfoldInfo) : MetaM (Array Con let mut ret := #[] for name in nameArr do let .some ui := nameMap.get? name - | throwError "topoSortUnfolds :: Unexpected error" + | throwError "{decl_name%} :: Unexpected error" ret := ret.push ui return ret.reverse where go (depMap : Std.HashMap Name (Std.HashSet Name)) (stack : Std.HashSet Name) (n : Name) : StateRefT (Std.HashSet Name × Array Name) MetaM Unit := do if stack.contains n then - throwError "topoSortUnfolds :: Cyclic dependency" + throwError "{decl_name%} :: Cyclic dependency" let (done, ret) ← get if done.contains n then return diff --git a/Auto/Translation/Reduction.lean b/Auto/Translation/Reduction.lean index 7bc758d..a5e53b5 100644 --- a/Auto/Translation/Reduction.lean +++ b/Auto/Translation/Reduction.lean @@ -29,14 +29,14 @@ def unfoldProj (e : Expr) : MetaM Expr := match e with | .proj name idx struct => do let some (.inductInfo indi) := (← getEnv).find? name - | throwError s!"unfoldProj :: {name} is not a inductive type" + | throwError s!"{decl_name%} :: {name} is not a inductive type" let some structInfo := getStructureInfo? (← getEnv) name - | throwError s!"unfoldProj :: {name} is not a structure" + | throwError s!"{decl_name%} :: {name} is not a structure" let nameMap : Std.HashMap Name StructureFieldInfo := Std.HashMap.ofList (structInfo.fieldInfo.map (fun sfi => (sfi.fieldName, sfi))).toList let sorted := structInfo.fieldNames.map (fun name => nameMap.get? name) let .some (.some sfi) := sorted[idx]? - | throwError s!"unfoldProj :: Projection index out of bound" + | throwError s!"{decl_name%} :: Projection index out of bound" let nones : List (Option Expr) := (List.range indi.numParams).map (fun _ => .none) Meta.mkAppOptM sfi.projFn ((Array.mk nones).push struct) | _ => return e diff --git a/Auto/Translation/ReifM.lean b/Auto/Translation/ReifM.lean index e5256ce..c4a414c 100644 --- a/Auto/Translation/ReifM.lean +++ b/Auto/Translation/ReifM.lean @@ -42,11 +42,11 @@ abbrev ReifM := StateT State MetaM @[inline] def resolveTy (e : Expr) : ReifM Expr := do match (← getTyCanMap).get? (Expr.eraseMData e) with | .some id => return id - | .none => throwError "resolveTy :: Unable to resolve {e}" + | .none => throwError "{decl_name%} :: Unable to resolve {e}" def mkAuxName (suffix : Name) : ReifM Name := do match (← getDeclName?) with - | none => throwError "ReifM.mkAuxName :: auxiliary declaration cannot be created when declaration name is not available" + | none => throwError "{decl_name%} :: auxiliary declaration cannot be created when declaration name is not available" | some declName => Lean.mkAuxName (declName ++ suffix) 1 end Auto.Reif diff --git a/Doc/LamPULift.lean b/Doc/LamPULift.lean index d3790d5..2c0524c 100644 --- a/Doc/LamPULift.lean +++ b/Doc/LamPULift.lean @@ -30,7 +30,7 @@ initialize Note that `types/terms depending on types` are not fully supported For example, if we have const/fvar `f : ∀ (α : Type), α → Prop` and `h : Nat → Nat`, then calling `termLift` on `f (Nat → Nat) h` - would fail. This is because + would fail. This is because (1) `h` will be lifted to `hLift : GLift Nat → GLift Nat`, so The lifted version of `f (Nat → Nat)` must have type `(GLift Nat → GLift Nat) → GLift Prop` @@ -78,11 +78,11 @@ mutual -- In the return type, the first `Expr` is `e↑`, and the second `Expr` is -- the type of `e↑` partial def cstULiftPos (u : Level) (e : Expr) : (ty : Expr) → MetaM (Expr × Expr) - | .bvar idx => throwError "Auto.cstULift :: Loose bound variable" - | .lam .. => throwError "Auto.cstULift :: Please β-reduce type before calling cstULift" - | .letE .. => throwError "Auto.cstULift :: Please ζ-reduce type before calling cstULift" - | .lit .. => throwError "Auto.cstULift :: Malformed type" - | .proj .. => throwError "Auto.cstULift :: Please unfold projections in type before calling cstULift" + | .bvar idx => throwError "{decl_name%} :: Loose bound variable" + | .lam .. => throwError "{decl_name%} :: Please β-reduce type before calling cstULift" + | .letE .. => throwError "{decl_name%} :: Please ζ-reduce type before calling cstULift" + | .lit .. => throwError "{decl_name%} :: Malformed type" + | .proj .. => throwError "{decl_name%} :: Please unfold projections in type before calling cstULift" | .forallE name biTy body binfo => do -- We want to reate a free variable (intended) of type `bity↑`. -- However, we still don't what's `bity↑`, so we first @@ -120,7 +120,7 @@ mutual -- The same holds for `cstULiftNeg` let sort ← instantiateMVars (← Meta.inferType ty) if !sort.isSort then - throwError "Auto.ULift :: Expected sort" + throwError "{decl_name%} :: Expected sort" let lvl := sort.sortLevel! let eUpTy := mkApp (.const ``GLift [lvl, u]) ty let eUp := mkAppN (.const ``GLift.up [lvl, u]) #[ty, e] @@ -128,11 +128,11 @@ mutual /-- In the return type, the first `Expr` is `eUp↓`, and the second `Expr` is the type of `e↑` -/ partial def cstULiftNeg (u : Level) (eUp : Expr) : (ty : Expr) → MetaM (Expr × Expr) - | .bvar idx => throwError "Auto.cstULift :: Loose bound variable" - | .lam .. => throwError "Auto.cstULift :: Please β-reduce type before calling cstULift" - | .letE .. => throwError "Auto.cstULift :: Please ζ-reduce type before calling cstULift" - | .lit .. => throwError "Auto.cstULift :: Malformed type" - | .proj .. => throwError "Auto.cstULift :: Please unfold projections in type before calling cstULift" + | .bvar idx => throwError "{decl_name%} :: Loose bound variable" + | .lam .. => throwError "{decl_name%} :: Please β-reduce type before calling cstULift" + | .letE .. => throwError "{decl_name%} :: Please ζ-reduce type before calling cstULift" + | .lit .. => throwError "{decl_name%} :: Malformed type" + | .proj .. => throwError "{decl_name%} :: Please unfold projections in type before calling cstULift" | .mdata data ty' => cstULiftNeg u eUp ty' | .forallE name biTy body binfo => do -- `upFunc` is such that `upFunc binder` is equivalent to `binder↑` @@ -163,7 +163,7 @@ mutual | ty => do let sort ← instantiateMVars (← Meta.inferType ty) if !sort.isSort then - throwError "Auto.ULift :: Expected sort" + throwError "{decl_name%} :: Expected sort" let lvl := sort.sortLevel! let eUpTy := mkApp (.const ``GLift [lvl, u]) ty let eUpDown := mkAppN (.const ``GLift.down [lvl, u]) #[ty, eUp] @@ -191,10 +191,9 @@ section TestcstULift let (eup, eupTy) ← cstULiftPos (.param `tmp) e ety let eup ← postProcessULift eup logInfo eup - + set_option pp.universes true set_option pp.funBinderTypes true - set_option pp.structureProjections false private def f₁ := fun (x : Nat) => x #getExprAndApply [f₁ | ulift] @@ -270,7 +269,7 @@ noncomputable def A_Constant.Lift.{u} := fun /- Now, we want to implement a function `termULift` which lifts an expressions - `e` to `GLift.up e`, such that all the constants occurring in `e↑` are [lifted + `e` to `GLift.up e`, such that all the constants occurring in `e↑` are [lifted counterparts of constants in `e`]. To do this, the function requires that all the constants in `e` has already had their lifted counterparts calculated. Before we implement this function, let's first look at what we'll obtain @@ -325,4 +324,4 @@ noncomputable def example₁.Lift.{u} := fun (*) Introduce a let binder `let fvarp↑ : ty↑ := body↑` into the local context -/ -end Auto.LamPULift \ No newline at end of file +end Auto.LamPULift From c89aa2f3af9bfa11f7347e77c8695598ba185c26 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 27 Nov 2024 00:58:23 -0800 Subject: [PATCH 12/25] Test code --- Auto.lean | 3 +- Auto/EvaluateAuto/ConstAnalysis.lean | 127 +++++++++++++++++++++++++++ Auto/EvaluateAuto/EnvAnalysis.lean | 11 +++ Auto/EvaluateAuto/TestCode.lean | 72 +++++++++++++++ Auto/Tactic.lean | 15 ++-- 5 files changed, 221 insertions(+), 7 deletions(-) create mode 100644 Auto/EvaluateAuto/ConstAnalysis.lean create mode 100644 Auto/EvaluateAuto/EnvAnalysis.lean create mode 100644 Auto/EvaluateAuto/TestCode.lean diff --git a/Auto.lean b/Auto.lean index 53b6336..0fd7d64 100644 --- a/Auto.lean +++ b/Auto.lean @@ -1,3 +1,4 @@ import Auto.Tactic +import Auto.EvaluateAuto.TestCode -def hello := "world" \ No newline at end of file +def hello := "world" diff --git a/Auto/EvaluateAuto/ConstAnalysis.lean b/Auto/EvaluateAuto/ConstAnalysis.lean new file mode 100644 index 0000000..859f81b --- /dev/null +++ b/Auto/EvaluateAuto/ConstAnalysis.lean @@ -0,0 +1,127 @@ +import Lean + +open Lean + +namespace Auto + +def Name.getConstsOfModule (module : Name) : CoreM (Array Name) := do + let mFile ← findOLean module + unless (← mFile.pathExists) do + throwError s!"object file '{mFile}' of module {module} does not exist" + let (mod, _) ← readModuleData mFile + return mod.constNames + +/-- Used as a wrapper in other functions -/ +def Name.getCi (name : Name) (parentFunc : Name) : CoreM ConstantInfo := do + let .some ci := (← getEnv).find? name + | throwError "{parentFunc} :: Cannot find name {name}" + return ci + +/-- Used as a wrapper in other functions -/ +def Name.hasValue (name : Name) (parentFunc : Name) : CoreM Bool := do + return (← Name.getCi name parentFunc).value?.isSome + +/-- Used as a wrapper in other functions -/ +def Name.getValue (name : Name) (parentFunc : Name) : CoreM Expr := do + let .some v := (← Name.getCi name parentFunc).value? + | throwError "{parentFunc} :: {name} has no value" + return v + +def Name.isTheorem (name : Name) : CoreM Bool := do + let .some ci := (← getEnv).find? name + | throwError "Name.isTheorem :: Cannot find name {name}" + let .thmInfo _ := ci + | return false + return true + +/-- + A constant is a human theorem iff it is a theorem and has a + declaration range. Roughly speaking, a constant have a declaration + range iff it is defined (presumably by a human) in a `.lean` file +-/ +def Name.isHumanTheorem (name : Name) : CoreM Bool := + return (← Lean.findDeclarationRanges? name).isSome && (← Name.isTheorem name) + +def allHumanTheorems : CoreM (Array Name) := do + let allConsts := (← getEnv).constants.toList.map Prod.fst + let allHumanTheorems ← allConsts.filterM Name.isHumanTheorem + return Array.mk allHumanTheorems + +/-- Return the theorems that occurs in an expression -/ +def Expr.getUsedTheorems (e : Expr) : CoreM (Array Name) := + e.getUsedConstants.filterM Name.isTheorem + +/-- Return the theorems that are used in the declaration body of a constant -/ +def Name.getUsedTheorems (name : Name) : CoreM (Array Name) := do + let v ← Name.getValue name decl_name% + Expr.getUsedTheorems v + +/-- Return true if the expression `e` only uses constants present in `names` -/ +def Expr.onlyUsesConsts (e : Expr) (names : Array Name) : Bool := + e.getUsedConstants.all (fun name => names.contains name) + +/-- Return true if the declaration body of `name` only + uses constants present in `names` -/ +def Name.onlyUsesConsts (name : Name) (names : Array Name) : CoreM Bool := do + let v ← Name.getValue name decl_name% + return Expr.onlyUsesConsts v names + +/-- Return true if the type `name` only uses constants present in `names` -/ +def Name.onlyUsesConstsInType (name : Name) (names : Array Name) : CoreM Bool := do + let ci ← Name.getCi name decl_name% + return Expr.onlyUsesConsts ci.type names + +/-- Used to filter out theorems known to SMT solvers and native provers-/ +def logicConsts : Array Name := #[ + ``True, ``False, + ``Not, ``And, ``Or, ``Iff, + ``Eq + ] + +/-- Used to filter out theorems known to SMT solvers -/ +def boolConsts : Array Name := #[ + ``Bool, + ``true, ``false, + ``Bool.and, ``Bool.or, ``Bool.xor, ``Bool.not, + ``BEq, ``BEq.beq, ``bne, ``instBEqOfDecidableEq, ``instDecidableEqBool, + ``ite, ``cond, + ``Decidable, ``Decidable.decide + ] + +/-- Used to filter out theorems known to SMT solvers -/ +def natConsts : Array Name := #[ + ``Nat, + ``OfNat.ofNat, ``instOfNatNat, + ``Nat.add, ``Nat.sub, ``Nat.mul, ``Nat.div, ``Nat.mod, + ``HAdd, ``HAdd.hAdd, ``instHAdd, ``instAddNat, + ``HSub, ``HSub.hSub, ``instHSub, ``instSubNat, + ``HMul, ``HMul.hMul, ``instHMul, ``instMulNat, + ``HDiv, ``HDiv.hDiv, ``instHDiv, ``Nat.instDiv, + ``HMod, ``HMod.hMod, ``instHMod, ``Nat.instMod, + ``LE, ``LE.le, ``instLENat, + ``LT, ``LT.lt, ``instLTNat, + ``GE.ge, ``GT.gt + ] + +/- **TODO:** Int theorems -/ + +def Name.onlyLogicInType (name : Name) := + Name.onlyUsesConstsInType name logicConsts + +def Name.onlyBoolLogicInType (name : Name) := + Name.onlyUsesConstsInType name (logicConsts ++ boolConsts) + +def Name.onlyNatBoolLogicInType (name : Name) := + Name.onlyUsesConstsInType name (logicConsts ++ boolConsts ++ natConsts) + +/-- Obtain Logic, Bool and Nat theorems -/ +def analyze : CoreM (Array (Array Name)) := do + let a ← allHumanTheorems + let logicThms ← a.filterM Name.onlyLogicInType + let boolThms ← a.filterM (fun name => + return (!(← Name.onlyLogicInType name)) && (← Name.onlyBoolLogicInType name)) + let natThms ← a.filterM (fun name => + return (!(← Name.onlyBoolLogicInType name)) && (← Name.onlyNatBoolLogicInType name)) + return #[logicThms, boolThms, natThms] + +end Auto diff --git a/Auto/EvaluateAuto/EnvAnalysis.lean b/Auto/EvaluateAuto/EnvAnalysis.lean new file mode 100644 index 0000000..e2bfcac --- /dev/null +++ b/Auto/EvaluateAuto/EnvAnalysis.lean @@ -0,0 +1,11 @@ +import Lean + +open Lean + +namespace Auto + +def mathlibModules : CoreM (Array Name) := do + let u := (← getEnv).header.moduleNames + return u.filter (fun name => name.components[0]? == .some `Mathlib) + +end Auto diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean new file mode 100644 index 0000000..2247708 --- /dev/null +++ b/Auto/EvaluateAuto/TestCode.lean @@ -0,0 +1,72 @@ +import Lean +import Auto.EvaluateAuto.ConstAnalysis +import Auto.EvaluateAuto.EnvAnalysis +import Auto.Tactic + +open Lean + +namespace Auto + +inductive Result + | success + | nonProp + | typeCheckFail + | typeUnequal + | autoException (e : Exception) + +instance : ToMessageData Result where + toMessageData : Result → MessageData + | .success => "Result.success" + | .nonProp => "Result.nonProp" + | .typeCheckFail => "Result.typeCheckFail" + | .typeUnequal => "Result.typeUnequal" + | .autoException e => m!"Result.autoException :: {e.toMessageData}" + +/-- + Run `Lean-auto` on `lem.type`, using premises collected from `lem.proof` + Premises which only contain logic constants are filtered because they + are assumed to be known by the prover +-/ +def runAutoOnAutoLemma (declName? : Option Name) (lem : Auto.Lemma) : MetaM Result := do + if !(← Meta.isProp lem.type) then + return .nonProp + -- **TODO: Aux theorem like those ending in `.proof_1`** + let usedThmNames ← (← Expr.getUsedTheorems lem.proof).filterM (fun name => + return !(← Name.onlyLogicInType name)) + let usedThms ← usedThmNames.mapM (fun n => Lemma.ofConst n (.leaf "collected by hammertest")) + let autoProofFn : MetaM Expr := Meta.forallTelescope lem.type fun bs body => do + let negGoal := Expr.app (.const ``Not []) body + let negGoalImpFalse ← Meta.withLocalDeclD `negGoal negGoal fun negGoalFVar => do + let inhLemmas ← Auto.Inhabitation.getInhFactsFromLCtx + let lctxLemmas ← Auto.collectLctxLemmas true #[] + let proofOfFalse ← Auto.runAuto declName? (lctxLemmas ++ usedThms) inhLemmas + Meta.mkLambdaFVars #[negGoalFVar] proofOfFalse + let goal := mkApp2 (.const ``Classical.byContradiction []) body negGoalImpFalse + Meta.mkLambdaFVars bs goal + let mut autoProof : Expr := Expr.sort .zero + try + autoProof ← autoProofFn + catch e => + return .autoException e + match Kernel.check (← getEnv) {} autoProof with + | Except.ok autoProofType => + match Kernel.isDefEq (← getEnv) {} autoProofType lem.type with + | Except.ok true => return .success + | _ => return .typeUnequal + | Except.error _ => return .typeCheckFail + +/-- + Run `Lean-auto` on the type of ``name``, using premises collected + from the proof of `name` + Premises which only contain logic constants are filtered because they + are assumed to be known by the prover +-/ +def runAutoOnConst (name : Name) : MetaM Result := do + let ci ← Name.getCi name decl_name% + let .some v := ci.value? + | throwError "{decl_name%} :: {name} has no value" + let lemmaPre ← Auto.Lemma.ofConst name (.leaf "ofConst") + let lemmaV := {lemmaPre with proof := v} + runAutoOnAutoLemma (.some name) lemmaV + +end Auto diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index ddee40c..1afa718 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -10,6 +10,7 @@ initialize registerTraceClass `auto.tactic registerTraceClass `auto.tactic.printProof registerTraceClass `auto.printLemmas + registerTraceClass `auto.runAuto.printLemmas namespace Auto @@ -203,13 +204,14 @@ def unfoldConstAndprepReduceDefeq (unfolds : Array Prep.ConstUnfoldInfo) (lem : let lem := {lem with type := type} return lem -def traceLemmas (pre : String) (lemmas : Array Lemma) : TacticM Unit := do +def traceLemmas (traceClass : Name) (pre : String) (lemmas : Array Lemma) : CoreM Unit := do let mut cnt : Nat := 0 let mut mdatas : Array MessageData := #[] for lem in lemmas do mdatas := mdatas.push m!"\n{cnt}: {lem}" cnt := cnt + 1 - trace[auto.printLemmas] mdatas.foldl MessageData.compose pre + if ← isTracingEnabledFor traceClass then + addTrace traceClass (mdatas.foldl MessageData.compose pre) def checkDuplicatedFact (terms : Array Term) : TacticM Unit := let n := terms.size @@ -236,19 +238,19 @@ def collectAllLemmas let startTime ← IO.monoMsNow let lctxLemmas ← collectLctxLemmas inputHints.lctxhyps ngoalAndBinders let lctxLemmas ← lctxLemmas.mapM (m:=MetaM) (unfoldConstAndPreprocessLemma unfoldInfos) - traceLemmas "Lemmas collected from local context:" lctxLemmas + traceLemmas `auto.printLemmas "Lemmas collected from local context:" lctxLemmas checkDuplicatedFact inputHints.terms checkDuplicatedLemmaDB inputHints.lemdbs let userLemmas := (← collectUserLemmas inputHints.terms) ++ (← collectHintDBLemmas inputHints.lemdbs) let userLemmas ← userLemmas.mapM (m:=MetaM) (unfoldConstAndPreprocessLemma unfoldInfos) - traceLemmas "Lemmas collected from user-provided terms:" userLemmas + traceLemmas `auto.printLemmas "Lemmas collected from user-provided terms:" userLemmas let defeqLemmas ← collectDefeqLemmas defeqNames let defeqLemmas ← defeqLemmas.mapM (m:=MetaM) (unfoldConstAndprepReduceDefeq unfoldInfos) - traceLemmas "Lemmas collected from user-provided defeq hints:" defeqLemmas + traceLemmas `auto.printLemmas "Lemmas collected from user-provided defeq hints:" defeqLemmas trace[auto.tactic] "Preprocessing took {(← IO.monoMsNow) - startTime}ms" let inhFacts ← Inhabitation.getInhFactsFromLCtx let inhFacts ← inhFacts.mapM (m:=MetaM) (unfoldConstAndPreprocessLemma unfoldInfos) - traceLemmas "Inhabitation lemmas :" inhFacts + traceLemmas `auto.printLemmas "Inhabitation lemmas :" inhFacts return (lctxLemmas ++ userLemmas ++ defeqLemmas, inhFacts) open Embedding.Lam in @@ -431,6 +433,7 @@ def queryNative -/ def runAuto (declName? : Option Name) (lemmas : Array Lemma) (inhFacts : Array Lemma) : MetaM Expr := do + traceLemmas `auto.runAuto.printLemmas s!"All lemmas received by {decl_name%}:" lemmas -- Simplify `ite` let ite_simp_lem ← Lemma.ofConst ``Auto.Bool.ite_simp (.leaf "hw Auto.Bool.ite_simp") let lemmas ← lemmas.mapM (fun lem => Lemma.rewriteUPolyRigid lem ite_simp_lem) From 769fc6cc9802e1c6d997d4ff17610c5bae04eeb4 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 27 Nov 2024 01:00:19 -0800 Subject: [PATCH 13/25] add evaluation bug --- Test/Bugs.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Test/Bugs.lean b/Test/Bugs.lean index c0f5896..a667780 100644 --- a/Test/Bugs.lean +++ b/Test/Bugs.lean @@ -1,4 +1,5 @@ import Auto.Tactic +import Auto.EvaluateAuto.TestCode -- Standard Preprocessing Configs set_option auto.redMode "reducible" @@ -79,3 +80,9 @@ end Set example (x : Nat) (primeset : Nat → Prop) (dvd : Nat → Nat → Prop) : ((∃ (i : _) (i_1 : primeset i), dvd i x) ↔ (∃ p, primeset p ∧ dvd p x)) := by auto + +-- bug +set_option trace.auto.tactic true in +#eval do + let m ← Auto.runAutoOnConst ``Nat.sub_one_lt_of_lt + trace[auto.tactic] m!"{m}" From 0867a4dfafdd39370faddb10182f456990c30528 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 27 Nov 2024 11:48:18 -0800 Subject: [PATCH 14/25] fix preprocessing bug in test code --- Auto/EvaluateAuto/TestCode.lean | 3 ++- Test/Bugs.lean | 6 ------ 2 files changed, 2 insertions(+), 7 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index 2247708..d176459 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -39,7 +39,8 @@ def runAutoOnAutoLemma (declName? : Option Name) (lem : Auto.Lemma) : MetaM Resu let negGoalImpFalse ← Meta.withLocalDeclD `negGoal negGoal fun negGoalFVar => do let inhLemmas ← Auto.Inhabitation.getInhFactsFromLCtx let lctxLemmas ← Auto.collectLctxLemmas true #[] - let proofOfFalse ← Auto.runAuto declName? (lctxLemmas ++ usedThms) inhLemmas + let allLemmas ← (lctxLemmas ++ usedThms).mapM (Auto.unfoldConstAndPreprocessLemma #[]) + let proofOfFalse ← Auto.runAuto declName? allLemmas inhLemmas Meta.mkLambdaFVars #[negGoalFVar] proofOfFalse let goal := mkApp2 (.const ``Classical.byContradiction []) body negGoalImpFalse Meta.mkLambdaFVars bs goal diff --git a/Test/Bugs.lean b/Test/Bugs.lean index a667780..a487b00 100644 --- a/Test/Bugs.lean +++ b/Test/Bugs.lean @@ -80,9 +80,3 @@ end Set example (x : Nat) (primeset : Nat → Prop) (dvd : Nat → Nat → Prop) : ((∃ (i : _) (i_1 : primeset i), dvd i x) ↔ (∃ p, primeset p ∧ dvd p x)) := by auto - --- bug -set_option trace.auto.tactic true in -#eval do - let m ← Auto.runAutoOnConst ``Nat.sub_one_lt_of_lt - trace[auto.tactic] m!"{m}" From 9521975671fd62e33d9d16027c168f67212434e3 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 27 Nov 2024 18:38:03 -0800 Subject: [PATCH 15/25] batch test --- Auto/EvaluateAuto/ConstAnalysis.lean | 4 +-- Auto/EvaluateAuto/EnvAnalysis.lean | 4 +-- Auto/EvaluateAuto/TestCode.lean | 46 +++++++++++++++++++++++++--- 3 files changed, 46 insertions(+), 8 deletions(-) diff --git a/Auto/EvaluateAuto/ConstAnalysis.lean b/Auto/EvaluateAuto/ConstAnalysis.lean index 859f81b..2099a76 100644 --- a/Auto/EvaluateAuto/ConstAnalysis.lean +++ b/Auto/EvaluateAuto/ConstAnalysis.lean @@ -2,7 +2,7 @@ import Lean open Lean -namespace Auto +namespace EvalAuto def Name.getConstsOfModule (module : Name) : CoreM (Array Name) := do let mFile ← findOLean module @@ -124,4 +124,4 @@ def analyze : CoreM (Array (Array Name)) := do return (!(← Name.onlyBoolLogicInType name)) && (← Name.onlyNatBoolLogicInType name)) return #[logicThms, boolThms, natThms] -end Auto +end EvalAuto diff --git a/Auto/EvaluateAuto/EnvAnalysis.lean b/Auto/EvaluateAuto/EnvAnalysis.lean index e2bfcac..e08f856 100644 --- a/Auto/EvaluateAuto/EnvAnalysis.lean +++ b/Auto/EvaluateAuto/EnvAnalysis.lean @@ -2,10 +2,10 @@ import Lean open Lean -namespace Auto +namespace EvalAuto def mathlibModules : CoreM (Array Name) := do let u := (← getEnv).header.moduleNames return u.filter (fun name => name.components[0]? == .some `Mathlib) -end Auto +end EvalAuto diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index d176459..4f64085 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -3,9 +3,14 @@ import Auto.EvaluateAuto.ConstAnalysis import Auto.EvaluateAuto.EnvAnalysis import Auto.Tactic -open Lean +open Lean Auto -namespace Auto +initialize + registerTraceClass `auto.eval.printConfig + registerTraceClass `auto.eval.printProblem + registerTraceClass `auto.eval.printResult + +namespace EvalAuto inductive Result | success @@ -20,7 +25,20 @@ instance : ToMessageData Result where | .nonProp => "Result.nonProp" | .typeCheckFail => "Result.typeCheckFail" | .typeUnequal => "Result.typeUnequal" - | .autoException e => m!"Result.autoException :: {e.toMessageData}" + | .autoException e => m!"Result.autoException ::\n{e.toMessageData}" + +structure EvalConfig where + maxHeartbeats : Nat := 65536 + logFile : Option String := .none + +instance : ToString EvalConfig where + toString : EvalConfig → String + | ⟨maxHeartbeats, logFile⟩ => + let logFileStr := + match logFile with + | .some logFile => s!", logFile := {logFile}" + | .none => "" + s!"\{maxHeartbeats := {maxHeartbeats}{logFileStr}}" /-- Run `Lean-auto` on `lem.type`, using premises collected from `lem.proof` @@ -70,4 +88,24 @@ def runAutoOnConst (name : Name) : MetaM Result := do let lemmaV := {lemmaPre with proof := v} runAutoOnAutoLemma (.some name) lemmaV -end Auto +def runAutoOnConsts (config : EvalConfig) (names : Array Name) : MetaM Unit := do + let logFileHandle : Option IO.FS.Handle ← config.logFile.mapM (fun fname => IO.FS.Handle.mk fname .write) + trace[auto.eval.printConfig] m!"Config = {config}" + if let .some fhandle := logFileHandle then + fhandle.putStrLn s!"Config = {config}" + for name in names do + let ci ← Name.getCi name decl_name% + trace[auto.eval.printProblem] m!"Testing || {name} : {ci.type}" + if let .some fhandle := logFileHandle then + fhandle.putStrLn "" + fhandle.putStrLn s!"Testing || {name} : {← Lean.Meta.ppExpr ci.type}" + let result ← + -- **TODO: Setting maxHeartbeats not working** + withOptions (fun opts => Lean.maxHeartbeats.set opts config.maxHeartbeats) <| + withCurrHeartbeats <| do + runAutoOnConst name + trace[auto.eval.printResult] m!"{result}" + if let .some fhandle := logFileHandle then + fhandle.putStrLn (toString (← MessageData.format m!"{result}")) + +end EvalAuto From a7bef1889db5e88491277d5d2a3e1fa9fc6d2978 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Wed, 27 Nov 2024 19:06:30 -0800 Subject: [PATCH 16/25] fix maxheartbeat issue --- Auto/EvaluateAuto/TestCode.lean | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index 4f64085..73ec192 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -45,7 +45,7 @@ instance : ToString EvalConfig where Premises which only contain logic constants are filtered because they are assumed to be known by the prover -/ -def runAutoOnAutoLemma (declName? : Option Name) (lem : Auto.Lemma) : MetaM Result := do +private def runAutoOnAutoLemmaMeta (declName? : Option Name) (lem : Auto.Lemma) : MetaM Result := do if !(← Meta.isProp lem.type) then return .nonProp -- **TODO: Aux theorem like those ending in `.proof_1`** @@ -74,13 +74,16 @@ def runAutoOnAutoLemma (declName? : Option Name) (lem : Auto.Lemma) : MetaM Resu | _ => return .typeUnequal | Except.error _ => return .typeCheckFail +def runAutoOnAutoLemma (declName? : Option Name) (lem : Auto.Lemma) : CoreM Result := do + (runAutoOnAutoLemmaMeta declName? lem).run' + /-- Run `Lean-auto` on the type of ``name``, using premises collected from the proof of `name` Premises which only contain logic constants are filtered because they are assumed to be known by the prover -/ -def runAutoOnConst (name : Name) : MetaM Result := do +def runAutoOnConst (name : Name) : CoreM Result := do let ci ← Name.getCi name decl_name% let .some v := ci.value? | throwError "{decl_name%} :: {name} has no value" @@ -88,7 +91,7 @@ def runAutoOnConst (name : Name) : MetaM Result := do let lemmaV := {lemmaPre with proof := v} runAutoOnAutoLemma (.some name) lemmaV -def runAutoOnConsts (config : EvalConfig) (names : Array Name) : MetaM Unit := do +def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := do let logFileHandle : Option IO.FS.Handle ← config.logFile.mapM (fun fname => IO.FS.Handle.mk fname .write) trace[auto.eval.printConfig] m!"Config = {config}" if let .some fhandle := logFileHandle then @@ -98,10 +101,9 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : MetaM Unit := d trace[auto.eval.printProblem] m!"Testing || {name} : {ci.type}" if let .some fhandle := logFileHandle then fhandle.putStrLn "" - fhandle.putStrLn s!"Testing || {name} : {← Lean.Meta.ppExpr ci.type}" + fhandle.putStrLn s!"Testing || {name} : {← (Lean.Meta.ppExpr ci.type).run'}" let result ← - -- **TODO: Setting maxHeartbeats not working** - withOptions (fun opts => Lean.maxHeartbeats.set opts config.maxHeartbeats) <| + withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| withCurrHeartbeats <| do runAutoOnConst name trace[auto.eval.printResult] m!"{result}" From 53f85a0e638876591c0627e32aaafaefd843ef3e Mon Sep 17 00:00:00 2001 From: PratherConid Date: Fri, 29 Nov 2024 00:42:09 -0800 Subject: [PATCH 17/25] setting up test configuration --- Auto/EvaluateAuto/TestCode.lean | 63 +++++++++++++++++++++++++++++--- Auto/Solver/SMT.lean | 11 ++++-- Auto/Solver/TPTP.lean | 36 +++++++++++++++--- Auto/Tactic.lean | 4 -- Auto/Translation/LamFOL2SMT.lean | 2 +- 5 files changed, 97 insertions(+), 19 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index 73ec192..edd6f3c 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -27,18 +27,38 @@ instance : ToMessageData Result where | .typeUnequal => "Result.typeUnequal" | .autoException e => m!"Result.autoException ::\n{e.toMessageData}" +inductive SolverConfig where + | native + | leanSmt + | smt (solverName : Solver.SMT.SolverName) + | tptp (solverName : Solver.TPTP.SolverName) (path : String) + +instance : ToString SolverConfig where + toString : SolverConfig → String + | .native => "native" + | .leanSmt => "leanSmt" + | .smt sn => s!"smt {sn}" + | .tptp sn path => s!"tptp {sn} {path}" + structure EvalConfig where + /-- Timeout for native prover, e.g. Duper and Lean-smt -/ maxHeartbeats : Nat := 65536 + /-- Timeout for external provers, i.e. TPTP solvers and SMT solvers -/ + timeout : Nat := 10 + /-- Solver configuration -/ + solverConfig : SolverConfig + /-- Optional logfile for saving the result of the evaluation -/ logFile : Option String := .none instance : ToString EvalConfig where toString : EvalConfig → String - | ⟨maxHeartbeats, logFile⟩ => + | ⟨maxHeartbeats, timeout, solverConfig, logFile⟩ => let logFileStr := match logFile with | .some logFile => s!", logFile := {logFile}" | .none => "" - s!"\{maxHeartbeats := {maxHeartbeats}{logFileStr}}" + s!"\{maxHeartbeats := {maxHeartbeats}, timeout := {timeout}, " ++ + s!"solverConfig = {solverConfig}{logFileStr}}" /-- Run `Lean-auto` on `lem.type`, using premises collected from `lem.proof` @@ -91,6 +111,12 @@ def runAutoOnConst (name : Name) : CoreM Result := do let lemmaV := {lemmaPre with proof := v} runAutoOnAutoLemma (.some name) lemmaV +def disableAllSolvers (o : Options) : Options := + let o := auto.native.set o false + let o := auto.smt.set o false + let o := auto.tptp.set o false + o + def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := do let logFileHandle : Option IO.FS.Handle ← config.logFile.mapM (fun fname => IO.FS.Handle.mk fname .write) trace[auto.eval.printConfig] m!"Config = {config}" @@ -102,10 +128,35 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d if let .some fhandle := logFileHandle then fhandle.putStrLn "" fhandle.putStrLn s!"Testing || {name} : {← (Lean.Meta.ppExpr ci.type).run'}" - let result ← - withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| - withCurrHeartbeats <| do - runAutoOnConst name + let result : Result ← + match config.solverConfig with + | .native => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.native.set o true + o) <| + withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| + withCurrHeartbeats <| runAutoOnConst name + | .leanSmt => + throwError "Lean-SMT is currently not supported" + | .smt sn => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.smt.set o true + let o := auto.smt.solver.name.set o sn + let o := auto.smt.trust.set o true + o) <| runAutoOnConst name + | .tptp sn path => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.tptp.set o true + let o := auto.tptp.solver.name.set o sn + match sn with + | .zipperposition => auto.tptp.zipperposition.path.set o path + | .zeport _ => auto.tptp.eproverHo.path.set o path + | .eproverHo => auto.tptp.zeport.path.set o path + | .vampire => auto.tptp.vampire.path.set o path) <| + runAutoOnConst name trace[auto.eval.printResult] m!"{result}" if let .some fhandle := logFileHandle then fhandle.putStrLn (toString (← MessageData.format m!"{result}")) diff --git a/Auto/Solver/SMT.lean b/Auto/Solver/SMT.lean index 67fca0d..f965049 100644 --- a/Auto/Solver/SMT.lean +++ b/Auto/Solver/SMT.lean @@ -52,9 +52,6 @@ register_option auto.smt.rconsProof : Bool := { namespace Auto -open IR.SMT -open Parser.SMTSexp - namespace Solver.SMT inductive SolverName where @@ -80,11 +77,19 @@ instance : Lean.KVMap.Value SolverName where | "cvc5" => some .cvc5 | _ => none +end Auto.Solver.SMT + +open Auto.Solver.SMT in register_option auto.smt.solver.name : SolverName := { defValue := SolverName.z3 descr := "Name of the designated SMT solver. Use `none` to disable solver querying." } +namespace Auto.Solver.SMT + +open IR.SMT +open Parser.SMTSexp + abbrev SolverProc := IO.Process.Child ⟨.piped, .piped, .piped⟩ private def emitCommand (p : SolverProc) (c : IR.SMT.Command) : IO Unit := do diff --git a/Auto/Solver/TPTP.lean b/Auto/Solver/TPTP.lean index 121b236..a719edb 100644 --- a/Auto/Solver/TPTP.lean +++ b/Auto/Solver/TPTP.lean @@ -43,6 +43,7 @@ inductive SolverName where | zeport (zept : ZEPortType) -- E prover, higher-order version | eproverHo + | vampire deriving BEq, Hashable, Inhabited instance : ToString SolverName where @@ -53,6 +54,7 @@ instance : ToString SolverName where | .fo => "zeport-fo" | .lams => "zeport-lams" | .eproverHo => "eprover-ho" + | .vampire => "vampire" instance : Lean.KVMap.Value SolverName where toDataValue n := toString n @@ -61,8 +63,12 @@ instance : Lean.KVMap.Value SolverName where | "zeport-fo" => some (.zeport .fo) | "zeport-lams" => some (.zeport .lams) | "eprover-ho" => some .eproverHo + | "vampire" => some .vampire | _ => none +end Auto.Solver.TPTP + +open Auto.Solver.TPTP in register_option auto.tptp.solver.name : SolverName := { defValue := SolverName.zipperposition descr := "Name of the designated TPTP solver" @@ -73,16 +79,23 @@ register_option auto.tptp.zipperposition.path : String := { descr := "Path to zipperposition, defaults to \"zipperposition\"" } +register_option auto.tptp.zeport.path : String := { + defValue := "zeport" + descr := "Path to the zipperposition-E portfolio" +} + register_option auto.tptp.eproverHo.path : String := { defValue := "eprover-ho" descr := "Path to higher-order version of E theorem prover" } -register_option auto.tptp.zeport.path : String := { - defValue := "zeport" - descr := "Path to the zipperposition-E portfolio" +register_option auto.tptp.vampire.path : String := { + defValue := "vampire" + descr := "Path to vampire prover" } +namespace Auto.Solver.TPTP + abbrev SolverProc := IO.Process.Child ⟨.piped, .piped, .piped⟩ private def createAux (path : String) (args : Array String) : MetaM SolverProc := @@ -147,14 +160,27 @@ def queryE (query : String) : MetaM String := do solver.kill return stdout +def queryVampire (query : String) : MetaM String := do + let path := auto.tptp.vampire.path.get (← getOptions) + let tlim := auto.tptp.timeout.get (← getOptions) + let solver ← createAux path #["--mode", "casc", "--time_limit", s!"{tlim}"] + solver.stdin.putStr s!"{query}\n" + let (_, solver) ← solver.takeStdin + let stdout ← solver.stdout.readToEnd + let stderr ← solver.stderr.readToEnd + trace[auto.tptp.result] "Result: \nstderr:\n{stderr}\nstdout:\n{stdout}" + solver.kill + return stdout + def querySolver (query : String) : MetaM (Array Parser.TPTP.Command) := do if !(auto.tptp.get (← getOptions)) then throwError "{decl_name%} :: Unexpected error" let stdout ← (do match auto.tptp.solver.name.get (← getOptions) with | .zipperposition => queryZipperposition query - | .zeport zept => queryZEPort zept query - | .eproverHo => queryE query) + | .zeport zept => queryZEPort zept query + | .eproverHo => queryE query + | .vampire => queryVampire query) return ← Parser.TPTP.parse stdout end Solver.TPTP diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index 1afa718..a3b6a33 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -254,10 +254,6 @@ def collectAllLemmas return (lctxLemmas ++ userLemmas ++ defeqLemmas, inhFacts) open Embedding.Lam in -/-- - If TPTP succeeds, return unsat core - If TPTP fails, return none --/ def queryTPTP (exportFacts : Array REntry) : LamReif.ReifM (Array Embedding.Lam.REntry) := do let lamVarTy := (← LamReif.getVarVal).map Prod.snd let lamEVarTy ← LamReif.getLamEVarTy diff --git a/Auto/Translation/LamFOL2SMT.lean b/Auto/Translation/LamFOL2SMT.lean index 912c5b3..53f878c 100644 --- a/Auto/Translation/LamFOL2SMT.lean +++ b/Auto/Translation/LamFOL2SMT.lean @@ -319,7 +319,7 @@ private def lamBaseTerm2STerm_Arity1 (sni : SMTNamingInfo) (arg : STerm) : LamBa | t => throwError "{decl_name%} :: The arity of {repr t} is not 1" where solverName : MetaM Solver.SMT.SolverName := do - return Solver.SMT.auto.smt.solver.name.get (← getOptions) + return auto.smt.solver.name.get (← getOptions) mkSMTMsbExpr (n : Nat) (arg : STerm) : STerm := let andExpr := .qStrApp "bvand" #[arg, .qStrApp "bvshl" #[bitVec2STerm n 1, bitVec2STerm n (n - 1)]] .qStrApp "not" #[.qStrApp "=" #[andExpr, bitVec2STerm n 0]] From 052202306ba771ae0c15f9e49015dca5fc945d3f Mon Sep 17 00:00:00 2001 From: PratherConid Date: Fri, 29 Nov 2024 19:22:57 -0800 Subject: [PATCH 18/25] update TPTP and SMT invocation --- Auto/Solver/SMT.lean | 9 ++--- Auto/Solver/TPTP.lean | 37 ++++++++++++++------- Auto/Tactic.lean | 76 ++++++++++++++++++++++++++----------------- 3 files changed, 76 insertions(+), 46 deletions(-) diff --git a/Auto/Solver/SMT.lean b/Auto/Solver/SMT.lean index f965049..ff31deb 100644 --- a/Auto/Solver/SMT.lean +++ b/Auto/Solver/SMT.lean @@ -21,10 +21,13 @@ register_option auto.smt : Bool := { register_option auto.smt.trust : Bool := { defValue := false - descr := "When this option is set to `true`, auto closes the " ++ - "goal by `sorry` if SMT solver returns `unsat`" + descr := + "When this option is set to `true`, auto closes the " ++ + "goal by `autoSMTSorry` if SMT solver returns `unsat`" } +axiom autoSMTSorry.{u} (α : Sort u) : α + register_option auto.smt.timeout : Nat := { defValue := 10 descr := "Time limit for smt solvers (seconds)" @@ -111,8 +114,6 @@ where IO.Process.spawn {stdin := .piped, stdout := .piped, stderr := .piped, cmd := path, args := args} -axiom autoSMTSorry.{u} (α : Sort u) : α - def getSexp (s : String) : MetaM (Sexp × String) := match parseSexp s ⟨0⟩ {} with | .complete se p => return (se, Substring.toString ⟨s, p, s.endPos⟩) diff --git a/Auto/Solver/TPTP.lean b/Auto/Solver/TPTP.lean index a719edb..478e815 100644 --- a/Auto/Solver/TPTP.lean +++ b/Auto/Solver/TPTP.lean @@ -15,6 +15,15 @@ register_option auto.tptp : Bool := { descr := "Enable/Disable TPTP" } +register_option auto.tptp.trust : Bool := { + defValue := false + descr := + "When this option is set to `true`, auto closes the " ++ + "goal by `autoTPTPSorry` if TPTP solver proves the problem" +} + +axiom autoTPTPSorry.{u} (α : Sort u) : α + register_option auto.tptp.premiseSelection : Bool := { defValue := true descr := "Enable/Disable premise selection by TPTP solvers" @@ -102,7 +111,7 @@ private def createAux (path : String) (args : Array String) : MetaM SolverProc : IO.Process.spawn {stdin := .piped, stdout := .piped, stderr := .piped, cmd := path, args := args} -def queryZipperposition (query : String) : MetaM String := do +def queryZipperposition (query : String) : MetaM (Bool × String) := do let path := auto.tptp.zipperposition.path.get (← getOptions) let tlim := auto.tptp.timeout.get (← getOptions) let solver ← createAux path #["-i=tptp", "-o=tptp", "--mode=ho-competitive", s!"-t={tlim}"] @@ -112,9 +121,10 @@ def queryZipperposition (query : String) : MetaM String := do let stderr ← solver.stderr.readToEnd trace[auto.tptp.result] "Result: \nstderr:\n{stderr}\nstdout:\n{stdout}" solver.kill - return stdout + let proven := (stdout.splitOn "SZS status Unsatisfiable").length >= 2 + return (proven, stdout) -def queryZEPort (zept : ZEPortType) (query : String) : MetaM String := do +def queryZEPort (zept : ZEPortType) (query : String) : MetaM (Bool × String) := do let path := auto.tptp.zeport.path.get (← getOptions) -- To avoid concurrency issue, use `attempt` attempt <| IO.FS.createDir "./.zeport_ignore" @@ -138,7 +148,8 @@ def queryZEPort (zept : ZEPortType) (query : String) : MetaM String := do IO.FS.removeFile s!"./.zeport_ignore/problem{idx}.p" -- For synchronization, remove directory in the end IO.FS.removeDirAll s!"./.zeport_ignore/tmp{idx}" - return stdout + let proven := (stdout.splitOn "SZS status Unsatisfiable").length >= 2 + return (proven, stdout) where attempt (action : MetaM Unit) : MetaM Unit := try action catch _ => pure () createSolver (path : String) (idx : Nat) := do @@ -148,19 +159,20 @@ where | .fo => createAux "python3" #[path ++ "portfolio.fo.parallel.py", s!"./.zeport_ignore/problem{idx}.p", s!"{tlim}", "true"] | .lams => createAux "python3" #[path ++ "portfolio.lams.parallel.py", s!"./.zeport_ignore/problem{idx}.p", s!"{tlim}", s!"./.zeport_ignore/tmp{idx}", "true"] -def queryE (query : String) : MetaM String := do +def queryE (query : String) : MetaM (Bool × String) := do let path := auto.tptp.eproverHo.path.get (← getOptions) let tlim := auto.tptp.timeout.get (← getOptions) - let solver ← createAux path #["--tptp-format", s!"--cpu-limit={tlim}"] + let solver ← createAux path #["--tstp-format", s!"--cpu-limit={tlim}"] solver.stdin.putStr s!"{query}\n" let (_, solver) ← solver.takeStdin let stdout ← solver.stdout.readToEnd let stderr ← solver.stderr.readToEnd trace[auto.tptp.result] "Result: \nstderr:\n{stderr}\nstdout:\n{stdout}" solver.kill - return stdout + let proven := (stdout.splitOn "Proof found!").length >= 2 + return (proven, stdout) -def queryVampire (query : String) : MetaM String := do +def queryVampire (query : String) : MetaM (Bool × String) := do let path := auto.tptp.vampire.path.get (← getOptions) let tlim := auto.tptp.timeout.get (← getOptions) let solver ← createAux path #["--mode", "casc", "--time_limit", s!"{tlim}"] @@ -170,18 +182,19 @@ def queryVampire (query : String) : MetaM String := do let stderr ← solver.stderr.readToEnd trace[auto.tptp.result] "Result: \nstderr:\n{stderr}\nstdout:\n{stdout}" solver.kill - return stdout + let proven := (stdout.splitOn "Refutation found. Thanks to Tanya!").length >= 2 + return (proven, stdout) -def querySolver (query : String) : MetaM (Array Parser.TPTP.Command) := do +def querySolver (query : String) : MetaM (Bool × Array Parser.TPTP.Command) := do if !(auto.tptp.get (← getOptions)) then throwError "{decl_name%} :: Unexpected error" - let stdout ← (do + let (proven, stdout) ← (do match auto.tptp.solver.name.get (← getOptions) with | .zipperposition => queryZipperposition query | .zeport zept => queryZEPort zept query | .eproverHo => queryE query | .vampire => queryVampire query) - return ← Parser.TPTP.parse stdout + return (proven, ← Parser.TPTP.parse stdout) end Solver.TPTP diff --git a/Auto/Tactic.lean b/Auto/Tactic.lean index a3b6a33..d75af4a 100644 --- a/Auto/Tactic.lean +++ b/Auto/Tactic.lean @@ -254,32 +254,38 @@ def collectAllLemmas return (lctxLemmas ++ userLemmas ++ defeqLemmas, inhFacts) open Embedding.Lam in -def queryTPTP (exportFacts : Array REntry) : LamReif.ReifM (Array Embedding.Lam.REntry) := do - let lamVarTy := (← LamReif.getVarVal).map Prod.snd - let lamEVarTy ← LamReif.getLamEVarTy - let exportLamTerms ← exportFacts.mapM (fun re => do - match re with - | .valid [] t => return t - | _ => throwError "{decl_name%} :: Unexpected error") - let query ← lam2TH0 lamVarTy lamEVarTy exportLamTerms - trace[auto.tptp.printQuery] "\n{query}" - let tptpProof ← Solver.TPTP.querySolver query - try - let proofSteps ← Parser.TPTP.getProof lamVarTy lamEVarTy tptpProof - for step in proofSteps do - trace[auto.tptp.printProof] "{step}" - catch e => - trace[auto.tptp.printProof] "TPTP proof reification failed with {e.toMessageData}" - let unsatCore ← Parser.TPTP.unsatCore tptpProof - let mut ret := #[] - for n in unsatCore do - let .some re := exportFacts[n]? - | throwError "{decl_name%} :: Index {n} out of range" - ret := ret.push re - return ret +def queryTPTP (exportFacts : Array REntry) : LamReif.ReifM (Option Expr × Option (Array REntry)) := do + let lamVarTy := (← LamReif.getVarVal).map Prod.snd + let lamEVarTy ← LamReif.getLamEVarTy + let exportLamTerms ← exportFacts.mapM (fun re => do + match re with + | .valid [] t => return t + | _ => throwError "{decl_name%} :: Unexpected error") + let query ← lam2TH0 lamVarTy lamEVarTy exportLamTerms + trace[auto.tptp.printQuery] "\n{query}" + let (proven, tptpProof) ← Solver.TPTP.querySolver query + if !proven then + return (.none, .none) + try + let proofSteps ← Parser.TPTP.getProof lamVarTy lamEVarTy tptpProof + for step in proofSteps do + trace[auto.tptp.printProof] "{step}" + catch e => + trace[auto.tptp.printProof] "TPTP proof reification failed with {e.toMessageData}" + let unsatCoreIds ← Parser.TPTP.unsatCore tptpProof + let mut unsatCore := #[] + for n in unsatCoreIds do + let .some re := exportFacts[n]? + | throwError "{decl_name%} :: Index {n} out of range" + unsatCore := unsatCore.push re + if (auto.tptp.trust.get (← getOptions)) then + logWarning "Trusting TPTP solvers. `autoTPTPSorry` is used to discharge the goal." + return (← Meta.mkAppM ``autoTPTPSorry #[Expr.const ``False []], unsatCore) + else + return (.none, unsatCore) open Embedding.Lam in -def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : LamReif.ReifM (Option Expr) := do +def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : LamReif.ReifM (Option Expr × Option (Array REntry)) := do let lamVarTy := (← LamReif.getVarVal).map Prod.snd let lamEVarTy ← LamReif.getLamEVarTy let exportLamTerms ← exportFacts.mapM (fun re => do @@ -294,7 +300,7 @@ def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : L if (auto.smt.save.get (← getOptions)) then Solver.SMT.saveQuery commands let .some (unsatCore, proof) ← Solver.SMT.querySolver commands - | return .none + | return (.none, .none) let unsatCoreIds ← Solver.SMT.validFactOfUnsatCore unsatCore -- **Print valuation of SMT atoms** SMT.withExprValuation sni state.h2lMap (fun tyValMap varValMap etomValMap => do @@ -321,14 +327,20 @@ def querySMT (exportFacts : Array REntry) (exportInds : Array MutualIndInfo) : L | throwError "{decl_name%} :: Index {id} of `exportLamTerm` out of range" let vderiv ← LamReif.collectDerivFor (.valid [] t) trace[auto.smt.unsatCore.deriv] "|valid_fact_{id}| : {vderiv}" + -- **Getting unsatCore** + let mut unsatCore := #[] + for id in unsatCoreIds do + let .some re := exportFacts[id]? + | throwError "{decl_name%} :: Index {id} of `exportFacts` out of range" + unsatCore := unsatCore.push re if auto.smt.rconsProof.get (← getOptions) then let (_, _) ← Solver.SMT.getSexp proof logWarning "Proof reconstruction is not implemented." if (auto.smt.trust.get (← getOptions)) then logWarning "Trusting SMT solvers. `autoSMTSorry` is used to discharge the goal." - return .some (← Meta.mkAppM ``Solver.SMT.autoSMTSorry #[Expr.const ``False []]) + return (← Meta.mkAppM ``autoSMTSorry #[Expr.const ``False []], unsatCore) else - return .none + return (.none, unsatCore) open LamReif Embedding.Lam in /-- @@ -452,13 +464,17 @@ def runAuto exportFacts := exportFacts' -- **TPTP invocation and Premise Selection** if auto.tptp.get (← getOptions) then + let (proof, unsatCore) ← queryTPTP exportFacts + if let .some proof := proof then + return proof let premiseSel? := auto.tptp.premiseSelection.get (← getOptions) - let unsatCore ← queryTPTP exportFacts if premiseSel? then - exportFacts := unsatCore + if let .some unsatCore := unsatCore then + exportFacts := unsatCore -- **SMT** if auto.smt.get (← getOptions) then - if let .some proof ← querySMT exportFacts exportInds then + let (proof, _) ← querySMT exportFacts exportInds + if let .some proof := proof then return proof -- **Native Prover** exportFacts := exportFacts.append (← LamReif.auxLemmas exportFacts) From a0ce592128ffc4dd82621f679d7d45c48e2b85ff Mon Sep 17 00:00:00 2001 From: PratherConid Date: Fri, 29 Nov 2024 19:59:22 -0800 Subject: [PATCH 19/25] fix parsing and evaluation issue --- Auto/EvaluateAuto/TestCode.lean | 5 +++-- Auto/Parser/TPTP.lean | 4 ++-- 2 files changed, 5 insertions(+), 4 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index edd6f3c..fa70e0d 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -151,10 +151,11 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := disableAllSolvers o let o := auto.tptp.set o true let o := auto.tptp.solver.name.set o sn + let o := auto.tptp.trust.set o true match sn with | .zipperposition => auto.tptp.zipperposition.path.set o path - | .zeport _ => auto.tptp.eproverHo.path.set o path - | .eproverHo => auto.tptp.zeport.path.set o path + | .zeport _ => auto.tptp.zeport.path.set o path + | .eproverHo => auto.tptp.eproverHo.path.set o path | .vampire => auto.tptp.vampire.path.set o path) <| runAutoOnConst name trace[auto.eval.printResult] m!"{result}" diff --git a/Auto/Parser/TPTP.lean b/Auto/Parser/TPTP.lean index fd66be9..b0ef5c4 100644 --- a/Auto/Parser/TPTP.lean +++ b/Auto/Parser/TPTP.lean @@ -859,8 +859,8 @@ def getProof (lamVarTy lamEVarTy : Array LamSort) (cmds : Array Command) : MetaM else match val with | ⟨.ident name, [ty]⟩ => - -- In zipperposition, skolems start with `sk_` - if name.take 3 == "sk_" then + -- In zipperposition, skolems start with `sk_` or `#sk_` + if name.take 3 == "sk_" || name.take 3 == "#sk" then match term2LamTerm pi ty with | .sort s => pi := pi.addSkolem name s | lc => throwError "{decl_name%} :: Unexpected LamConstr {lc}, expected sort" From 28de2068a965f15b76400b8970876310dacd282d Mon Sep 17 00:00:00 2001 From: PratherConid Date: Fri, 29 Nov 2024 20:56:35 -0800 Subject: [PATCH 20/25] add randeval --- Auto/EvaluateAuto/EnvAnalysis.lean | 9 +++++++++ Auto/EvaluateAuto/TestCode.lean | 10 ++++++++-- 2 files changed, 17 insertions(+), 2 deletions(-) diff --git a/Auto/EvaluateAuto/EnvAnalysis.lean b/Auto/EvaluateAuto/EnvAnalysis.lean index e08f856..d16437e 100644 --- a/Auto/EvaluateAuto/EnvAnalysis.lean +++ b/Auto/EvaluateAuto/EnvAnalysis.lean @@ -8,4 +8,13 @@ def mathlibModules : CoreM (Array Name) := do let u := (← getEnv).header.moduleNames return u.filter (fun name => name.components[0]? == .some `Mathlib) +/-- Pick `n` elements from array `xs`. Elements may duplicate -/ +def Array.randPick {α} (xs : Array α) (n : Nat) : IO (Array α) := do + let mut ret := #[] + for _ in [0:n] do + let rd ← IO.rand 0 (xs.size - 1) + if h : rd < xs.size then + ret := ret.push (xs[rd]'h) + return ret + end EvalAuto diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index fa70e0d..b2f5ba4 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -135,8 +135,9 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := disableAllSolvers o let o := auto.native.set o true o) <| - withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| - withCurrHeartbeats <| runAutoOnConst name + withCurrHeartbeats <| + withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| + runAutoOnConst name | .leanSmt => throwError "Lean-SMT is currently not supported" | .smt sn => @@ -162,4 +163,9 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d if let .some fhandle := logFileHandle then fhandle.putStrLn (toString (← MessageData.format m!"{result}")) +def randEval (cfg : EvalConfig) (n : Nat) : CoreM Unit := do + let hthms ← allHumanTheorems + let asel ← Array.randPick hthms n + runAutoOnConsts cfg asel + end EvalAuto From 436da908161c9814c463296ba27d2c9333db6399 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Mon, 2 Dec 2024 11:10:07 -0800 Subject: [PATCH 21/25] passing timeout option --- Auto/EvaluateAuto/TestCode.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index b2f5ba4..703116a 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -145,6 +145,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := disableAllSolvers o let o := auto.smt.set o true let o := auto.smt.solver.name.set o sn + let o := auto.smt.timeout.set o config.timeout let o := auto.smt.trust.set o true o) <| runAutoOnConst name | .tptp sn path => @@ -152,6 +153,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := disableAllSolvers o let o := auto.tptp.set o true let o := auto.tptp.solver.name.set o sn + let o := auto.tptp.timeout.set o config.timeout let o := auto.tptp.trust.set o true match sn with | .zipperposition => auto.tptp.zipperposition.path.set o path From 8fc8cad27421e098b55eeae4a0bb9b71a0b9052b Mon Sep 17 00:00:00 2001 From: PratherConid Date: Tue, 3 Dec 2024 17:16:06 -0800 Subject: [PATCH 22/25] add --enum-inst option for cvc5 --- Auto/Solver/SMT.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Auto/Solver/SMT.lean b/Auto/Solver/SMT.lean index ff31deb..0b31907 100644 --- a/Auto/Solver/SMT.lean +++ b/Auto/Solver/SMT.lean @@ -108,7 +108,7 @@ def createSolver (name : SolverName) : MetaM SolverProc := do | .none => throwError "{decl_name%} :: Unexpected error" | .z3 => createAux "z3" #["-in", "-smt2", s!"-T:{tlim}"] | .cvc4 => throwError "cvc4 is not supported" - | .cvc5 => createAux "cvc5" #[s!"--tlimit={tlim * 1000}", "--produce-models"] + | .cvc5 => createAux "cvc5" #[s!"--tlimit={tlim * 1000}", "--produce-models", "--enum-inst"] where createAux (path : String) (args : Array String) : MetaM SolverProc := IO.Process.spawn {stdin := .piped, stdout := .piped, stderr := .piped, From a958f1e9715985cb302894a20cc2fc8f0ef43001 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Mon, 9 Dec 2024 00:49:08 -0800 Subject: [PATCH 23/25] add evaluation file support --- Auto/EvaluateAuto/NameArr.lean | 80 +++++++++++++++++++++++++++++++++ Auto/EvaluateAuto/TestCode.lean | 9 ++++ 2 files changed, 89 insertions(+) create mode 100644 Auto/EvaluateAuto/NameArr.lean diff --git a/Auto/EvaluateAuto/NameArr.lean b/Auto/EvaluateAuto/NameArr.lean new file mode 100644 index 0000000..0393e7e --- /dev/null +++ b/Auto/EvaluateAuto/NameArr.lean @@ -0,0 +1,80 @@ +import Lean +open Lean + +namespace EvalAuto + +/-- + Unique string representation of array of names + We use `.` to separate fields of a name, and `\n` to separate names. + A `.` is appended to the end of each name. + A `\n` is further appended to the end of the last name. + We use `\` to escape `.` and `\n` occurring in the fields of names, i.e., + `\.` represents literal `.`, `\\n` represents literal `\n`, and + `\\` represents literal `\`. + To distinguish `mkStr` and `mkNum`, we prepend `\` to all `mkNum` fields +-/ +def NameArray.repr (ns : Array Name) : String := + let strRepr (s : String) : String := + ((s.replace "\\" "\\\\").replace "." "\\.").replace "\n" "\\\n" + let compRepr (c : Name) : String := + match c with + | .anonymous => "" + | .mkNum _ n => s!"\\{n}" + | .mkStr _ s => strRepr s + let nameRepr (n : Name) : String := + String.join (n.components.map (fun c => compRepr c ++ ".")) + String.join (ns.map (fun n => nameRepr n ++ "\n")).toList + +/- + Parse string representation of array of names + We use `.` to separate fields of a name, and `\n` to separate names + We use `\` to escape `.` and `\n` occurring in the fields of names, i.e., + `\.` represents literal `.`, `\\n` represents literal `\n`, and + `\\` represents literal `\`. + To distinguish `mkStr` and `mkNum`, we prepend `\` to all `mkNum` fields +-/ +def NameArray.parse (repr : String) : Array Name := Id.run <| do + let mut curField := "" + let mut curFields : Array (String ⊕ Nat) := #[] + let mut allNames : Array Name := #[] + let mut escape := false + let mut num := false + for c in repr.data do + if !escape && c == '.' then + if num then + curFields := curFields.push (.inr ((String.toNat? curField).getD 0)) + else + curFields := curFields.push (.inl curField) + curField := "" + num := false + continue + if !escape && c == '\n' then + let curName := curFields.foldl (fun prev sn => + match sn with + | .inl s => Name.mkStr prev s + | .inr n => Name.mkNum prev n) .anonymous + allNames := allNames.push curName + curFields := #[] + continue + if escape && c.isDigit then + num := true + if c == '\\' then + if escape then + escape := false + curField := curField.push c + else + escape := true + else + escape := false + curField := curField.push c + return allNames + +def NameArray.save (ns : Array Name) (fname : String) : IO Unit := do + let fd ← IO.FS.Handle.mk fname .write + fd.putStr (NameArray.repr ns) + +def NameArray.load (fname : String) : IO (Array Name) := do + let fd ← IO.FS.Handle.mk fname .read + return NameArray.parse (← fd.readToEnd) + +end EvalAuto diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index 703116a..a4aa3df 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -1,6 +1,7 @@ import Lean import Auto.EvaluateAuto.ConstAnalysis import Auto.EvaluateAuto.EnvAnalysis +import Auto.EvaluateAuto.NameArr import Auto.Tactic open Lean Auto @@ -122,6 +123,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d trace[auto.eval.printConfig] m!"Config = {config}" if let .some fhandle := logFileHandle then fhandle.putStrLn s!"Config = {config}" + let startTime ← IO.monoMsNow for name in names do let ci ← Name.getCi name decl_name% trace[auto.eval.printProblem] m!"Testing || {name} : {ci.type}" @@ -164,6 +166,13 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d trace[auto.eval.printResult] m!"{result}" if let .some fhandle := logFileHandle then fhandle.putStrLn (toString (← MessageData.format m!"{result}")) + if let .some fhandle := logFileHandle then + fhandle.putStrLn "" + fhandle.putStrLn s!"Elapsed time: {(← IO.monoMsNow) - startTime} ms" + +def namesFileEval (cfg : EvalConfig) (fname : String) : CoreM Unit := do + let names ← NameArray.load fname + runAutoOnConsts cfg names def randEval (cfg : EvalConfig) (n : Nat) : CoreM Unit := do let hthms ← allHumanTheorems From c024824dd470e1692056d55672e85bc39767e9d1 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Mon, 9 Dec 2024 21:54:13 -0800 Subject: [PATCH 24/25] change timeout setting --- Auto/EvaluateAuto/TestCode.lean | 64 ++++++++++++++++----------------- 1 file changed, 31 insertions(+), 33 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index a4aa3df..73f91a7 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -42,7 +42,7 @@ instance : ToString SolverConfig where | .tptp sn path => s!"tptp {sn} {path}" structure EvalConfig where - /-- Timeout for native prover, e.g. Duper and Lean-smt -/ + /-- Timeout for Lean code (Lean-auto + native provers) -/ maxHeartbeats : Nat := 65536 /-- Timeout for external provers, i.e. TPTP solvers and SMT solvers -/ timeout : Nat := 10 @@ -130,39 +130,37 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d if let .some fhandle := logFileHandle then fhandle.putStrLn "" fhandle.putStrLn s!"Testing || {name} : {← (Lean.Meta.ppExpr ci.type).run'}" - let result : Result ← - match config.solverConfig with - | .native => - withOptions (fun o => - let o := disableAllSolvers o - let o := auto.native.set o true - o) <| - withCurrHeartbeats <| - withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| + let result : Result ← withCurrHeartbeats <| + withReader (fun ctx => {ctx with maxHeartbeats := config.maxHeartbeats * 1000}) <| + match config.solverConfig with + | .native => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.native.set o true + o) <| runAutoOnConst name + | .leanSmt => + throwError "Lean-SMT is currently not supported" + | .smt sn => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.smt.set o true + let o := auto.smt.solver.name.set o sn + let o := auto.smt.timeout.set o config.timeout + let o := auto.smt.trust.set o true + o) <| runAutoOnConst name + | .tptp sn path => + withOptions (fun o => + let o := disableAllSolvers o + let o := auto.tptp.set o true + let o := auto.tptp.solver.name.set o sn + let o := auto.tptp.timeout.set o config.timeout + let o := auto.tptp.trust.set o true + match sn with + | .zipperposition => auto.tptp.zipperposition.path.set o path + | .zeport _ => auto.tptp.zeport.path.set o path + | .eproverHo => auto.tptp.eproverHo.path.set o path + | .vampire => auto.tptp.vampire.path.set o path) <| runAutoOnConst name - | .leanSmt => - throwError "Lean-SMT is currently not supported" - | .smt sn => - withOptions (fun o => - let o := disableAllSolvers o - let o := auto.smt.set o true - let o := auto.smt.solver.name.set o sn - let o := auto.smt.timeout.set o config.timeout - let o := auto.smt.trust.set o true - o) <| runAutoOnConst name - | .tptp sn path => - withOptions (fun o => - let o := disableAllSolvers o - let o := auto.tptp.set o true - let o := auto.tptp.solver.name.set o sn - let o := auto.tptp.timeout.set o config.timeout - let o := auto.tptp.trust.set o true - match sn with - | .zipperposition => auto.tptp.zipperposition.path.set o path - | .zeport _ => auto.tptp.zeport.path.set o path - | .eproverHo => auto.tptp.eproverHo.path.set o path - | .vampire => auto.tptp.vampire.path.set o path) <| - runAutoOnConst name trace[auto.eval.printResult] m!"{result}" if let .some fhandle := logFileHandle then fhandle.putStrLn (toString (← MessageData.format m!"{result}")) From d437b8d40700efcf39251e7fdd0c5520f296bd26 Mon Sep 17 00:00:00 2001 From: PratherConid Date: Mon, 9 Dec 2024 22:34:36 -0800 Subject: [PATCH 25/25] update monomode --- Auto/EvaluateAuto/TestCode.lean | 3 +++ Auto/Translation/Monomorphization.lean | 8 ++++++-- 2 files changed, 9 insertions(+), 2 deletions(-) diff --git a/Auto/EvaluateAuto/TestCode.lean b/Auto/EvaluateAuto/TestCode.lean index 73f91a7..bc53a9c 100644 --- a/Auto/EvaluateAuto/TestCode.lean +++ b/Auto/EvaluateAuto/TestCode.lean @@ -137,6 +137,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d withOptions (fun o => let o := disableAllSolvers o let o := auto.native.set o true + let o := auto.mono.mode.set o MonoMode.hol o) <| runAutoOnConst name | .leanSmt => throwError "Lean-SMT is currently not supported" @@ -147,6 +148,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := auto.smt.solver.name.set o sn let o := auto.smt.timeout.set o config.timeout let o := auto.smt.trust.set o true + let o := auto.mono.mode.set o MonoMode.fol o) <| runAutoOnConst name | .tptp sn path => withOptions (fun o => @@ -155,6 +157,7 @@ def runAutoOnConsts (config : EvalConfig) (names : Array Name) : CoreM Unit := d let o := auto.tptp.solver.name.set o sn let o := auto.tptp.timeout.set o config.timeout let o := auto.tptp.trust.set o true + let o := auto.mono.mode.set o MonoMode.fol match sn with | .zipperposition => auto.tptp.zipperposition.path.set o path | .zeport _ => auto.tptp.zeport.path.set o path diff --git a/Auto/Translation/Monomorphization.lean b/Auto/Translation/Monomorphization.lean index bc41cdd..83b61a8 100644 --- a/Auto/Translation/Monomorphization.lean +++ b/Auto/Translation/Monomorphization.lean @@ -30,6 +30,8 @@ register_option auto.mono.recordInstInst : Bool := { descr := "Whether to record instances of constants with the `instance` attribute" } +namespace Auto + inductive MonoMode where | fol -- First-order logic | hol -- Monomorphic higher-order logic @@ -47,8 +49,10 @@ instance : Lean.KVMap.Value MonoMode where | "hol" => some .hol | _ => none -register_option auto.mono.mode : MonoMode := { - defValue := MonoMode.hol +end Auto + +register_option auto.mono.mode : Auto.MonoMode := { + defValue := Auto.MonoMode.hol descr := "Operation mode of monomorphization" }