Intuitionistic tautology (itauto
) decision procedure #
The itauto
tactic will prove any intuitionistic tautology. It implements the well known
G4ip
algorithm:
Dyckhoff, Contraction-free sequent calculi for intuitionistic logic.
All built in propositional connectives are supported: true
, false
, and
, or
, implies
,
not
, iff
, xor
, as well as eq
and ne
on propositions. Anything else, including definitions
and predicate logical connectives (forall
and exists
), are not supported, and will have to be
simplified or instantiated before calling this tactic.
The resulting proofs will never use any axioms except possibly propext
, and propext
is only
used if the input formula contains an equality of propositions p = q
.
Implementation notes #
The core logic of the prover is in three functions:
prove : context → prop → state_t ℕ option proof
: The main entry point. Gets a context and a goal, and returns aproof
object or fails, usingstate_t ℕ
for the name generator.search : context → prop → state_t ℕ option proof
: Same meaning asproof
, called during the search phase (see below).context.add : prop → proof → context → except (prop → proof) context
: Adds a proposition with its proof into the context, but it also does some simplifications on the spot while doing so. It will either return the new context, or if it happens to notice a proof of false, it will return a function to compute a proof of any proposition in the original context.
The intuitionistic logic rules are separated into three groups:
- level 1: No splitting, validity preserving: apply whenever you can.
Left rules in
context.add
, right rules inprove
.context.add
:- simplify
Γ, ⊤ ⊢ B
toΓ ⊢ B
Γ, ⊥ ⊢ B
is true- simplify
Γ, A ∧ B ⊢ C
toΓ, A, B ⊢ C
- simplify
Γ, ⊥ → A ⊢ B
toΓ ⊢ B
- simplify
Γ, ⊤ → A ⊢ B
toΓ, A ⊢ B
- simplify
Γ, A ∧ B → C ⊢ D
toΓ, A → B → C ⊢ D
- simplify
Γ, A ∨ B → C ⊢ D
toΓ, A → C, B → C ⊢ D
- simplify
prove
:Γ ⊢ ⊤
is true- simplify
Γ ⊢ A → B
toΓ, A ⊢ B
search
:Γ, P ⊢ P
is true- simplify
Γ, P, P → A ⊢ B
toΓ, P, A ⊢ B
- level 2: Splitting rules, validity preserving: apply after level 1 rules. Done in
prove
- simplify
Γ ⊢ A ∧ B
toΓ ⊢ A
andΓ ⊢ B
- simplify
Γ, A ∨ B ⊢ C
toΓ, A ⊢ C
andΓ, B ⊢ C
- simplify
- level 3: Splitting rules, not validity preserving: apply only if nothing else applies.
Done in
search
Γ ⊢ A ∨ B
follows fromΓ ⊢ A
Γ ⊢ A ∨ B
follows fromΓ ⊢ B
Γ, (A₁ → A₂) → C ⊢ B
follows fromΓ, A₂ → C, A₁ ⊢ A₂
andΓ, C ⊢ B
This covers the core algorithm, which only handles true
, false
, and
, or
, and implies
.
For iff
and eq
, we treat them essentially the same as (p → q) ∧ (q → p)
, although we use
a different prop
representation because we have to remember to apply different theorems during
replay. For definitions like not
and xor
, we just eagerly unfold them. (This could potentially
cause a blowup issue for xor
, but it isn't used very often anyway. We could add it to the prop
grammar if it matters.)
Tags #
propositional logic, intuitionistic logic, decision procedure
- and : tactic.itauto.and_kind
- iff : tactic.itauto.and_kind
- eq : tactic.itauto.and_kind
Different propositional constructors that are variants of "and" for the purposes of the theorem prover.
Instances for tactic.itauto.and_kind
- tactic.itauto.and_kind.has_sizeof_inst
- tactic.itauto.and_kind.has_reflect
- tactic.itauto.and_kind.inhabited
Equations
- var : ℕ → tactic.itauto.prop
- true : tactic.itauto.prop
- false : tactic.itauto.prop
- and' : tactic.itauto.and_kind → tactic.itauto.prop → tactic.itauto.prop → tactic.itauto.prop
- or : tactic.itauto.prop → tactic.itauto.prop → tactic.itauto.prop
- imp : tactic.itauto.prop → tactic.itauto.prop → tactic.itauto.prop
A reified inductive type for propositional logic.
Instances for tactic.itauto.prop
- tactic.itauto.prop.has_sizeof_inst
- tactic.itauto.prop.has_reflect
- tactic.itauto.prop.inhabited
- tactic.itauto.prop.has_to_format
- tactic.itauto.prop.has_lt
Constructor for p ∧ q
.
Constructor for p ↔ q
.
Constructor for p = q
.
Constructor for ¬ p
.
Equations
Constructor for xor p q
.
Equations
Given the contents of an and
variant, return the two conjuncts.
Equations
- tactic.itauto.and_kind.eq.sides A B = (A.imp B, B.imp A)
- tactic.itauto.and_kind.iff.sides A B = (A.imp B, B.imp A)
- tactic.itauto.and_kind.and.sides A B = (A, B)
A comparator for and_kind
. (There should really be a derive handler for this.)
Equations
- p.cmp q = p.cases_on (q.cases_on ordering.eq ordering.lt ordering.lt) (q.cases_on ordering.gt ordering.eq ordering.lt) (q.cases_on ordering.gt ordering.gt ordering.eq)
A comparator for propositions. (There should really be a derive handler for this.)
Equations
- p.cmp q = tactic.itauto.prop.rec (λ (p : ℕ) (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), cmp p q) ordering.lt ordering.lt (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt)) (λ (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), ordering.gt) ordering.eq ordering.lt (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt)) (λ (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), ordering.gt) ordering.gt ordering.eq (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt)) (λ (ap : tactic.itauto.and_kind) (p_ᾰ p_ᾰ_1 : tactic.itauto.prop) (p₁ p₂ : tactic.itauto.prop → ordering) (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), ordering.gt) ordering.gt ordering.gt (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), (λ (q₁ q₂ : tactic.itauto.prop), (ap.cmp q_ᾰ).or_else ((p₁ q₁).or_else (p₂ q₂))) q_ᾰ_1 q_ᾰ_2) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt)) (λ (p_ᾰ p_ᾰ_1 : tactic.itauto.prop) (p₁ p₂ : tactic.itauto.prop → ordering) (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), ordering.gt) ordering.gt ordering.gt (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), ordering.gt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), (λ (q₂ : tactic.itauto.prop), (p₁ q_ᾰ).or_else (p₂ q₂)) q_ᾰ_1) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.lt)) (λ (p_ᾰ p_ᾰ_1 : tactic.itauto.prop) (p₁ p₂ : tactic.itauto.prop → ordering) (q : tactic.itauto.prop), q.cases_on (λ (q : ℕ), ordering.gt) ordering.gt ordering.gt (λ (q_ᾰ : tactic.itauto.and_kind) (q_ᾰ_1 q_ᾰ_2 : tactic.itauto.prop), ordering.gt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), ordering.gt) (λ (q_ᾰ q_ᾰ_1 : tactic.itauto.prop), (λ (q₂ : tactic.itauto.prop), (p₁ q_ᾰ).or_else (p₂ q₂)) q_ᾰ_1)) p q
Equations
- tactic.itauto.prop.has_lt = {lt := λ (p q : tactic.itauto.prop), p.cmp q = ordering.lt}
Equations
- tactic.itauto.has_lt.lt.decidable_rel = λ (_x _x_1 : tactic.itauto.prop), ordering.decidable_eq (_x.cmp _x_1) ordering.lt
- sorry : tactic.itauto.proof
- hyp : name → tactic.itauto.proof
- triv : tactic.itauto.proof
- exfalso' : tactic.itauto.proof → tactic.itauto.proof
- intro : name → tactic.itauto.proof → tactic.itauto.proof
- and_left : tactic.itauto.and_kind → tactic.itauto.proof → tactic.itauto.proof
- and_right : tactic.itauto.and_kind → tactic.itauto.proof → tactic.itauto.proof
- and_intro : tactic.itauto.and_kind → tactic.itauto.proof → tactic.itauto.proof → tactic.itauto.proof
- curry : tactic.itauto.and_kind → tactic.itauto.proof → tactic.itauto.proof
- curry₂ : tactic.itauto.and_kind → tactic.itauto.proof → tactic.itauto.proof → tactic.itauto.proof
- app' : tactic.itauto.proof → tactic.itauto.proof → tactic.itauto.proof
- or_imp_left : tactic.itauto.proof → tactic.itauto.proof
- or_imp_right : tactic.itauto.proof → tactic.itauto.proof
- or_inl : tactic.itauto.proof → tactic.itauto.proof
- or_inr : tactic.itauto.proof → tactic.itauto.proof
- or_elim' : tactic.itauto.proof → name → tactic.itauto.proof → tactic.itauto.proof → tactic.itauto.proof
- decidable_elim : bool → name → name → tactic.itauto.proof → tactic.itauto.proof → tactic.itauto.proof
- em : bool → name → tactic.itauto.proof
- imp_imp_simp : name → tactic.itauto.proof → tactic.itauto.proof
A reified inductive proof type for intuitionistic propositional logic.
Instances for tactic.itauto.proof
- tactic.itauto.proof.has_sizeof_inst
- tactic.itauto.proof.has_reflect
- tactic.itauto.proof.inhabited
- tactic.itauto.proof.has_to_format
Equations
Map a function over the proof (regardless of whether the proof is successful or not).
Equations
- tactic.itauto.map_proof f (b, p, n) = (b, f p, n)
Convert a value-with-success to an optional value.
Equations
- tactic.itauto.is_ok (bool.tt, p) = option.some p
- tactic.itauto.is_ok (bool.ff, p) = option.none
Skip the continuation and return a failed proof if the boolean is false.
Equations
- tactic.itauto.when_ok bool.tt f n = f n
- tactic.itauto.when_ok bool.ff f n = (bool.ff, tactic.itauto.proof.sorry, n)
A decision procedure for intuitionistic propositional logic. Unlike finish
and tauto!
this
tactic never uses the law of excluded middle (without the !
option), and the proof search is
tailored for this use case. (itauto!
will work as a classical SAT solver, but the algorithm is
not very good in this situation.)
example (p : Prop) : ¬ (p ↔ ¬ p) := by itauto
itauto [a, b]
will additionally attempt case analysis on a
and b
assuming that it can derive
decidable a
and decidable b
. itauto *
will case on all decidable propositions that it can
find among the atomic propositions, and itauto! *
will case on all propositional atoms.
Warning: This can blow up the proof search, so it should be used sparingly.