doc(doc/lean/tutorial): update tutorial
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
4d25cb7f47
commit
ceff335bb8
6 changed files with 361 additions and 36 deletions
|
@ -81,4 +81,17 @@ The Lean `let` construct can also be used to build calculational-like proofs.
|
|||
in s3.
|
||||
```
|
||||
|
||||
Finally, the [Nat (natural number) builtin library](../../src/builtin/Nat.lean) makes extensive use of calculational proofs.
|
||||
Finally, the [Nat (natural number) builtin library](../../src/builtin/Nat.lean) makes extensive use of calculational proofs.
|
||||
|
||||
The Lean simplifier can be used to reduce the size of calculational proofs.
|
||||
In the following example, we create a rewrite rule set with basic theorems from the Natural number library, and some of the axioms
|
||||
declared above.
|
||||
|
||||
```lean
|
||||
import tactic
|
||||
rewrite_set simple
|
||||
add_rewrite Nat::add_comm Nat::add_assoc Nat::add_left_comm Ax1 Ax2 Ax3 eq_id : simple
|
||||
theorem T'' : a = e
|
||||
:= calc a = 1 + d : (by simp simple)
|
||||
... = e : symm Ax4
|
||||
```
|
||||
|
|
|
@ -55,9 +55,9 @@ type of a given expression.
|
|||
The last command returns `Nat → Nat`. That is, the type of double is a function
|
||||
from `Nat` to `Nat`, where `Nat` is the type of the natural numbers.
|
||||
|
||||
The command `import` loads existing libraries and extensions. For example,
|
||||
the following command imports the command `find` that searches the Lean environment
|
||||
using regular expressions
|
||||
The command `import` loads existing libraries and extensions. The
|
||||
following command imports the command `find` that searches the Lean
|
||||
environment using regular expressions
|
||||
|
||||
```lean
|
||||
import find
|
||||
|
@ -89,7 +89,7 @@ The command `variables n m o : Nat` can be used a shorthand for the three comman
|
|||
|
||||
In Lean, proofs are also expressions, and all functionality provided for manipulating
|
||||
expressions is also available for manipulating proofs. For example, `refl n` is a proof
|
||||
for `n == n`. In Lean, `refl` is the reflexivity axiom.
|
||||
for `n = n`. In Lean, `refl` is the reflexivity axiom.
|
||||
|
||||
```lean
|
||||
check refl n
|
||||
|
@ -105,7 +105,7 @@ The following commands postulate two axioms `Ax1` and `Ax2` that state that `n =
|
|||
```
|
||||
|
||||
`Ax1` and `Ax2` are not just names. For example, `trans Ax1 Ax2` is a proof that
|
||||
`n == o`, where `trans` is the transitivity axiom.
|
||||
`n = o`, where `trans` is the transitivity axiom.
|
||||
|
||||
```lean
|
||||
check trans Ax1 Ax2
|
||||
|
@ -120,7 +120,7 @@ will reject the type incorrect term `trans Ax2 Ax1`.
|
|||
Because we use _proposition as types_, we must support _empty types_. For example,
|
||||
the type `false` must be empty, since we don't have a proof for `false`.
|
||||
|
||||
Most systems based on _propositions as types_ paradigm are based on constructive logic.
|
||||
Most systems based on the _propositions as types_ paradigm are based on constructive logic.
|
||||
Lean on the other hand is based on classical logic. The _excluded middle_ is a theorem
|
||||
in Lean, and `em p` is a proof for `p ∨ ¬ p`.
|
||||
|
||||
|
@ -141,7 +141,7 @@ called `nat_trans3`
|
|||
:= trans (trans H1 (symm H2)) H3
|
||||
```
|
||||
|
||||
The theorem `nat_trans3` has 7 parameters, it takes for natural numbers `a`, `b`, `c` and `d,
|
||||
The theorem `nat_trans3` has 7 parameters, it takes for natural numbers `a`, `b`, `c` and `d`,
|
||||
and three proofs showing that `a = b`, `c = b` and `c = d`, and returns a proof that `a = d`.
|
||||
In the example above, `symm` is the symmetry theorem. Now, we use `nat_trans3` in a simple
|
||||
example.
|
||||
|
@ -155,8 +155,9 @@ example.
|
|||
```
|
||||
|
||||
The theorem `nat_trans3` is somewhat inconvenient to use because it has 7 parameters.
|
||||
However, the first four parameters can be inferred from the last 3. We can `_` as placeholder
|
||||
that instructs Lean to synthesize this expression.
|
||||
However, the first four parameters can be inferred from the last 3. We can use `_` as placeholder
|
||||
that instructs Lean to synthesize this expression. The synthesis process is based on type inference, and it is
|
||||
the most basic forms of automation provided by Lean.
|
||||
|
||||
```lean
|
||||
check nat_trans3 _ _ _ _ Hxy Hzy Hzw
|
||||
|
@ -180,8 +181,8 @@ is quite simple, we are just instructing Lean to automatically insert the placeh
|
|||
check nat_trans3i Hxy Hzy Hzw
|
||||
```
|
||||
|
||||
Sometimes, Lean will not be able to infer the parameters automatically. So, whenever we
|
||||
define a theorem/definition/axiom/variable containing implicit arguments, Lean will
|
||||
Sometimes, Lean will not be able to infer the parameters automatically.
|
||||
So, whenever we define a theorem/definition/axiom/variable containing implicit arguments, Lean will
|
||||
automatically create an _explicit_ version where all parameters are explicit.
|
||||
The explicit version uses the same name with a `@` prefix.
|
||||
|
||||
|
@ -197,8 +198,8 @@ The axiom `refl`, and the theorems `trans` and `symm` all have implicit argument
|
|||
check @symm
|
||||
```
|
||||
|
||||
We can also instruct Lean to display all implicit arguments. This is useful
|
||||
when debugging non-trivial problems.
|
||||
We can also instruct Lean to display all implicit arguments when it prints expressions.
|
||||
This is useful when debugging non-trivial problems.
|
||||
|
||||
```lean
|
||||
set_option pp::implicit true -- show implicit arguments
|
||||
|
@ -206,27 +207,10 @@ when debugging non-trivial problems.
|
|||
set_option pp::implicit false -- hide implicit arguments
|
||||
```
|
||||
|
||||
Note that, in the examples above, we have seen two symbols for equality: `=` and `==`.
|
||||
Moreover, in the previous example, the `check` command stated that `nat_trans3i Hxy Hzy Hzw`
|
||||
has type `@eq ℕ x w`. For technical reasons, in Lean, we have heterogenous and homogeneous
|
||||
equality. In a nutshell, heterogenous is mainly used internally, and homogeneous are used externally
|
||||
when writing programs and specifications in Lean. Heterogenous equality allows us to compare
|
||||
elements of any type, and homogenous equality is just a definition on top of the heterogenous
|
||||
equality that expects arguments of the same type. The expression `t == s` is a heterogenous equality that is true
|
||||
iff `t` and `s` have the same interpretation. On the other hand `t = s` is a homogeneous equality,
|
||||
and is only type correct if `t` and `s` have the same type. The symbol `=` is actually notation for
|
||||
`eq t s`, where `eq` is defined (using heterogenous equality) as:
|
||||
In the previous example, the `check` command stated that `nat_trans3i Hxy Hzy Hzw`
|
||||
has type `@eq ℕ x w`. The expression `x = w` is just notational convenience.
|
||||
|
||||
```
|
||||
definition eq {A : TypeU} (a b : A) : Bool
|
||||
:= a == b
|
||||
```
|
||||
|
||||
We strongly discourage users from directly using heterogeneous equality. The main problem is that it is very easy to
|
||||
write nonsensical expressions such as `2 == true`. On the other hand, the expression `2 = true` is type incorrect,
|
||||
and Lean will report the mistake to the user.
|
||||
|
||||
We have seen many occurrences of `TypeU`. It is just a definition: `(Type U)`, where `U` is a _universe variable_.
|
||||
We have seen many occurrences of `TypeU`. It is just a definition for: `(Type U)`, where `U` is a _universe variable_.
|
||||
In Lean, the type of `Nat` and `Bool` is `Type`.
|
||||
|
||||
```lean
|
||||
|
@ -244,12 +228,319 @@ Lean returns `(Type 1)`. Similarly, the type of `(Type 1)` is `(Type 2)`. In Lea
|
|||
That is, we can provide an element of type `(Type i)` where an element of type `(Type j)` is expected when `i ≤ j`.
|
||||
This makes the system more convenient to use. Otherwise, we would need a reflexivity axiom for `Type` (i.e., `(Type 0)`),
|
||||
`Type 1`, `Type 2`, etc. Universe cumulativity improves usability, but it is not enough because
|
||||
We would still have the question: how big should `i` be? Moreover, if we choose an `i` that is not big enough
|
||||
we would still have the question: how big should `i` be? Moreover, if we choose an `i` that is not big enough
|
||||
we have to go back and correct all libraries. This is not satisfactory and not modular.
|
||||
So, in Lean, we allow user to declare _universe variables_ and simple constraints between them. The Lean kernel defines
|
||||
one universe variable `U`, and states that `U ≥ 1` using the command `universe U ≥ 1`.
|
||||
The Lean type casting library defines another universe variable called `M` and states that `universe M ≥ 1` and `universe M ≥ U + 1`.
|
||||
Lean reports an universe inconsistency if the universe constraints are inconsistency. For example, it will return an error
|
||||
Lean reports an universe inconsistency if the universe constraints are inconsistent. For example, it will return an error
|
||||
if execute the command `universe M ≥ U`. We can view universe variables as placeholders, and we can always solve
|
||||
the universe constraints and find and assignment for the universe variables used in our developments.
|
||||
This assignment allows us to produce a Lean specification that is not based on this particular feature.
|
||||
|
||||
Propositional logic
|
||||
-------------------
|
||||
|
||||
To manipulate formulas with a richer logical structure, it is important to master the notation Lean uses for building
|
||||
composite logical expressions out of basic formulas using _logical connectives_. The logical connectives (`and`, `or`, `not`, etc)
|
||||
are defined in the Lean [kernel](../../src/builtin/kernel.lean). The kernel also defines notational convention for rewriting formulas
|
||||
in a natural way. Here is a table showing the notation for the so called propositional (or Boolean) connectives.
|
||||
|
||||
|
||||
| Ascii | Ascii alt. | Unicode | Definition |
|
||||
|-------|--------------|---------|--------------|
|
||||
| true | | ⊤ | true |
|
||||
| false | | ⊥ | false |
|
||||
| not | | ¬ | not |
|
||||
| /\ | && | ∧ | and |
|
||||
| \/ | || | ∨ | or |
|
||||
| -> | | → | implies |
|
||||
| <-> | | ↔ | iff |
|
||||
|
||||
`true` and `false` are logical constants to denote the true and false propositions. Logical negation is a unary operator just like
|
||||
arithmetical negation on numbers. The other connectives are all binary operators. The meaning of the operators is the usual one.
|
||||
The table above makes clear that Lean supports unicode characters. We can use Ascii or/and unicode versions.
|
||||
Here is a simple example using the connectives above.
|
||||
|
||||
```lean
|
||||
variable q : Bool
|
||||
check p → q → p ∧ q
|
||||
check ¬ p → p ↔ false
|
||||
check p ∨ q → q ∨ p
|
||||
-- Ascii version
|
||||
check p -> q -> p && q
|
||||
check not p -> p <-> false
|
||||
check p || q -> q \/ p
|
||||
```
|
||||
|
||||
Depending on the platform, Lean uses unicode characters by default when printing expressions. The following commands can be used to
|
||||
change this behavior.
|
||||
|
||||
```lean
|
||||
set_option pp::unicode false
|
||||
check p → q → p ∧ q
|
||||
set_option pp::unicode true
|
||||
check p → q → p ∧ q
|
||||
```
|
||||
|
||||
Note that, it may seem that the symbols `->` and `→` are overloaded, and Lean uses them to represent Boolean implication and the type
|
||||
of functions. Actually, they are not overloaded, they are the same symbols. In Lean, the Boolean `p → q` expression is also the type
|
||||
of the functions that given a proof for `p`, returns a proof for `q`. This is very convenient for writing proofs.
|
||||
|
||||
```lean
|
||||
-- Hpq is a function that takes a proof for p and returns a proof for q
|
||||
axiom Hpq : p → q
|
||||
-- Hq is a proof/certificate for p
|
||||
axiom Hp : p
|
||||
-- The expression Hpq Hp is a proof/certificate for q
|
||||
check Hpq Hp
|
||||
```
|
||||
|
||||
In composite expressions, the precedences of the various binary
|
||||
connectives are in order of the above table, with `and` being the
|
||||
strongest and `iff` the weakest. For example, `a ∧ b → c ∨ d ∧ e`
|
||||
means `(a ∧ b) → (c ∨ (d ∧ e))`. All of them are right-associative.
|
||||
So, `p ∧ q ∧ r` means `p ∧ (q ∧ r)`. The actual precedence and fixity of all
|
||||
logical connectives is defined in the Lean [kernel definition file](../../src/builtin/kernel.lean).
|
||||
|
||||
Finally, `not`, `and`, `or` and `iff` are the actual names used when
|
||||
defining the Boolean connectives. They can be used as any other function.
|
||||
|
||||
```lean
|
||||
check and
|
||||
check or
|
||||
check not
|
||||
```
|
||||
|
||||
Lean supports _currying_ `and true` is a function from `Bool` to `Bool`.
|
||||
|
||||
```lean
|
||||
check and true
|
||||
definition id := and true
|
||||
eval id true
|
||||
eval id false
|
||||
```
|
||||
|
||||
Functions
|
||||
---------
|
||||
|
||||
There are many variable-binding constructs in mathematics. Lean expresses
|
||||
all of them using just one _abstraction_, which is a converse operation to
|
||||
function application. Given a variable `x`, a type `A`, and a term `t` that
|
||||
may or may not contain `x`, one can construct the so-called _lambda abstraction_
|
||||
`fun x : A, t`, or using unicode notation `λ x : A, t`. Here is some simple
|
||||
examples.
|
||||
|
||||
```lean
|
||||
check fun x : Nat, x + 1
|
||||
check fun x y : Nat, x + 2 * y
|
||||
check fun x y : Bool, not (x ∧ y)
|
||||
check λ x : Nat, x + 1
|
||||
check λ (x : Nat) (p : Bool), x = 0 ∨ p
|
||||
```
|
||||
|
||||
In many cases, Lean can automatically infer the type of the variable. Actually,
|
||||
In all examples above, the type can be automatically inferred.
|
||||
|
||||
```lean
|
||||
check fun x, x + 1
|
||||
check fun x y, x + 2 * y
|
||||
check fun x y, not (x ∧ y)
|
||||
check λ x, x + 1
|
||||
check λ x p, x = 0 ∨ p
|
||||
```
|
||||
|
||||
However, Lean will complain that it cannot infer the type of the
|
||||
variable in `fun x, x` because any type would work in this example.
|
||||
|
||||
The following example shows how to use lambda abstractions in
|
||||
function applications
|
||||
|
||||
```lean
|
||||
eval (fun x y, x + 2 * y) 1
|
||||
eval (fun x y, x + 2 * y) 1 2
|
||||
eval (fun x y, not (x ∧ y)) true false
|
||||
```
|
||||
|
||||
Lambda abstractions are also used to create proofs for propositions of the form `A → B`.
|
||||
This should be natural since Lean views `A → B` as the type of functions that given
|
||||
a proof for `A` returns a proof for `B`.
|
||||
For example, a proof for `p → p` is just `fun H : p, H` (the identity function).
|
||||
|
||||
```lean
|
||||
check fun H : p, H
|
||||
```
|
||||
|
||||
Definitional equality
|
||||
---------------------
|
||||
|
||||
Recall that the command `eval t` computes a normal form for the term `t`.
|
||||
In Lean, we say two terms are _definitionally equal_ if the have the same
|
||||
normal proof. For example, the terms `(λ x : Nat, x + 1) a` and `a + 1`
|
||||
are definitionally equal. The Lean type/proof checker uses the normalizer/evaluator when
|
||||
checking types/proofs. So, we can prove that two definitionally equal terms
|
||||
are equal using just `refl`. Here is a simple example.
|
||||
|
||||
```lean
|
||||
theorem def_eq_th (a : Nat) : ((λ x : Nat, x + 1) a) = a + 1
|
||||
:= refl (a+1)
|
||||
```
|
||||
|
||||
Provable equality
|
||||
-----------------
|
||||
|
||||
In the previous examples, we have used `nat_trans3 x y z w Hxy Hzy Hzw`
|
||||
to show that `x = w`. In this case, `x` and `w` are not definitionally equal,
|
||||
but they are provably equal in the environment that contains `nat_trans3` and
|
||||
axioms `Hxy`, `Hzy` and `Hzw`.
|
||||
|
||||
Proving
|
||||
-------
|
||||
|
||||
The Lean kernel contains basic theorems for creating proof terms. The
|
||||
basic theorems are useful for creating manual proofs. The are also the
|
||||
basic building blocks used by all automated proof engines available in
|
||||
Lean. The theorems can be broken into three different categories:
|
||||
introduction, elimination, and rewriting. First, we cover the introduction
|
||||
and elimination theorems for the basic Boolean connectives.
|
||||
|
||||
### And (conjuction)
|
||||
|
||||
The expression `and_intro H1 H2` creates a proof for `a ∧ b` using proofs
|
||||
`H1 : a` and `H2 : b`. We say `and_intro` is the _and-introduction_ operation.
|
||||
In the following example we use `and_intro` for creating a proof for
|
||||
`p → q → p ∧ q`.
|
||||
|
||||
```lean
|
||||
check fun (Hp : p) (Hq : q), and_intro Hp Hq
|
||||
```
|
||||
|
||||
The expression `and_eliml H` creates a proof `a` from a proof `H : a ∧ b`.
|
||||
Similarly `and_elimr H` is a proof for `b`. We say they are the _left/right and-elimination_.
|
||||
|
||||
```lean
|
||||
-- Proof for p ∧ q → p
|
||||
check fun H : p ∧ q, and_eliml H
|
||||
-- Proof for p ∧ q → q
|
||||
check fun H : p ∧ q, and_elimr H
|
||||
```
|
||||
|
||||
Now, we prove `p ∧ q → q ∧ p` with the following simple proof term.
|
||||
|
||||
```lean
|
||||
check fun H : p ∧ q, and_intro (and_elimr H) (and_eliml H)
|
||||
```
|
||||
|
||||
Note that the proof term is very similar to a function that just swaps the
|
||||
elements of a pair.
|
||||
|
||||
### Or (disjuction)
|
||||
|
||||
The expression `or_introl H1 b` creates a proof for `a ∨ b` using a proof `H1 : a`.
|
||||
Similarly, `or_intror a H2` creates a proof for `a ∨ b` using a proof `H2 : b`.
|
||||
We say they are the _left/right or-introduction_.
|
||||
|
||||
```lean
|
||||
-- Proof for p → p ∨ q
|
||||
check fun H : p, or_introl H q
|
||||
-- Proof for q → p ∨ q
|
||||
check fun H : q, or_intror p H
|
||||
```
|
||||
|
||||
The or-elimination rule is slightly more complicated. The basic idea is the
|
||||
following, we can prove `c` from `a ∨ b`, by showing we can prove `c`
|
||||
by assuming `a` or by assuming `b`. It is essentially a proof by cases.
|
||||
`or_elim Hab Hac Hbc` takes three arguments `Hab : a ∨ b`, `Hac : a → c` and `Hbc : b → c` and produces a proof for `c`.
|
||||
In the following example, we use `or_elim` to prove that `p v q → q ∨ p`.
|
||||
|
||||
```lean
|
||||
check fun H : p ∨ q,
|
||||
or_elim H
|
||||
(fun Hp : p, or_intror q Hp)
|
||||
(fun Hq : q, or_introl Hq p)
|
||||
```
|
||||
|
||||
### Not (negation)
|
||||
|
||||
`not_intro H` produces a proof for `¬ a` from `H : a → false`. That is,
|
||||
we obtain `¬ a` if we can derive `false` from `a`. The expression
|
||||
`absurd_elim b Ha Hna` produces a proof for `b` from `Ha : a` and `Hna : ¬ a`.
|
||||
That is, we can deduce anything if we have `a` and `¬ a`.
|
||||
We now use `not_intro` and `absurd_elim` to produce a proof term for
|
||||
`(a → b) → ¬ b → ¬ a`
|
||||
|
||||
```lean
|
||||
variables a b : Bool
|
||||
check fun (Hab : a → b) (Hnb : ¬ b),
|
||||
not_intro (fun Ha : a, absurd_elim false (Hab Ha) Hnb)
|
||||
```
|
||||
|
||||
Here is the proof term for `¬ a → b → (b → a) → c`
|
||||
|
||||
```lean
|
||||
variable c : Bool
|
||||
check fun (Hna : ¬ a) (Hb : b) (Hba : b → a),
|
||||
absurd_elim c (Hba Hb) Hna
|
||||
```
|
||||
|
||||
### Iff (if-and-only-if)
|
||||
|
||||
The expression `iff_intro H1 H2` produces a proof for `a ↔ b` from `H1 : a → b` and `H2 : b → a`.
|
||||
`iff_eliml H` produces a proof for `a → b` from `H : a ↔ b`. Similarly,
|
||||
`iff_elimr H` produces a proof for `b → a` from `H : a ↔ b`.
|
||||
Note that, in Lean, `a ↔ b` is definitionally equal to `a = b` when `a` and `b` have type `Bool`.
|
||||
Here is the proof term for `a ∧ b ↔ b ∧ a`
|
||||
|
||||
```lean
|
||||
check iff_intro (fun H : a ∧ b, and_intro (and_elimr H) (and_eliml H))
|
||||
(fun H : b ∧ a, and_intro (and_elimr H) (and_eliml H))
|
||||
```
|
||||
|
||||
### True and False
|
||||
|
||||
The expression `trivial` is a proof term for `true`, and `false_elim a H`
|
||||
produces a proof for `a` from `H : false`.
|
||||
|
||||
Other basic operators used in proof construction are `eqt_intro`, `eqt_elim`, `eqf_intro` and `eqf_elim`.
|
||||
`eqt_intro H` produces a proof for `a ↔ true` from `H : a`.
|
||||
`eqt_elim H` produces a proof for `a` from `H : a ↔ true`.
|
||||
`eqf_intro H` produces a proof for `a ↔ false` from `H : ¬ a`.
|
||||
`eqf_elim H` produces a proof for `¬ a` from `H : a ↔ false`.
|
||||
|
||||
```lean
|
||||
check @eqt_intro
|
||||
check @eqt_elim
|
||||
check @eqf_intro
|
||||
check @eqf_elim
|
||||
```
|
||||
|
||||
### Rewrite rules
|
||||
|
||||
The Lean kernel also contains many theorems that are meant to be used as rewriting/simplification rules.
|
||||
The conclusion of these theorems is of the form `t = s` or `t ↔ s`. For example, `and_id a` is proof term for
|
||||
`a ∧ a ↔ a`. The Lean simplifier can use these theorems to automatically create proof terms for us.
|
||||
The expression `(by simp [rule-set])` is similar to `_`, but it tells Lean to synthesize the proof term using the simplifier
|
||||
using the rewrite rule set named `[rule-set]`. In the following example, we create a simple rewrite rule set
|
||||
and use it to prove a theorem that would be quite tedious to prove by hand.
|
||||
|
||||
```lean
|
||||
-- import module that defines several tactics/strategies including "simp"
|
||||
import tactic
|
||||
-- create a rewrite rule set with name 'simple'
|
||||
rewrite_set simple
|
||||
-- add some theorems to the rewrite rule set 'simple'
|
||||
add_rewrite and_id and_truer and_truel and_comm and_assoc and_left_comm iff_id : simple
|
||||
theorem th1 (a b : Bool) : a ∧ b ∧ true ∧ b ∧ true ∧ b ↔ a ∧ b
|
||||
:= (by simp simple)
|
||||
```
|
||||
|
||||
In Lean, we can combine manual and automated proofs in a natural way. We can manually write the proof
|
||||
skeleton and use the `by` construct to invoke automated proof engines like the simplifier for filling the
|
||||
tedious steps. Here is a very simple example.
|
||||
|
||||
```lean
|
||||
theorem th2 (a b : Bool) : a ∧ b ↔ b ∧ a
|
||||
:= iff_intro
|
||||
(fun H : a ∧ b, (by simp simple))
|
||||
(fun H : b ∧ a, (by simp simple))
|
||||
```
|
||||
|
|
|
@ -242,6 +242,12 @@ theorem boolext {a b : Bool} (Hab : a → b) (Hba : b → a) : a = b
|
|||
theorem iff_intro {a b : Bool} (Hab : a → b) (Hba : b → a) : a ↔ b
|
||||
:= boolext Hab Hba
|
||||
|
||||
theorem iff_eliml {a b : Bool} (H : a ↔ b) : a → b
|
||||
:= (λ Ha : a, eqmp H Ha)
|
||||
|
||||
theorem iff_elimr {a b : Bool} (H : a ↔ b) : b → a
|
||||
:= (λ Hb : b, eqmpr H Hb)
|
||||
|
||||
theorem skolem_th {A : TypeU} {B : A → TypeU} {P : ∀ x : A, B x → Bool} :
|
||||
(∀ x, ∃ y, P x y) ↔ ∃ f, (∀ x, P x (f x))
|
||||
:= iff_intro
|
||||
|
@ -264,6 +270,9 @@ theorem neq_elim {A : TypeU} {a b : A} (H : a ≠ b) : a = b ↔ false
|
|||
theorem eq_id {A : TypeU} (a : A) : (a = a) ↔ true
|
||||
:= eqt_intro (refl a)
|
||||
|
||||
theorem iff_id (a : Bool) : (a ↔ a) ↔ true
|
||||
:= eqt_intro (refl a)
|
||||
|
||||
theorem or_comm (a b : Bool) : (a ∨ b) = (b ∨ a)
|
||||
:= boolext (assume H, or_elim H (λ H1, or_intror b H1) (λ H2, or_introl H2 a))
|
||||
(assume H, or_elim H (λ H1, or_intror a H1) (λ H2, or_introl H2 b))
|
||||
|
|
Binary file not shown.
|
@ -72,11 +72,14 @@ MK_CONSTANT(exists_to_eps_fn, name("exists_to_eps"));
|
|||
MK_CONSTANT(axiom_of_choice_fn, name("axiom_of_choice"));
|
||||
MK_CONSTANT(boolext_fn, name("boolext"));
|
||||
MK_CONSTANT(iff_intro_fn, name("iff_intro"));
|
||||
MK_CONSTANT(iff_eliml_fn, name("iff_eliml"));
|
||||
MK_CONSTANT(iff_elimr_fn, name("iff_elimr"));
|
||||
MK_CONSTANT(skolem_th_fn, name("skolem_th"));
|
||||
MK_CONSTANT(eqt_intro_fn, name("eqt_intro"));
|
||||
MK_CONSTANT(eqf_intro_fn, name("eqf_intro"));
|
||||
MK_CONSTANT(neq_elim_fn, name("neq_elim"));
|
||||
MK_CONSTANT(eq_id_fn, name("eq_id"));
|
||||
MK_CONSTANT(iff_id_fn, name("iff_id"));
|
||||
MK_CONSTANT(or_comm_fn, name("or_comm"));
|
||||
MK_CONSTANT(or_assoc_fn, name("or_assoc"));
|
||||
MK_CONSTANT(or_id_fn, name("or_id"));
|
||||
|
|
|
@ -209,6 +209,12 @@ inline expr mk_boolext_th(expr const & e1, expr const & e2, expr const & e3, exp
|
|||
expr mk_iff_intro_fn();
|
||||
bool is_iff_intro_fn(expr const & e);
|
||||
inline expr mk_iff_intro_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4) { return mk_app({mk_iff_intro_fn(), e1, e2, e3, e4}); }
|
||||
expr mk_iff_eliml_fn();
|
||||
bool is_iff_eliml_fn(expr const & e);
|
||||
inline expr mk_iff_eliml_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4) { return mk_app({mk_iff_eliml_fn(), e1, e2, e3, e4}); }
|
||||
expr mk_iff_elimr_fn();
|
||||
bool is_iff_elimr_fn(expr const & e);
|
||||
inline expr mk_iff_elimr_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4) { return mk_app({mk_iff_elimr_fn(), e1, e2, e3, e4}); }
|
||||
expr mk_skolem_th_fn();
|
||||
bool is_skolem_th_fn(expr const & e);
|
||||
inline expr mk_skolem_th_th(expr const & e1, expr const & e2, expr const & e3) { return mk_app({mk_skolem_th_fn(), e1, e2, e3}); }
|
||||
|
@ -224,6 +230,9 @@ inline expr mk_neq_elim_th(expr const & e1, expr const & e2, expr const & e3, ex
|
|||
expr mk_eq_id_fn();
|
||||
bool is_eq_id_fn(expr const & e);
|
||||
inline expr mk_eq_id_th(expr const & e1, expr const & e2) { return mk_app({mk_eq_id_fn(), e1, e2}); }
|
||||
expr mk_iff_id_fn();
|
||||
bool is_iff_id_fn(expr const & e);
|
||||
inline expr mk_iff_id_th(expr const & e1) { return mk_app({mk_iff_id_fn(), e1}); }
|
||||
expr mk_or_comm_fn();
|
||||
bool is_or_comm_fn(expr const & e);
|
||||
inline expr mk_or_comm_th(expr const & e1, expr const & e2) { return mk_app({mk_or_comm_fn(), e1, e2}); }
|
||||
|
|
Loading…
Reference in a new issue