chore: Version 0.3 #136
|
@ -24,8 +24,11 @@ def env_catalog (env: Environment): Array Name := env.constants.fold (init := #[
|
|||
|
||||
@[export pantograph_environment_module_of_name]
|
||||
def module_of_name (env: Environment) (name: Name): Option Name := do
|
||||
let moduleId ← env.getModuleIdxFor? name
|
||||
return env.allImportedModuleNames.get! moduleId.toNat
|
||||
let moduleId ← env.getModuleIdxFor? name
|
||||
if h : moduleId.toNat < env.allImportedModuleNames.size then
|
||||
return env.allImportedModuleNames.get moduleId.toNat h
|
||||
else
|
||||
.none
|
||||
|
||||
def toCompactSymbolName (n: Name) (info: ConstantInfo): String :=
|
||||
let pref := match info with
|
||||
|
|
|
@ -74,13 +74,14 @@ protected def GoalState.metaState (state: GoalState): Meta.State :=
|
|||
protected def GoalState.coreState (state: GoalState): Core.SavedState :=
|
||||
state.savedState.term.meta.core
|
||||
|
||||
protected def GoalState.withContext (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do
|
||||
protected def GoalState.withContext' (state: GoalState) (mvarId: MVarId) (m: MetaM α): MetaM α := do
|
||||
mvarId.withContext m |>.run' (← read) state.metaState
|
||||
|
||||
protected def GoalState.withContext { m } [MonadControlT MetaM m] [Monad m] (state: GoalState) (mvarId: MVarId) : m α → m α :=
|
||||
Meta.mapMetaM <| state.withContext' mvarId
|
||||
protected def GoalState.withParentContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||
Meta.mapMetaM <| state.withContext state.parentMVar?.get!
|
||||
Meta.mapMetaM <| state.withContext' state.parentMVar?.get!
|
||||
protected def GoalState.withRootContext { n } [MonadControlT MetaM n] [Monad n] (state: GoalState): n α → n α :=
|
||||
Meta.mapMetaM <| state.withContext state.root
|
||||
Meta.mapMetaM <| state.withContext' state.root
|
||||
|
||||
private def GoalState.mvars (state: GoalState): SSet MVarId :=
|
||||
state.mctx.decls.foldl (init := .empty) fun acc k _ => acc.insert k
|
||||
|
|
|
@ -191,22 +191,6 @@ protected def GoalState.tryDefine (state: GoalState) (goal: MVarId) (binderName:
|
|||
runTermElabM do
|
||||
state.restoreElabM
|
||||
state.tryTacticM goal (Tactic.evalDefine binderName.toName expr)
|
||||
@[export pantograph_goal_try_motivated_apply_m]
|
||||
protected def GoalState.tryMotivatedApply (state: GoalState) (goal: MVarId) (recursor: String):
|
||||
Elab.TermElabM TacticResult := do
|
||||
state.restoreElabM
|
||||
let recursor ← match (← parseTermM recursor) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => return .parseError error
|
||||
state.tryTacticM goal (tacticM := Tactic.evalMotivatedApply recursor)
|
||||
@[export pantograph_goal_try_no_confuse_m]
|
||||
protected def GoalState.tryNoConfuse (state: GoalState) (goal: MVarId) (eq: String):
|
||||
Elab.TermElabM TacticResult := do
|
||||
state.restoreElabM
|
||||
let eq ← match (← parseTermM eq) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => return .parseError error
|
||||
state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq)
|
||||
@[export pantograph_goal_try_draft_m]
|
||||
protected def GoalState.tryDraft (state: GoalState) (goal: MVarId) (expr: String): CoreM TacticResult := do
|
||||
let expr ← match (← parseTermM expr) with
|
||||
|
|
|
@ -1,5 +1,2 @@
|
|||
import Pantograph.Tactic.Assign
|
||||
import Pantograph.Tactic.Congruence
|
||||
import Pantograph.Tactic.MotivatedApply
|
||||
import Pantograph.Tactic.NoConfuse
|
||||
import Pantograph.Tactic.Prograde
|
||||
|
|
|
@ -1,98 +0,0 @@
|
|||
import Lean
|
||||
|
||||
open Lean
|
||||
|
||||
namespace Pantograph.Tactic
|
||||
|
||||
def congruenceArg (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
|
||||
mvarId.checkNotAssigned `Pantograph.Tactic.congruenceArg
|
||||
let target ← mvarId.getType
|
||||
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
|
||||
let userName := (← mvarId.getDecl).userName
|
||||
|
||||
let u ← Meta.mkFreshLevelMVar
|
||||
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
|
||||
(tag := userName ++ `α)
|
||||
let f ← Meta.mkFreshExprSyntheticOpaqueMVar (.forallE .anonymous α β .default)
|
||||
(tag := userName ++ `f)
|
||||
let a₁ ← Meta.mkFreshExprSyntheticOpaqueMVar α
|
||||
(tag := userName ++ `a₁)
|
||||
let a₂ ← Meta.mkFreshExprSyntheticOpaqueMVar α
|
||||
(tag := userName ++ `a₂)
|
||||
let h ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq a₁ a₂)
|
||||
(tag := userName ++ `h)
|
||||
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f a₁) (.app f a₂)) target
|
||||
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
|
||||
(tag := userName ++ `conduit)
|
||||
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrArg f h)
|
||||
let result := [α, a₁, a₂, f, h, conduit]
|
||||
return result.map (·.mvarId!)
|
||||
|
||||
def evalCongruenceArg: Elab.Tactic.TacticM Unit := do
|
||||
let goal ← Elab.Tactic.getMainGoal
|
||||
let nextGoals ← congruenceArg goal
|
||||
Elab.Tactic.replaceMainGoal nextGoals
|
||||
|
||||
def congruenceFun (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
|
||||
mvarId.checkNotAssigned `Pantograph.Tactic.congruenceFun
|
||||
let target ← mvarId.getType
|
||||
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
|
||||
let userName := (← mvarId.getDecl).userName
|
||||
let u ← Meta.mkFreshLevelMVar
|
||||
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
|
||||
(tag := userName ++ `α)
|
||||
let fType := .forallE .anonymous α β .default
|
||||
let f₁ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
|
||||
(tag := userName ++ `f₁)
|
||||
let f₂ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
|
||||
(tag := userName ++ `f₂)
|
||||
let a ← Meta.mkFreshExprSyntheticOpaqueMVar α
|
||||
(tag := userName ++ `a)
|
||||
let h ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq f₁ f₂)
|
||||
(tag := userName ++ `h)
|
||||
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a) (.app f₂ a)) target
|
||||
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
|
||||
(tag := userName ++ `conduit)
|
||||
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongrFun h a)
|
||||
let result := [α, f₁, f₂, h, a, conduit]
|
||||
return result.map (·.mvarId!)
|
||||
|
||||
def evalCongruenceFun: Elab.Tactic.TacticM Unit := do
|
||||
let goal ← Elab.Tactic.getMainGoal
|
||||
let nextGoals ← congruenceFun goal
|
||||
Elab.Tactic.replaceMainGoal nextGoals
|
||||
|
||||
def congruence (mvarId: MVarId): MetaM (List MVarId) := mvarId.withContext do
|
||||
mvarId.checkNotAssigned `Pantograph.Tactic.congruence
|
||||
let target ← mvarId.getType
|
||||
let .some (β, _, _) := (← instantiateMVars target).eq? | throwError "Goal is not an Eq"
|
||||
let userName := (← mvarId.getDecl).userName
|
||||
let u ← Meta.mkFreshLevelMVar
|
||||
let α ← Meta.mkFreshExprSyntheticOpaqueMVar (mkSort u)
|
||||
(tag := userName ++ `α)
|
||||
let fType := .forallE .anonymous α β .default
|
||||
let f₁ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
|
||||
(tag := userName ++ `f₁)
|
||||
let f₂ ← Meta.mkFreshExprSyntheticOpaqueMVar fType
|
||||
(tag := userName ++ `f₂)
|
||||
let a₁ ← Meta.mkFreshExprSyntheticOpaqueMVar α
|
||||
(tag := userName ++ `a₁)
|
||||
let a₂ ← Meta.mkFreshExprSyntheticOpaqueMVar α
|
||||
(tag := userName ++ `a₂)
|
||||
let h₁ ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq f₁ f₂)
|
||||
(tag := userName ++ `h₁)
|
||||
let h₂ ← Meta.mkFreshExprSyntheticOpaqueMVar (← Meta.mkEq a₁ a₂)
|
||||
(tag := userName ++ `h₂)
|
||||
let conduitType ← Meta.mkEq (← Meta.mkEq (.app f₁ a₁) (.app f₂ a₂)) target
|
||||
let conduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType
|
||||
(tag := userName ++ `conduit)
|
||||
mvarId.assign $ ← Meta.mkEqMP conduit (← Meta.mkCongr h₁ h₂)
|
||||
let result := [α, f₁, f₂, a₁, a₂, h₁, h₂, conduit]
|
||||
return result.map (·.mvarId!)
|
||||
|
||||
def evalCongruence: Elab.Tactic.TacticM Unit := do
|
||||
let goal ← Elab.Tactic.getMainGoal
|
||||
let nextGoals ← congruence goal
|
||||
Elab.Tactic.replaceMainGoal nextGoals
|
||||
|
||||
end Pantograph.Tactic
|
|
@ -1,106 +0,0 @@
|
|||
import Lean
|
||||
|
||||
open Lean
|
||||
|
||||
namespace Pantograph.Tactic
|
||||
|
||||
def getForallArgsBody: Expr → List Expr × Expr
|
||||
| .forallE _ d b _ =>
|
||||
let (innerArgs, innerBody) := getForallArgsBody b
|
||||
(d :: innerArgs, innerBody)
|
||||
| e => ([], e)
|
||||
|
||||
def replaceForallBody: Expr → Expr → Expr
|
||||
| .forallE param domain body binderInfo, target =>
|
||||
let body := replaceForallBody body target
|
||||
.forallE param domain body binderInfo
|
||||
| _, target => target
|
||||
|
||||
structure RecursorWithMotive where
|
||||
args: List Expr
|
||||
body: Expr
|
||||
|
||||
-- .bvar index for the motive and major from the body
|
||||
iMotive: Nat
|
||||
|
||||
namespace RecursorWithMotive
|
||||
|
||||
protected def nArgs (info: RecursorWithMotive): Nat := info.args.length
|
||||
|
||||
protected def getMotiveType (info: RecursorWithMotive): Expr :=
|
||||
let level := info.nArgs - info.iMotive - 1
|
||||
let a := info.args.get! level
|
||||
a
|
||||
|
||||
protected def surrogateMotiveType (info: RecursorWithMotive) (mvars: Array Expr) (resultant: Expr): MetaM Expr := do
|
||||
let motiveType := Expr.instantiateRev info.getMotiveType mvars
|
||||
let resultantType ← Meta.inferType resultant
|
||||
return replaceForallBody motiveType resultantType
|
||||
|
||||
protected def conduitType (info: RecursorWithMotive) (mvars: Array Expr) (resultant: Expr): MetaM Expr := do
|
||||
let motiveCall := Expr.instantiateRev info.body mvars
|
||||
Meta.mkEq motiveCall resultant
|
||||
|
||||
end RecursorWithMotive
|
||||
|
||||
def getRecursorInformation (recursorType: Expr): Option RecursorWithMotive := do
|
||||
let (args, body) := getForallArgsBody recursorType
|
||||
if ¬ body.isApp then
|
||||
.none
|
||||
let iMotive ← match body.getAppFn with
|
||||
| .bvar iMotive => pure iMotive
|
||||
| _ => .none
|
||||
return {
|
||||
args,
|
||||
body,
|
||||
iMotive,
|
||||
}
|
||||
|
||||
def collectMotiveArguments (forallBody: Expr): SSet Nat :=
|
||||
match forallBody with
|
||||
| .app (.bvar i) _ => SSet.empty.insert i
|
||||
| _ => SSet.empty
|
||||
|
||||
/-- Applies a symbol of the type `∀ (motive: α → Sort u) (a: α)..., (motive α)` -/
|
||||
def motivatedApply (mvarId: MVarId) (recursor: Expr) : MetaM (Array Meta.InductionSubgoal) := mvarId.withContext do
|
||||
mvarId.checkNotAssigned `Pantograph.Tactic.motivatedApply
|
||||
let recursorType ← Meta.inferType recursor
|
||||
let resultant ← mvarId.getType
|
||||
let tag ← mvarId.getTag
|
||||
|
||||
let info ← match getRecursorInformation recursorType with
|
||||
| .some info => pure info
|
||||
| .none => throwError "Recursor return type does not correspond with the invocation of a motive: {← Meta.ppExpr recursorType}"
|
||||
|
||||
let rec go (i: Nat) (prev: Array Expr): MetaM (Array Expr) := do
|
||||
if i ≥ info.nArgs then
|
||||
return prev
|
||||
else
|
||||
let argType := info.args.get! i
|
||||
-- If `argType` has motive references, its goal needs to be placed in it
|
||||
let argType := argType.instantiateRev prev
|
||||
let bvarIndex := info.nArgs - i - 1
|
||||
let argGoal ← if bvarIndex = info.iMotive then
|
||||
let surrogateMotiveType ← info.surrogateMotiveType prev resultant
|
||||
Meta.mkFreshExprSyntheticOpaqueMVar surrogateMotiveType (tag := tag ++ `motive)
|
||||
else
|
||||
Meta.mkFreshExprSyntheticOpaqueMVar argType (tag := .anonymous)
|
||||
let prev := prev ++ [argGoal]
|
||||
go (i + 1) prev
|
||||
termination_by info.nArgs - i
|
||||
let mut newMVars ← go 0 #[]
|
||||
|
||||
-- Create the conduit type which proves the result of the motive is equal to the goal
|
||||
let conduitType ← info.conduitType newMVars resultant
|
||||
let goalConduit ← Meta.mkFreshExprSyntheticOpaqueMVar conduitType (tag := `conduit)
|
||||
mvarId.assign $ ← Meta.mkEqMP goalConduit (mkAppN recursor newMVars)
|
||||
newMVars := newMVars ++ [goalConduit]
|
||||
|
||||
return newMVars.map (λ mvar => { mvarId := mvar.mvarId!})
|
||||
|
||||
def evalMotivatedApply : Elab.Tactic.Tactic := fun stx => Elab.Tactic.withMainContext do
|
||||
let recursor ← Elab.Term.elabTerm (stx := stx) .none
|
||||
let nextGoals ← motivatedApply (← Elab.Tactic.getMainGoal) recursor
|
||||
Elab.Tactic.replaceMainGoal $ nextGoals.toList.map (·.mvarId)
|
||||
|
||||
end Pantograph.Tactic
|
|
@ -1,22 +0,0 @@
|
|||
import Lean
|
||||
|
||||
open Lean
|
||||
|
||||
namespace Pantograph.Tactic
|
||||
|
||||
def noConfuse (mvarId: MVarId) (h: Expr): MetaM Unit := mvarId.withContext do
|
||||
mvarId.checkNotAssigned `Pantograph.Tactic.noConfuse
|
||||
let target ← mvarId.getType
|
||||
let noConfusion ← Meta.mkNoConfusion (target := target) (h := h)
|
||||
|
||||
unless ← Meta.isDefEq (← Meta.inferType noConfusion) target do
|
||||
throwError "invalid noConfuse call: The resultant type {← Meta.ppExpr $ ← Meta.inferType noConfusion} cannot be unified with {← Meta.ppExpr target}"
|
||||
mvarId.assign noConfusion
|
||||
|
||||
def evalNoConfuse: Elab.Tactic.Tactic := λ stx => do
|
||||
let goal ← Elab.Tactic.getMainGoal
|
||||
let h ← goal.withContext $ Elab.Term.elabTerm (stx := stx) .none
|
||||
noConfuse goal h
|
||||
Elab.Tactic.replaceMainGoal []
|
||||
|
||||
end Pantograph.Tactic
|
|
@ -54,9 +54,6 @@ def main (args: List String) := do
|
|||
("Delate", Delate.suite env_default),
|
||||
("Serial", Serial.suite env_default),
|
||||
("Tactic/Assign", Tactic.Assign.suite env_default),
|
||||
("Tactic/Congruence", Tactic.Congruence.suite env_default),
|
||||
("Tactic/Motivated Apply", Tactic.MotivatedApply.suite env_default),
|
||||
("Tactic/No Confuse", Tactic.NoConfuse.suite env_default),
|
||||
("Tactic/Prograde", Tactic.Prograde.suite env_default),
|
||||
]
|
||||
let tests: List (String × IO LSpec.TestSeq) := suites.foldl (λ acc (name, suite) => acc ++ (addPrefix name suite)) []
|
||||
|
|
175
Test/Proofs.lean
175
Test/Proofs.lean
|
@ -543,179 +543,6 @@ def test_calc: TestM Unit := do
|
|||
("h1", "a + b = b + c"), ("h2", "b + c = c + d")] ++ free
|
||||
buildGoal free target userName?
|
||||
|
||||
def test_nat_zero_add: TestM Unit := do
|
||||
let state? ← startProof (.expr "∀ (n: Nat), n + 0 = n")
|
||||
let state0 ← match state? with
|
||||
| .some state => pure state
|
||||
| .none => do
|
||||
addTest $ assertUnreachable "Goal could not parse"
|
||||
return ()
|
||||
let tactic := "intro n"
|
||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[buildGoal [("n", "Nat")] "n + 0 = n"])
|
||||
let recursor := "@Nat.brecOn"
|
||||
let state2 ← match ← state1.tryMotivatedApply (state1.get! 0) (recursor := recursor) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
let [mvarMotive, mvarMajor, mvarInduct, mvarConduit] := state2.goals |
|
||||
fail "Incorrect number of goals"
|
||||
let .num _ major := mvarMajor.name | fail "Incorrect form of mvar id"
|
||||
addTest $ LSpec.check s!"mapply {recursor}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilizeVars) =
|
||||
#[
|
||||
buildNamedGoal mvarMotive.name.toString [("n", "Nat")] "Nat → Prop" (.some "motive"),
|
||||
buildNamedGoal mvarMajor.name.toString [("n", "Nat")] "Nat",
|
||||
buildNamedGoal mvarInduct.name.toString [("n", "Nat")] "∀ (t : Nat), Nat.below t → ?motive t",
|
||||
buildNamedGoal mvarConduit.name.toString [("n", "Nat")] s!"?motive ?m.{major} = (n + 0 = n)" (.some "conduit")
|
||||
])
|
||||
|
||||
let tactic := "exact n"
|
||||
let state3b ← match ← state2.tacticOn (goalId := 1) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state3b.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[])
|
||||
let state2b ← match state3b.continue state2 with
|
||||
| .ok state => pure state
|
||||
| .error e => do
|
||||
addTest $ assertUnreachable e
|
||||
return ()
|
||||
let tactic := "exact (λ x => x + 0 = x)"
|
||||
let state3c ← match ← state2b.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state3c.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[])
|
||||
let state2c ← match state3c.continue state2b with
|
||||
| .ok state => pure state
|
||||
| .error e => do
|
||||
addTest $ assertUnreachable e
|
||||
return ()
|
||||
let tactic := "intro t h"
|
||||
let state3 ← match ← state2c.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state3.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[buildGoal [("n", "Nat"), ("t", "Nat"), ("h", "Nat.below t")] "t + 0 = t"])
|
||||
|
||||
let tactic := "simp"
|
||||
let state3d ← match ← state3.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
let state2d ← match state3d.continue state2c with
|
||||
| .ok state => pure state
|
||||
| .error e => do
|
||||
addTest $ assertUnreachable e
|
||||
return ()
|
||||
let tactic := "rfl"
|
||||
let stateF ← match ← state2d.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← stateF.serializeGoals (options := ← read)) =
|
||||
#[])
|
||||
|
||||
let expr := stateF.mctx.eAssignment.find! stateF.root
|
||||
let (expr, _) := instantiateMVarsCore (mctx := stateF.mctx) (e := expr)
|
||||
addTest $ LSpec.check "(F root)" stateF.rootExpr?.isSome
|
||||
|
||||
def test_nat_zero_add_alt: TestM Unit := do
|
||||
let state? ← startProof (.expr "∀ (n: Nat), n + 0 = n")
|
||||
let state0 ← match state? with
|
||||
| .some state => pure state
|
||||
| .none => do
|
||||
addTest $ assertUnreachable "Goal could not parse"
|
||||
return ()
|
||||
let tactic := "intro n"
|
||||
let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state1.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[buildGoal [("n", "Nat")] "n + 0 = n"])
|
||||
let recursor := "@Nat.brecOn"
|
||||
let state2 ← match ← state1.tryMotivatedApply (state1.get! 0) (recursor := recursor) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
let [mvarMotive, mvarMajor, mvarInduct, mvarConduit] := state2.goals |
|
||||
fail "Incorrect number of goals"
|
||||
let .num _ major := mvarMajor.name | fail "Incorrect form of mvar id"
|
||||
addTest $ LSpec.check s!"mapply {recursor}" ((← state2.serializeGoals (options := ← read)).map (·.devolatilizeVars) =
|
||||
#[
|
||||
buildNamedGoal mvarMotive.name.toString [("n", "Nat")] "Nat → Prop" (.some "motive"),
|
||||
buildNamedGoal mvarMajor.name.toString [("n", "Nat")] "Nat",
|
||||
buildNamedGoal mvarInduct.name.toString [("n", "Nat")] "∀ (t : Nat), Nat.below t → ?motive t",
|
||||
buildNamedGoal mvarConduit.name.toString [("n", "Nat")] s!"?motive ?m.{major} = (n + 0 = n)" (.some "conduit")
|
||||
])
|
||||
|
||||
let tactic := "intro x"
|
||||
let state3m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
addTest $ LSpec.check tactic ((← state3m.serializeGoals (options := ← read)).map (·.devolatilize) =
|
||||
#[buildGoal [("n", "Nat"), ("x", "Nat")] "Prop" (.some "motive")])
|
||||
let tactic := "apply Eq"
|
||||
let state3m2 ← match ← state3m.tacticOn (goalId := 0) (tactic := tactic) with
|
||||
| .success state => pure state
|
||||
| other => do
|
||||
addTest $ assertUnreachable $ other.toString
|
||||
return ()
|
||||
let [eqL, eqR, eqT] := state3m2.goals | fail "Incorrect number of goals"
|
||||
let [_motive, _major, _step, conduit] := state2.goals | panic! "Goals conflict"
|
||||
let state2b ← match state3m2.resume [conduit] with
|
||||
| .ok state => pure state
|
||||
| .error e => do
|
||||
addTest $ assertUnreachable e
|
||||
return ()
|
||||
|
||||
let cNatAdd := "(:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat))"
|
||||
let cNat0 := "((:c OfNat.ofNat) (:c Nat) (:lit 0) ((:c instOfNatNat) (:lit 0)))"
|
||||
let fvN ← state2b.withContext conduit do
|
||||
let lctx ← getLCtx
|
||||
pure $ lctx.getFVarIds.get! 0 |>.name
|
||||
let conduitRight := s!"((:c Eq) (:c Nat) ({cNatAdd} (:fv {fvN}) {cNat0}) (:fv {fvN}))"
|
||||
let substOf (mvarId: MVarId) := s!"(:subst (:mv {mvarId.name}) (:fv {fvN}) (:mv {mvarMajor}))"
|
||||
let .num _ nL := eqL.name | fail "Incorrect form of mvar id"
|
||||
let .num _ nR := eqR.name | fail "Incorrect form of mvar id"
|
||||
let nL' := nL + 4
|
||||
let nR' := nR + 5
|
||||
addTest $ LSpec.check "resume" ((← state2b.serializeGoals (options := { ← read with printExprAST := true })) =
|
||||
#[
|
||||
{
|
||||
name := mvarConduit.name.toString,
|
||||
userName? := .some "conduit",
|
||||
target := {
|
||||
pp? := .some s!"(?m.{nL'} ?m.{major} = ?m.{nR'} ?m.{major}) = (n + 0 = n)",
|
||||
sexp? := .some s!"((:c Eq) (:sort 0) ((:c Eq) {substOf eqT} {substOf eqL} {substOf eqR}) {conduitRight})",
|
||||
},
|
||||
vars := #[{
|
||||
name := fvN.toString,
|
||||
userName := "n",
|
||||
type? := .some { pp? := .some "Nat", sexp? := .some "(:c Nat)" },
|
||||
}],
|
||||
}
|
||||
])
|
||||
|
||||
def test_tactic_failure_unresolved_goals : TestM Unit := do
|
||||
let state? ← startProof (.expr "∀ (p : Nat → Prop), ∃ (x : Nat), p (0 + x + 0)")
|
||||
let state0 ← match state? with
|
||||
|
@ -778,8 +605,6 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
|||
("Or.comm", test_or_comm),
|
||||
("conv", test_conv),
|
||||
("calc", test_calc),
|
||||
("Nat.zero_add", test_nat_zero_add),
|
||||
("Nat.zero_add alt", test_nat_zero_add_alt),
|
||||
("tactic failure with unresolved goals", test_tactic_failure_unresolved_goals),
|
||||
("tactic failure with synthesize placeholder", test_tactic_failure_synthesize_placeholder),
|
||||
]
|
||||
|
|
|
@ -1,5 +1,2 @@
|
|||
import Test.Tactic.Assign
|
||||
import Test.Tactic.Congruence
|
||||
import Test.Tactic.MotivatedApply
|
||||
import Test.Tactic.NoConfuse
|
||||
import Test.Tactic.Prograde
|
||||
|
|
|
@ -1,88 +0,0 @@
|
|||
import LSpec
|
||||
import Lean
|
||||
import Test.Common
|
||||
|
||||
open Lean
|
||||
open Pantograph
|
||||
|
||||
namespace Pantograph.Test.Tactic.Congruence
|
||||
|
||||
def test_congr_arg_list : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ {α} (l1 l2 : List α) (h: l1 = l2) => l1.reverse = l2.reverse"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let newGoals ← runTacticOnMVar Tactic.evalCongruenceArg target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.30"),
|
||||
(`a₁, "?α"),
|
||||
(`a₂, "?α"),
|
||||
(`f, "?α → List α"),
|
||||
(`h, "?a₁ = ?a₂"),
|
||||
(`conduit, "(?f ?a₁ = ?f ?a₂) = (l1.reverse = l2.reverse)"),
|
||||
])
|
||||
let f := newGoals.get! 3
|
||||
let h := newGoals.get! 4
|
||||
let c := newGoals.get! 5
|
||||
let results ← Meta.withAssignableSyntheticOpaque do f.apply (← parseSentence "List.reverse")
|
||||
addTest $ LSpec.check "apply" (results.length = 0)
|
||||
addTest $ LSpec.check "h" ((← exprToStr $ ← h.getType) = "?a₁ = ?a₂")
|
||||
addTest $ LSpec.check "conduit" ((← exprToStr $ ← c.getType) = "(List.reverse ?a₁ = List.reverse ?a₂) = (l1.reverse = l2.reverse)")
|
||||
def test_congr_arg : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n m: Nat) (h: n = m) => n * n = m * m"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let newGoals ← runTacticOnMVar Tactic.evalCongruenceArg target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.73"),
|
||||
(`a₁, "?α"),
|
||||
(`a₂, "?α"),
|
||||
(`f, "?α → Nat"),
|
||||
(`h, "?a₁ = ?a₂"),
|
||||
(`conduit, "(?f ?a₁ = ?f ?a₂) = (n * n = m * m)"),
|
||||
])
|
||||
def test_congr_fun : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n m: Nat) => (n + m) + (n + m) = (n + m) * 2"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let newGoals ← runTacticOnMVar Tactic.evalCongruenceFun target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.165"),
|
||||
(`f₁, "?α → Nat"),
|
||||
(`f₂, "?α → Nat"),
|
||||
(`h, "?f₁ = ?f₂"),
|
||||
(`a, "?α"),
|
||||
(`conduit, "(?f₁ ?a = ?f₂ ?a) = (n + m + (n + m) = (n + m) * 2)"),
|
||||
])
|
||||
def test_congr : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (a b: Nat) => a = b"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let newGoals ← runTacticOnMVar Tactic.evalCongruence target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ x => mvarUserNameAndType x)) =
|
||||
[
|
||||
(`α, "Sort ?u.10"),
|
||||
(`f₁, "?α → Nat"),
|
||||
(`f₂, "?α → Nat"),
|
||||
(`a₁, "?α"),
|
||||
(`a₂, "?α"),
|
||||
(`h₁, "?f₁ = ?f₂"),
|
||||
(`h₂, "?a₁ = ?a₂"),
|
||||
(`conduit, "(?f₁ ?a₁ = ?f₂ ?a₂) = (a = b)"),
|
||||
])
|
||||
|
||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||
[
|
||||
("congrArg List.reverse", test_congr_arg_list),
|
||||
("congrArg", test_congr_arg),
|
||||
("congrFun", test_congr_fun),
|
||||
("congr", test_congr),
|
||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
||||
|
||||
end Pantograph.Test.Tactic.Congruence
|
|
@ -1,113 +0,0 @@
|
|||
import LSpec
|
||||
import Lean
|
||||
import Test.Common
|
||||
|
||||
open Lean
|
||||
open Pantograph
|
||||
|
||||
namespace Pantograph.Test.Tactic.MotivatedApply
|
||||
|
||||
def test_type_extract : TestT Elab.TermElabM Unit := do
|
||||
let recursor ← parseSentence "@Nat.brecOn"
|
||||
let recursorType ← Meta.inferType recursor
|
||||
addTest $ LSpec.check "recursorType" ("{motive : Nat → Sort ?u.1} → (t : Nat) → ((t : Nat) → Nat.below t → motive t) → motive t" =
|
||||
(← exprToStr recursorType))
|
||||
let info ← match Tactic.getRecursorInformation recursorType with
|
||||
| .some info => pure info
|
||||
| .none => throwError "Failed to extract recursor info"
|
||||
addTest $ LSpec.check "iMotive" (info.iMotive = 2)
|
||||
let motiveType := info.getMotiveType
|
||||
addTest $ LSpec.check "motiveType" ("Nat → Sort ?u.1" =
|
||||
(← exprToStr motiveType))
|
||||
|
||||
def test_nat_brec_on : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n t: Nat) => n + 0 = n"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "@Nat.brecOn")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let tactic := Tactic.evalMotivatedApply recursor
|
||||
let newGoals ← runTacticOnMVar tactic target.mvarId!
|
||||
let test := LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
|
||||
[
|
||||
"Nat → Prop",
|
||||
"Nat",
|
||||
"∀ (t : Nat), Nat.below t → ?motive t",
|
||||
"?motive ?m.74 = (n + 0 = n)",
|
||||
])
|
||||
addTest test
|
||||
|
||||
def test_list_brec_on : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ {α : Type} (l: List α) => l ++ [] = [] ++ l"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "@List.brecOn")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let tactic := Tactic.evalMotivatedApply recursor
|
||||
let newGoals ← runTacticOnMVar tactic target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
|
||||
[
|
||||
"Type ?u.90",
|
||||
"List ?m.92 → Prop",
|
||||
"List ?m.92",
|
||||
"∀ (t : List ?m.92), List.below t → ?motive t",
|
||||
"?motive ?m.94 = (l ++ [] = [] ++ l)",
|
||||
])
|
||||
|
||||
def test_partial_motive_instantiation : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n t: Nat) => n + 0 = n"
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "@Nat.brecOn")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let tactic := Tactic.evalMotivatedApply recursor
|
||||
let newGoals ← runTacticOnMVar tactic target.mvarId!
|
||||
let majorId := 74
|
||||
addTest $ (LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) =
|
||||
[
|
||||
"Nat → Prop",
|
||||
"Nat",
|
||||
"∀ (t : Nat), Nat.below t → ?motive t",
|
||||
s!"?motive ?m.{majorId} = (n + 0 = n)",
|
||||
]))
|
||||
let [motive, major, step, conduit] := newGoals | panic! "Incorrect goal number"
|
||||
addTest $ (LSpec.check "goal name" (major.name.toString = s!"_uniq.{majorId}"))
|
||||
|
||||
-- Assign motive to `λ x => x + _`
|
||||
let motive_assign ← parseSentence "λ (x: Nat) => @Nat.add x + 0 = _"
|
||||
motive.assign motive_assign
|
||||
|
||||
addTest $ ← conduit.withContext do
|
||||
let t := toString (← Meta.ppExpr $ ← conduit.getType)
|
||||
return LSpec.check "conduit" (t = s!"(Nat.add ?m.{majorId} + 0 = ?m.150 ?m.{majorId}) = (n + 0 = n)")
|
||||
|
||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||
[
|
||||
("type_extract", test_type_extract),
|
||||
("Nat.brecOn", test_nat_brec_on),
|
||||
("List.brecOn", test_list_brec_on),
|
||||
("Nat.brecOn partial motive instantiation", test_partial_motive_instantiation),
|
||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
||||
|
||||
end Pantograph.Test.Tactic.MotivatedApply
|
|
@ -1,72 +0,0 @@
|
|||
import LSpec
|
||||
import Lean
|
||||
import Test.Common
|
||||
|
||||
open Lean
|
||||
open Pantograph
|
||||
|
||||
namespace Pantograph.Test.Tactic.NoConfuse
|
||||
|
||||
def test_nat : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n: Nat) (h: 0 = n + 1) => False"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "h")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let tactic := Tactic.evalNoConfuse recursor
|
||||
let newGoals ← runTacticOnMVar tactic target.mvarId!
|
||||
addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
|
||||
|
||||
def test_nat_fail : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (n: Nat) (h: n = n) => False"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "h")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
try
|
||||
let tactic := Tactic.evalNoConfuse recursor
|
||||
let _ ← runTacticOnMVar tactic target.mvarId!
|
||||
addTest $ assertUnreachable "Tactic should fail"
|
||||
catch _ =>
|
||||
addTest $ LSpec.check "Tactic should fail" true
|
||||
|
||||
def test_list : TestT Elab.TermElabM Unit := do
|
||||
let expr := "λ (l: List Nat) (h: [] = 1 :: l) => False"
|
||||
let expr ← parseSentence expr
|
||||
Meta.lambdaTelescope expr $ λ _ body => do
|
||||
let recursor ← match Parser.runParserCategory
|
||||
(env := ← MonadEnv.getEnv)
|
||||
(catName := `term)
|
||||
(input := "h")
|
||||
(fileName := ← getFileName) with
|
||||
| .ok syn => pure syn
|
||||
| .error error => throwError "Failed to parse: {error}"
|
||||
-- Apply the tactic
|
||||
let target ← Meta.mkFreshExprSyntheticOpaqueMVar body
|
||||
let tactic := Tactic.evalNoConfuse recursor
|
||||
let newGoals ← runTacticOnMVar tactic target.mvarId!
|
||||
addTest $ LSpec.check "goals"
|
||||
((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = [])
|
||||
|
||||
def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
|
||||
[
|
||||
("Nat", test_nat),
|
||||
("Nat fail", test_nat_fail),
|
||||
("List", test_list),
|
||||
] |>.map (λ (name, t) => (name, runTestTermElabM env t))
|
||||
|
||||
end Pantograph.Test.Tactic.NoConfuse
|
48
flake.lock
48
flake.lock
|
@ -39,7 +39,9 @@
|
|||
"lean4-nix": {
|
||||
"inputs": {
|
||||
"flake-parts": "flake-parts_2",
|
||||
"nixpkgs": "nixpkgs"
|
||||
"nixpkgs": [
|
||||
"nixpkgs"
|
||||
]
|
||||
},
|
||||
"locked": {
|
||||
"lastModified": 1739073990,
|
||||
|
@ -55,35 +57,18 @@
|
|||
"type": "github"
|
||||
}
|
||||
},
|
||||
"lspec": {
|
||||
"flake": false,
|
||||
"locked": {
|
||||
"lastModified": 1728279187,
|
||||
"narHash": "sha256-ZMqbvCqR/gHXRuIkuo7b0Yp9N1vOQR7xnrcy/SeIBoQ=",
|
||||
"owner": "argumentcomputer",
|
||||
"repo": "LSpec",
|
||||
"rev": "504a8cecf8da601b9466ac727aebb6b511aae4ab",
|
||||
"type": "github"
|
||||
},
|
||||
"original": {
|
||||
"owner": "argumentcomputer",
|
||||
"ref": "504a8cecf8da601b9466ac727aebb6b511aae4ab",
|
||||
"repo": "LSpec",
|
||||
"type": "github"
|
||||
}
|
||||
},
|
||||
"nixpkgs": {
|
||||
"locked": {
|
||||
"lastModified": 1728500571,
|
||||
"narHash": "sha256-dOymOQ3AfNI4Z337yEwHGohrVQb4yPODCW9MDUyAc4w=",
|
||||
"lastModified": 1741332913,
|
||||
"narHash": "sha256-ri1e8ZliWS3Jnp9yqpKApHaOo7KBN33W8ECAKA4teAQ=",
|
||||
"owner": "nixos",
|
||||
"repo": "nixpkgs",
|
||||
"rev": "d51c28603def282a24fa034bcb007e2bcb5b5dd0",
|
||||
"rev": "20755fa05115c84be00b04690630cb38f0a203ad",
|
||||
"type": "github"
|
||||
},
|
||||
"original": {
|
||||
"owner": "nixos",
|
||||
"ref": "nixos-24.05",
|
||||
"ref": "nixos-24.11",
|
||||
"repo": "nixpkgs",
|
||||
"type": "github"
|
||||
}
|
||||
|
@ -112,28 +97,11 @@
|
|||
"url": "https://github.com/NixOS/nixpkgs/archive/fb192fec7cc7a4c26d51779e9bab07ce6fa5597a.tar.gz"
|
||||
}
|
||||
},
|
||||
"nixpkgs_2": {
|
||||
"locked": {
|
||||
"lastModified": 1735563628,
|
||||
"narHash": "sha256-OnSAY7XDSx7CtDoqNh8jwVwh4xNL/2HaJxGjryLWzX8=",
|
||||
"owner": "nixos",
|
||||
"repo": "nixpkgs",
|
||||
"rev": "b134951a4c9f3c995fd7be05f3243f8ecd65d798",
|
||||
"type": "github"
|
||||
},
|
||||
"original": {
|
||||
"owner": "nixos",
|
||||
"ref": "nixos-24.05",
|
||||
"repo": "nixpkgs",
|
||||
"type": "github"
|
||||
}
|
||||
},
|
||||
"root": {
|
||||
"inputs": {
|
||||
"flake-parts": "flake-parts",
|
||||
"lean4-nix": "lean4-nix",
|
||||
"lspec": "lspec",
|
||||
"nixpkgs": "nixpkgs_2"
|
||||
"nixpkgs": "nixpkgs"
|
||||
}
|
||||
}
|
||||
},
|
||||
|
|
16
flake.nix
16
flake.nix
|
@ -2,12 +2,11 @@
|
|||
description = "Pantograph";
|
||||
|
||||
inputs = {
|
||||
nixpkgs.url = "github:nixos/nixpkgs/nixos-24.05";
|
||||
nixpkgs.url = "github:nixos/nixpkgs/nixos-24.11";
|
||||
flake-parts.url = "github:hercules-ci/flake-parts";
|
||||
lean4-nix.url = "github:lenianiva/lean4-nix";
|
||||
lspec = {
|
||||
url = "github:argumentcomputer/LSpec?ref=504a8cecf8da601b9466ac727aebb6b511aae4ab";
|
||||
flake = false;
|
||||
lean4-nix = {
|
||||
url = "github:lenianiva/lean4-nix";
|
||||
inputs.nixpkgs.follows = "nixpkgs";
|
||||
};
|
||||
};
|
||||
|
||||
|
@ -16,7 +15,6 @@
|
|||
nixpkgs,
|
||||
flake-parts,
|
||||
lean4-nix,
|
||||
lspec,
|
||||
...
|
||||
}:
|
||||
flake-parts.lib.mkFlake {inherit inputs;} {
|
||||
|
@ -37,10 +35,12 @@
|
|||
inherit system;
|
||||
overlays = [(lean4-nix.readToolchainFile ./lean-toolchain)];
|
||||
};
|
||||
manifest = pkgs.lib.importJSON ./lake-manifest.json;
|
||||
manifest-lspec = builtins.head manifest;
|
||||
lspecLib = pkgs.lean.buildLeanPackage {
|
||||
name = "LSpec";
|
||||
roots = ["Main" "LSpec"];
|
||||
src = "${lspec}";
|
||||
roots = ["LSpec"];
|
||||
src = builtins.fetchGit { inherit (manifest-lspec) url rev; };
|
||||
};
|
||||
project = pkgs.lean.buildLeanPackage {
|
||||
name = "Pantograph";
|
||||
|
|
Loading…
Reference in New Issue