Node IDs #
Equations
- Aesop.instInhabitedGoalId = { default := { toNat := default } }
Instances For
Equations
- Aesop.GoalId.dummy = { toNat := 1000000000000000 }
Instances For
Equations
- Aesop.GoalId.instLT = { lt := fun (n m : Aesop.GoalId) => n.toNat < m.toNat }
Equations
- n.instDecidableRelLt m = inferInstanceAs (Decidable (n.toNat < m.toNat))
Equations
- Aesop.GoalId.instToString = { toString := fun (n : Aesop.GoalId) => toString n.toNat }
Equations
- Aesop.GoalId.instHashable = { hash := fun (n : Aesop.GoalId) => hash n.toNat }
Rule Application IDs #
Equations
- Aesop.instInhabitedRappId = { default := { toNat := default } }
Instances For
Equations
- Aesop.RappId.dummy = { toNat := 1000000000000000 }
Instances For
Equations
- Aesop.RappId.instLT = { lt := fun (n m : Aesop.RappId) => n.toNat < m.toNat }
Equations
- n.instDecidableRelLt m = inferInstanceAs (Decidable (n.toNat < m.toNat))
Equations
- Aesop.RappId.instToString = { toString := fun (n : Aesop.RappId) => toString n.toNat }
Equations
- Aesop.RappId.instHashable = { hash := fun (n : Aesop.RappId) => hash n.toNat }
Iterations #
@[inline]
Equations
Instances For
Equations
Instances For
Equations
Equations
Equations
- Aesop.Iteration.instDecidableRelLt = inferInstanceAs (DecidableRel fun (x x_1 : Nat) => x < x_1)
Equations
- Aesop.Iteration.instDecidableRelLe = inferInstanceAs (DecidableRel fun (x x_1 : Nat) => x ≤ x_1)
The Tree #
At each point during the search, every node of the tree (goal, rapp or mvar cluster) is in one of these states:
- proven: the node is proven.
- unprovable: the node is unprovable, i.e. it will never be proven regardless of any future expansions that might be performed.
- unknown: neither of the above.
Every node starts in the unknown state and may later become either proven or
unprovable. After this, the state does not change any more.
- unknown: Aesop.NodeState
- proven: Aesop.NodeState
- unprovable: Aesop.NodeState
Instances For
Equations
- Aesop.instInhabitedNodeState = { default := Aesop.NodeState.unknown }
Equations
- Aesop.instBEqNodeState = { beq := Aesop.beqNodeState✝ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- x.isUnknown = match x with | Aesop.NodeState.unknown => true | x => false
Instances For
Equations
- x.isProven = match x with | Aesop.NodeState.proven => true | x => false
Instances For
Equations
- x.isUnprovable = match x with | Aesop.NodeState.unprovable => true | x => false
Instances For
Equations
- x.isIrrelevant = match x with | Aesop.NodeState.proven => true | Aesop.NodeState.unprovable => true | Aesop.NodeState.unknown => false
Instances For
Equations
- x.toEmoji = match x with | Aesop.NodeState.proven => Aesop.nodeProvedEmoji | Aesop.NodeState.unprovable => Aesop.nodeUnprovableEmoji | Aesop.NodeState.unknown => Aesop.nodeUnknownEmoji
Instances For
A refinement of the NodeState, distinguishing between goals proven during
normalisation and goals proven by a child rule application.
- unknown: Aesop.GoalState
- provenByRuleApplication: Aesop.GoalState
- provenByNormalization: Aesop.GoalState
- unprovable: Aesop.GoalState
Instances For
Equations
- Aesop.instInhabitedGoalState = { default := Aesop.GoalState.unknown }
Equations
- Aesop.instBEqGoalState = { beq := Aesop.beqGoalState✝ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- x.isProvenByRuleApplication = match x with | Aesop.GoalState.provenByRuleApplication => true | x => false
Instances For
Equations
- x.isProvenByNormalization = match x with | Aesop.GoalState.provenByNormalization => true | x => false
Instances For
Equations
- x.isProven = match x with | Aesop.GoalState.provenByRuleApplication => true | Aesop.GoalState.provenByNormalization => true | x => false
Instances For
Equations
- x.isUnprovable = match x with | Aesop.GoalState.unprovable => true | x => false
Instances For
Equations
- x.isUnknown = match x with | Aesop.GoalState.unknown => true | x => false
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- s.isIrrelevant = s.toNodeState.isIrrelevant
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
- notNormal: Aesop.NormalizationState
- normal: Lean.MVarId → Lean.Meta.SavedState → Array (Aesop.DisplayRuleName × Option (Array Aesop.Script.LazyStep)) → Aesop.NormalizationState
- provenByNormalization: Lean.Meta.SavedState → Array (Aesop.DisplayRuleName × Option (Array Aesop.Script.LazyStep)) → Aesop.NormalizationState
Instances For
Equations
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
A goal G can be added to the tree for three reasons:
- Gwas produced by its parent rule as a subgoal. This is the most common reason.
- Gwas copied because it contains some metavariables which were assigned by its parent rule. In this case, we record goal of which- Gis a copy. We also record the representative of the equivalence class of goals which are copies of each other. E.g. if goal- 1is copied to goal- 2and goal- 2is copied to goal- 3, they are all part of the equivalence class with representative- 1.
- subgoal: Aesop.GoalOrigin
- copied: Aesop.GoalId → Aesop.GoalId → Aesop.GoalOrigin
- droppedMVar: Aesop.GoalOrigin
Instances For
Equations
- Aesop.instInhabitedGoalOrigin = { default := Aesop.GoalOrigin.subgoal }
Equations
- x.originalGoalId? = match x with | Aesop.GoalOrigin.copied from rep => some rep | x => none
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
- id : Aesop.GoalId
- parent : IO.Ref MVarCluster
- origin : Aesop.GoalOrigin
- depth : Nat
- state : Aesop.GoalState
- isIrrelevant : Bool
- isForcedUnprovable : Bool
- preNormGoal : Lean.MVarId
- normalizationState : Aesop.NormalizationState
- mvars : Aesop.UnorderedArraySet Lean.MVarId
- successProbability : Aesop.Percent
- addedInIteration : Aesop.Iteration
- lastExpandedInIteration : Aesop.Iteration
- unsafeRulesSelected : Bool
- unsafeQueue : Aesop.UnsafeQueue
- failedRapps : Array Aesop.RegularRule
Instances For
instance
Aesop.instNonemptyGoalData :
∀ {Rapp MVarCluster : Type} [inst : Nonempty MVarCluster], Nonempty (Aesop.GoalData Rapp MVarCluster)
Equations
- ⋯ = ⋯
instance
Aesop.instInhabitedMVarClusterData :
{a a_1 : Type} → Inhabited (Aesop.MVarClusterData a a_1)
Equations
- Aesop.instInhabitedMVarClusterData = { default := { parent? := default, goals := default, isIrrelevant := default, state := default } }
- id : Aesop.RappId
- parent : IO.Ref Goal
- state : Aesop.NodeState
- isIrrelevant : Bool
- appliedRule : Aesop.RegularRule
- scriptSteps? : Option (Array Aesop.Script.LazyStep)
- originalSubgoals : Array Lean.MVarId
- successProbability : Aesop.Percent
- metaState : Lean.Meta.SavedState
- introducedMVars : Aesop.UnorderedArraySet Lean.MVarId
- assignedMVars : Aesop.UnorderedArraySet Lean.MVarId
Instances For
instance
Aesop.instNonemptyRappData :
∀ {Goal : Type} [inst : Nonempty Goal] {MVarCluster : Type}, Nonempty (Aesop.RappData Goal MVarCluster)
Equations
- ⋯ = ⋯
Instances For
Instances For
Instances For
- Goal : Type
- Rapp : Type
- MVarCluster : Type
- introGoal : Aesop.GoalData self.Rapp self.MVarCluster → self.Goal
- elimGoal : self.Goal → Aesop.GoalData self.Rapp self.MVarCluster
- introRapp : Aesop.RappData self.Goal self.MVarCluster → self.Rapp
- elimRapp : self.Rapp → Aesop.RappData self.Goal self.MVarCluster
- introMVarCluster : Aesop.MVarClusterData self.Goal self.Rapp → self.MVarCluster
- elimMVarCluster : self.MVarCluster → Aesop.MVarClusterData self.Goal self.Rapp
Instances For
Equations
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[reducible, inline]
Equations
Instances For
@[reducible, inline]
Equations
Instances For
@[reducible, inline]
Equations
Instances For
Equations
@[inline]
def
Aesop.MVarCluster.modify
(f : Aesop.MVarClusterData Aesop.Goal Aesop.Rapp → Aesop.MVarClusterData Aesop.Goal Aesop.Rapp)
(c : Aesop.MVarCluster)
 :
Equations
- Aesop.MVarCluster.modify f c = Aesop.MVarCluster.mk (f c.elim)
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Goal.modify
(f : Aesop.GoalData Aesop.Rapp Aesop.MVarCluster → Aesop.GoalData Aesop.Rapp Aesop.MVarCluster)
(g : Aesop.Goal)
 :
Equations
- Aesop.Goal.modify f g = Aesop.Goal.mk (f g.elim)
Instances For
@[inline]
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Goal.setNormalizationState
(normalizationState : Aesop.NormalizationState)
(g : Aesop.Goal)
 :
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Goal.setLastExpandedInIteration
(lastExpandedInIteration : Aesop.Iteration)
(g : Aesop.Goal)
 :
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Aesop.Goal.instBEq = { beq := fun (g₁ g₂ : Aesop.Goal) => g₁.id == g₂.id }
Equations
- Aesop.Goal.instHashable = { hash := fun (g : Aesop.Goal) => hash g.id }
@[inline]
def
Aesop.Rapp.modify
(f : Aesop.RappData Aesop.Goal Aesop.MVarCluster → Aesop.RappData Aesop.Goal Aesop.MVarCluster)
(r : Aesop.Rapp)
 :
Equations
- Aesop.Rapp.modify f r = Aesop.Rapp.mk (f r.elim)
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Rapp.setScriptSteps?
(scriptSteps? : Option (Array Aesop.Script.LazyStep))
(r : Aesop.Rapp)
 :
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Rapp.setIntroducedMVars
(introducedMVars : Aesop.UnorderedArraySet Lean.MVarId)
(r : Aesop.Rapp)
 :
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
def
Aesop.Rapp.setAssignedMVars
(assignedMVars : Aesop.UnorderedArraySet Lean.MVarId)
(r : Aesop.Rapp)
 :
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Aesop.Rapp.instBEq = { beq := fun (r₁ r₂ : Aesop.Rapp) => r₁.id == r₂.id }
Equations
- Aesop.Rapp.instHashable = { hash := fun (r : Aesop.Rapp) => hash r.id }
Miscellaneous Queries #
Instances For
@[inline]
Equations
- g.postNormGoalAndMetaState? = match g.normalizationState with | Aesop.NormalizationState.normal postGoal postState script => some (postGoal, postState) | x => none
Instances For
Equations
- g.postNormGoal? = Option.map (fun (x : Lean.MVarId × Lean.Meta.SavedState) => x.fst) g.postNormGoalAndMetaState?
Instances For
Equations
- g.currentGoal = g.postNormGoal?.getD g.preNormGoal
Instances For
Equations
- g.parentRapp? = do let __do_lift ← ST.Ref.get g.parent pure __do_lift.parent?
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- g.safeRapps = Array.filterM (fun (rref : Aesop.RappRef) => do let __do_lift ← ST.Ref.get rref pure __do_lift.isSafe) g.children 0
Instances For
Equations
- g.hasSafeRapp = Array.anyM (fun (rref : Aesop.RappRef) => do let __do_lift ← ST.Ref.get rref pure __do_lift.isSafe) g.children 0
Instances For
Equations
- g.isUnsafeExhausted = (g.unsafeRulesSelected && Subarray.isEmpty g.unsafeQueue)
Instances For
Instances For
Instances For
Equations
- g.hasProvableRapp = Array.anyM (fun (r : Aesop.RappRef) => do let __do_lift ← ST.Ref.get r pure !__do_lift.state.isUnprovable) g.children 0
Instances For
Equations
- g.firstProvenRapp? = g.children.findSomeM? fun (rref : Aesop.RappRef) => do let __do_lift ← ST.Ref.get rref pure (if __do_lift.state.isProven = true then some rref else none)
Instances For
Instances For
Equations
- g.priority = g.successProbability * Aesop.unificationGoalPenalty ^ g.mvars.size
Instances For
Equations
- g.isNormal = g.normalizationState.isNormal
Instances For
Equations
- g.originalGoalId = g.origin.originalGoalId?.getD g.id
Instances For
Instances For
Instances For
Equations
- r.parentPostNormMetaState rootMetaState = do let __do_lift ← ST.Ref.get r.parent __do_lift.parentMetaState rootMetaState
Instances For
def
Aesop.Rapp.foldSubgoalsM
{m : Type → Type}
{σ : Type}
[Monad m]
[MonadLiftT (ST IO.RealWorld) m]
(init : σ)
(f : σ → Aesop.GoalRef → m σ)
(r : Aesop.Rapp)
 :
m σ
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Aesop.Rapp.forSubgoalsM
{m : Type → Type}
[Monad m]
[MonadLiftT (ST IO.RealWorld) m]
(f : Aesop.GoalRef → m Unit)
(r : Aesop.Rapp)
 :
m Unit
Equations
- Aesop.Rapp.forSubgoalsM f r = Array.forM (fun (cref : Aesop.MVarClusterRef) => do let __do_lift ← ST.Ref.get cref Array.forM f __do_lift.goals 0) r.children 0
Instances For
def
Aesop.Rapp.subgoals
{m : Type → Type}
[Monad m]
[MonadLiftT (ST IO.RealWorld) m]
(r : Aesop.Rapp)
 :
m (Array Aesop.GoalRef)
Equations
- r.subgoals = Aesop.Rapp.foldSubgoalsM #[] (fun (subgoals : Array Aesop.GoalRef) (gref : Aesop.GoalRef) => pure (subgoals.push gref)) r
Instances For
Equations
- r.depth = do let __do_lift ← ST.Ref.get r.parent pure __do_lift.depth
Instances For
Equations
- c.provenGoal? = c.goals.findM? fun (gref : Aesop.GoalRef) => do let __do_lift ← ST.Ref.get gref pure __do_lift.state.isProven