Simproc for ∃ a', ... ∧ a' = a ∧ ... #
This module implements the existsAndEq simproc, which triggers on goals of the form ∃ a, P.
It checks whether P allows only one possible value for a, and if so, substitutes it, eliminating
the leading quantifier.
The procedure traverses the body, branching at each ∧ and entering existential quantifiers,
searching for a subexpression of the form a = a' or a' = a for a' that is independent of a.
If such an expression is found, all occurrences of a are replaced with a'. If a' depends on
variables bound by existential quantifiers, those quantifiers are moved outside.
For example, ∃ a, p a ∧ ∃ b, a = f b ∧ q b will be rewritten as ∃ b, p (f b) ∧ q b.
Equations
- ExistsAndEq.instBEqGoTo.beq x✝ y✝ = (x✝.ctorIdx == y✝.ctorIdx)
Instances For
Equations
Equations
Instances For
Qq-fied version of Expr. Here, we use it to store free variables introduced when unpacking
existential quantifiers.
Equations
- ExistsAndEq.VarQ = ((u : Lean.Level) × (α : Q(Sort u)) × Q(«$α»))
Instances For
Qq-fied version of Expr proving some P : Prop.
Equations
- ExistsAndEq.HypQ = ((P : Q(Prop)) × Q(«$P»))
Instances For
Constructs ∃ f₁ f₂ ... fₙ, body, where [f₁, ..., fₙ] = fvars.
Equations
- One or more equations did not get rendered due to their size.
- ExistsAndEq.mkNestedExists [] body = pure body
Instances For
Finds a Path for findEq. It leads to a subexpression a = a' or a' = a, where
a' doesn't contain the free variable a.
This is a fast version that quickly returns none when the simproc
is not applicable.
Given P : Prop and a : α, traverses the expression P to find a subexpression of
the form a = a' or a' = a for some a'. It branches at each And and walks into
existential quantifiers.
Returns a tuple (fvars, lctx, P', a'), where:
fvarsis a list of all variables bound by existential quantifiers along the path.lctxis the local context containing all these free variables.P'isPwith all existential quantifiers along the path removed, and corresponding bound variables replaced withfvars.a'is the expression found that must be equal toa. It may contain free variables fromfvars.
Equations
- ExistsAndEq.findEq a P path = ExistsAndEq.findEq.go a P path
Instances For
Recursive part of findEq.
When P = ∃ f₁ ... fₙ, body, where exs = [f₁, ..., fₙ], this function takes
act : body → goal and proves P → goal using Exists.elim.
Example:
exs = []: act h
exs = [b]:
P := ∃ b, body
Exists.elim h (fun b hb ↦ act hb)
exs = [b, c]:
P := ∃ b c, body
Exists.elim h (fun b hb ↦
Exists.elim hb (fun c hc ↦ act hc)
)
...
Instances For
Generates a proof of P' → ∃ a, p a. We assume that fvars = [f₁, ..., fₙ] are free variables
and P' = ∃ f₁ ... fₙ, newBody, and path leads to a = a' in ∃ a, p a.
The proof follows the following structure:
example {α β : Type} (f : β → α) {p : α → Prop} :
(∃ b, p (f b) ∧ f b = f b) → (∃ a, p a ∧ ∃ b, a = f b) := by
-- withLocalDeclQ
intro h
-- withNestedExistsElim : we unpack all quantifiers in `P` to get `h : newBody`.
refine h.elim (fun b h ↦ ?_)
-- use `a'` in the leading existential quantifier
refine Exists.intro (f b) ?_
-- then we traverse `newBody` and goal simultaneously
refine And.intro ?_ ?_
-- at branches outside the path `h` must concide with goal
· replace h := h.left
exact h
-- inside path we substitute variables from `fvars` into existential quantifiers.
· replace h := h.right
refine Exists.intro b ?_
-- at the end the goal must be `x' = x'`.
rfl
Equations
- One or more equations did not get rendered due to their size.
Instances For
Traverses P and goal simultaneously, proving goal.
Recursive implementation for withExistsElimAlongPath.
Given act : (a = a') → hb₁ → hb₂ → ... → hbₙ → goal where hb₁, ..., hbₙ are hypotheses
obtained when unpacking existential quantifiers with variables from exs, it proves goal using
Exists.elim. We use this to prove implication in the forward direction.
Equations
- ExistsAndEq.withExistsElimAlongPath h exs path act = ExistsAndEq.withExistsElimAlongPathImp h exs path [] act
Instances For
When P = ∃ f₁ ... fₙ, body, where exs = [f₁, ..., fₙ], this function takes
act : body and proves P using Exists.intro.
Example:
exs = []: act
exs = [b]:
P := ∃ b, body
Exists.intro b act
exs = [b, c]:
P := ∃ b c, body
Exists.intro b (Exists.intro c act)
...
Instances For
Generates a proof of ∃ a, p a → P'. We assume that fvars = [f₁, ..., fₙ] are free variables
and P' = ∃ f₁ ... fₙ, newBody, and path leads to a = a' in ∃ a, p a.
The proof follows the following structure:
example {α β : Type} (f : β → α) {p : α → Prop} :
(∃ a, p a ∧ ∃ b, a = f b) → (∃ b, p (f b) ∧ f b = f b) := by
-- withLocalDeclQ
intro h
refine h.elim (fun a ha ↦ ?_)
-- withExistsElimAlongPath: following the path we unpack all existential quantifiers.
-- at the end `hs = [hb]`.
have h' := ha
replace h' := h'.right
refine Exists.elim h' (fun b hb ↦ ?_)
replace h' := hb
have h_eq := h'
clear h'
-- go: we traverse `P` and `goal` simultaneously
have h' := ha
refine Exists.intro b ?_
refine And.intro ?_ ?_
-- outside the path goal must concide with `h_eq ▸ h'`
· replace h' := h'.left
exact Eq.mp (congrArg (fun t ↦ p t) h_eq) h'
-- inside the path:
· replace h' := h'.right
-- when `h'` starts with existential quantifier we replace it with next hypothesis from `hs`.
replace h' := hb
-- at the end the goal must be `x' = x'`.
rfl
Equations
- One or more equations did not get rendered due to their size.
Instances For
Traverses P and goal simultaneously, proving goal.
Triggers at goals of the form ∃ a, body and checks if body allows a single value a'
for a. If so, replaces a with a' and removes quantifier.
It looks through nested quantifiers and conjuctions searching for a a = a'
or a' = a subexpression.
Equations
- One or more equations did not get rendered due to their size.