feat(builtin/num): define lt predicate, and prove basic theorems
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
0760b5b53d
commit
fa4b60963b
2 changed files with 62 additions and 0 deletions
|
@ -94,11 +94,73 @@ theorem induction {P : num → Bool} (H1 : P zero) (H2 : ∀ n, P n → P (succ
|
||||||
show P a,
|
show P a,
|
||||||
from subst Qa abst_eq
|
from subst Qa abst_eq
|
||||||
|
|
||||||
|
theorem induction_on {P : num → Bool} (a : num) (H1 : P zero) (H2 : ∀ n, P n → P (succ n)) : P a
|
||||||
|
:= induction H1 H2 a
|
||||||
|
|
||||||
|
definition lt (m n : num) := ∃ P, (∀ i, P (succ i) → P i) ∧ P m ∧ ¬ P n
|
||||||
|
infix 50 < : lt
|
||||||
|
|
||||||
|
theorem lt_elim {m n : num} {B : Bool} (H1 : m < n)
|
||||||
|
(H2 : ∀ (P : num → Bool), (∀ i, P (succ i) → P i) → P m → ¬ P n → B)
|
||||||
|
: B
|
||||||
|
:= obtain P Pdef, from H1,
|
||||||
|
H2 P (and_eliml Pdef) (and_eliml (and_elimr Pdef)) (and_elimr (and_elimr Pdef))
|
||||||
|
|
||||||
|
theorem lt_intro {m n : num} {P : num → Bool} (H1 : ∀ i, P (succ i) → P i) (H2 : P m) (H3 : ¬ P n) : m < n
|
||||||
|
:= exists_intro P (and_intro H1 (and_intro H2 H3))
|
||||||
|
|
||||||
|
set_opaque lt true
|
||||||
|
|
||||||
|
theorem lt_nrefl (n : num) : ¬ (n < n)
|
||||||
|
:= not_intro
|
||||||
|
(assume N : n < n,
|
||||||
|
lt_elim N (λ P Pred Pn nPn, absurd Pn nPn))
|
||||||
|
|
||||||
|
theorem lt_succ {m n : num} : succ m < n → m < n
|
||||||
|
:= assume H : succ m < n,
|
||||||
|
lt_elim H
|
||||||
|
(λ (P : num → Bool) (Pred : ∀ i, P (succ i) → P i) (Psm : P (succ m)) (nPn : ¬ P n),
|
||||||
|
have Pm : P m,
|
||||||
|
from Pred m Psm,
|
||||||
|
show m < n,
|
||||||
|
from lt_intro Pred Pm nPn)
|
||||||
|
|
||||||
|
theorem not_lt_zero (n : num) : ¬ (n < zero)
|
||||||
|
:= induction_on n
|
||||||
|
(show ¬ (zero < zero),
|
||||||
|
from lt_nrefl zero)
|
||||||
|
(λ (n : num) (iH : ¬ (n < zero)),
|
||||||
|
show ¬ (succ n < zero),
|
||||||
|
from not_intro
|
||||||
|
(assume N : succ n < zero,
|
||||||
|
have nLTzero : n < zero,
|
||||||
|
from lt_succ N,
|
||||||
|
show false,
|
||||||
|
from absurd nLTzero iH))
|
||||||
|
|
||||||
|
theorem z_lt_succ_z : zero < succ zero
|
||||||
|
:= let P : num → Bool := λ x, x = zero
|
||||||
|
in have Pred : ∀ i, P (succ i) → P i,
|
||||||
|
from take i, assume Ps : P (succ i),
|
||||||
|
have si_eq_z : succ i = zero,
|
||||||
|
from Ps,
|
||||||
|
have si_ne_z : succ i ≠ zero,
|
||||||
|
from succ_nz i,
|
||||||
|
show P i,
|
||||||
|
from absurd_elim (P i) si_eq_z (succ_nz i),
|
||||||
|
have Pz : P zero,
|
||||||
|
from (refl zero),
|
||||||
|
have nPs : ¬ P (succ zero),
|
||||||
|
from succ_nz zero,
|
||||||
|
show zero < succ zero,
|
||||||
|
from lt_intro Pred Pz nPs
|
||||||
|
|
||||||
set_opaque num true
|
set_opaque num true
|
||||||
set_opaque Z true
|
set_opaque Z true
|
||||||
set_opaque S true
|
set_opaque S true
|
||||||
set_opaque zero true
|
set_opaque zero true
|
||||||
set_opaque succ true
|
set_opaque succ true
|
||||||
|
set_opaque lt true
|
||||||
end
|
end
|
||||||
|
|
||||||
definition num := num::num
|
definition num := num::num
|
||||||
|
|
Binary file not shown.
Loading…
Reference in a new issue