feat(*): new numeral encoding

This commit is contained in:
Leonardo de Moura 2015-10-11 20:29:31 -07:00
parent 8ee214f133
commit 26eb6fa849
38 changed files with 236 additions and 172 deletions

View file

@ -14,15 +14,14 @@ namespace binary
local notation a * b := op₁ a b
local notation a ⁻¹ := inv a
local notation 1 := one
definition commutative [reducible] := ∀a b, a * b = b * a
definition associative [reducible] := ∀a b c, (a * b) * c = a * (b * c)
definition left_identity [reducible] := ∀a, 1 * a = a
definition right_identity [reducible] := ∀a, a * 1 = a
definition left_inverse [reducible] := ∀a, a⁻¹ * a = 1
definition right_inverse [reducible] := ∀a, a * a⁻¹ = 1
definition left_cancelative [reducible] := ∀a b c, a * b = a * c → b = c
definition commutative [reducible] := ∀a b, a * b = b * a
definition associative [reducible] := ∀a b c, (a * b) * c = a * (b * c)
definition left_identity [reducible] := ∀a, one * a = a
definition right_identity [reducible] := ∀a, a * one = a
definition left_inverse [reducible] := ∀a, a⁻¹ * a = one
definition right_inverse [reducible] := ∀a, a * a⁻¹ = one
definition left_cancelative [reducible] := ∀a b c, a * b = a * c → b = c
definition right_cancelative [reducible] := ∀a b c, a * b = c * b → a = c
definition inv_op_cancel_left [reducible] := ∀a b, a⁻¹ * (a * b) = b

View file

@ -121,8 +121,6 @@ namespace category
namespace ops
notation `type`:max := Type_category
notation 1 := Category_one
notation 2 := Category_two
postfix `ᵒᵖ`:max := opposite.Opposite
infixr `×c`:30 := product.Prod_category
attribute type_category [instance]

View file

@ -191,10 +191,8 @@ theorem pow_ge_one_of_ge_one {a : A} (H : a ≥ 1) (n : ) : a ^ n ≥ 1 :=
apply gt_of_ge_of_gt H zero_lt_one
end
local notation 2 := (1 : A) + 1
theorem pow_two_add (n : ) : 2^n + 2^n = 2^(succ n) :=
by rewrite [pow_succ', left_distrib, *mul_one]
theorem pow_two_add (n : ) : (2:A)^n + 2^n = 2^(succ n) :=
by rewrite [pow_succ', -one_add_one_eq_two, left_distrib, *mul_one]
end ordered_ring

View file

@ -281,16 +281,18 @@ section linear_ordered_field
theorem two_pos : (1 : A) + 1 > 0 :=
add_pos zero_lt_one zero_lt_one
theorem two_ne_zero : (1 : A) + 1 ≠ 0 :=
ne.symm (ne_of_lt two_pos)
theorem one_add_one_ne_zero : 1 + 1 ≠ (0:A) :=
ne.symm (ne_of_lt two_pos)
notation 2 := 1 + 1
theorem two_ne_zero : 2 ≠ (0:A) :=
by unfold bit0; apply one_add_one_ne_zero
theorem add_halves (a : A) : a / 2 + a / 2 = a :=
calc
a / 2 + a / 2 = (a + a) / 2 : by rewrite div_add_div_same
... = (a * 1 + a * 1) / 2 : by rewrite mul_one
... = (a * 2) / 2 : by rewrite left_distrib
... = (a * (1 + 1)) / 2 : by rewrite left_distrib
... = (a * 2) / 2 : by rewrite one_add_one_eq_two
... = a : by rewrite [@mul_div_cancel A _ _ _ two_ne_zero]
theorem sub_self_div_two (a : A) : a - a / 2 = a / 2 :=
@ -305,28 +307,29 @@ section linear_ordered_field
end
theorem div_two_sub_self (a : A) : a / 2 - a = - (a / 2) :=
by rewrite [-{a}add_halves at {2}, sub_add_eq_sub_sub, sub_self, zero_sub]
by rewrite [-{a}add_halves at {2}, sub_add_eq_sub_sub, sub_self, zero_sub]
theorem add_self_div_two (a : A) : (a + a) / 2 = a :=
symm (iff.mpr (!eq_div_iff_mul_eq (ne_of_gt (add_pos zero_lt_one zero_lt_one)))
(by rewrite [left_distrib, *mul_one]))
symm (iff.mpr (!eq_div_iff_mul_eq (ne_of_gt (add_pos zero_lt_one zero_lt_one)))
(by rewrite [left_distrib, *mul_one]))
theorem two_ge_one : (2 : A) ≥ 1 :=
by rewrite -(add_zero 1) at {3}; apply add_le_add_left; apply zero_le_one
theorem two_ge_one : (2:A) ≥ 1 :=
calc (2:A) = 1+1 : one_add_one_eq_two
... ≥ 1+0 : add_le_add_left (le_of_lt zero_lt_one)
... = 1 : add_zero
theorem mul_le_mul_of_mul_div_le (H : a * (b / c) ≤ d) (Hc : c > 0) : b * a ≤ d * c :=
begin
rewrite [-mul_div_assoc at H, mul.comm b],
apply le_mul_of_div_le Hc H
end
begin
rewrite [-mul_div_assoc at H, mul.comm b],
apply le_mul_of_div_le Hc H
end
theorem div_two_lt_of_pos (H : a > 0) : a / (1 + 1) < a :=
have Ha : a / (1 + 1) > 0, from div_pos_of_pos_of_pos H (add_pos zero_lt_one zero_lt_one),
calc
have Ha : a / (1 + 1) > 0, from div_pos_of_pos_of_pos H (add_pos zero_lt_one zero_lt_one),
calc
a / (1 + 1) < a / (1 + 1) + a / (1 + 1) : lt_add_of_pos_left Ha
... = a : add_halves
theorem div_mul_le_div_mul_of_div_le_div_pos {e : A} (Hb : b ≠ 0) (Hd : d ≠ 0) (H : a / b ≤ c / d)
(He : e > 0) : a / (b * e) ≤ c / (d * e) :=
begin
@ -338,26 +341,26 @@ section linear_ordered_field
end
theorem exists_add_lt_and_pos_of_lt (H : b < a) : ∃ c : A, b + c < a ∧ c > 0 :=
exists.intro ((a - b) / (1 + 1))
exists.intro ((a - b) / (1 + 1))
(and.intro (assert H2 : a + a > (b + b) + (a - b), from calc
a + a > b + a : add_lt_add_right H
... = b + a + b - b : add_sub_cancel
... = b + b + a - b : add.right_comm
... = (b + b) + (a - b) : add_sub,
assert H3 : (a + a) / (1 + 1) > ((b + b) + (a - b)) / (1 + 1),
assert H3 : (a + a) / 2 > ((b + b) + (a - b)) / 2,
from div_lt_div_of_lt_of_pos H2 two_pos,
by rewrite [add_self_div_two at H3, -div_add_div_same at H3, add_self_div_two at H3];
by rewrite [one_add_one_eq_two, sub_eq_add_neg, add_self_div_two at H3, -div_add_div_same at H3, add_self_div_two at H3];
exact H3)
(div_pos_of_pos_of_pos (iff.mpr !sub_pos_iff_lt H) two_pos))
theorem ge_of_forall_ge_sub {a b : A} (H : ∀ ε : A, ε > 0 → a ≥ b - ε) : a ≥ b :=
begin
begin
apply le_of_not_gt,
intro Hb,
cases exists_add_lt_and_pos_of_lt Hb with [c, Hc],
let Hc' := H c (and.right Hc),
apply (not_le_of_gt (and.left Hc)) (iff.mpr !le_add_iff_sub_right_le Hc')
end
end
end linear_ordered_field

View file

@ -48,6 +48,9 @@ section semiring
variables [s : semiring A] (a b c : A)
include s
theorem one_add_one_eq_two : 1 + 1 = (2:A) :=
by unfold bit0
theorem ne_zero_of_mul_ne_zero_right {a b : A} (H : a * b ≠ 0) : a ≠ 0 :=
suppose a = 0,
have a * b = 0, from this⁻¹ ▸ zero_mul b,

View file

@ -69,15 +69,15 @@ else
open decidable
theorem decode_encode_sum : ∀ s : sum A B, decode_sum (encode_sum s) = some s
| (sum.inl a) :=
assert aux : 2 > 0, from dec_trivial,
assert aux : 2 > (0:nat), from dec_trivial,
begin
esimp [encode_sum, decode_sum],
rewrite [mul_mod_right, if_pos (eq.refl (0 : nat)), mul_div_cancel_left _ aux, encodable.encodek]
end
| (sum.inr b) :=
assert aux₁ : 2 > 0, from dec_trivial,
assert aux₂ : 1 mod 2 = 1, by rewrite [nat.modulo_def],
assert aux₃ : 1 ≠ 0, from dec_trivial,
assert aux₁ : 2 > (0:nat), from dec_trivial,
assert aux₂ : 1 mod 2 = (1:nat), by rewrite [nat.modulo_def],
assert aux₃ : 1 ≠ (0:nat), from dec_trivial,
begin
esimp [encode_sum, decode_sum],
rewrite [add.comm, add_mul_mod_self_left, aux₂, if_neg aux₃, add_sub_cancel_left,

View file

@ -44,7 +44,7 @@ section depchoice
Recall that ⟨f, H⟩ is notation for (sigma.mk f H)
-/
theorem stronger_dependent_choice_of_encodable_of_decidable
: (∀ a, ∃ b, R a b) → (∀ a, Σ f, f 0 = a ∧ ∀ n, f n ~ f (n+1)) :=
: (∀ a, ∃ b, R a b) → (∀ a, Σ f, f (0:nat) = a ∧ ∀ n, f n ~ f (n+1)) :=
assume H : ∀ a, ∃ b, a ~ b,
take a : A,
let f : nat → A := f_aux a H in
@ -71,7 +71,7 @@ section sigma_depchoice
| 0 := a
| (n+1) := (H (f_aux n)).1
theorem sigma_dependent_choice : (∀ a, Σ b, R a b) → (∀ a, Σ f, f 0 = a ∧ ∀ n, f n ~ f (n+1)) :=
theorem sigma_dependent_choice : (∀ a, Σ b, R a b) → (∀ a, Σ f, f (0:nat) = a ∧ ∀ n, f n ~ f (n+1)) :=
assume H : ∀ a, Σ b, a ~ b,
take a : A,
let f : nat → A := f_aux a H in

View file

@ -7,7 +7,7 @@ Finite ordinal types.
-/
import data.list.basic data.finset.basic data.fintype.card algebra.group data.equiv
open eq.ops nat function list finset fintype
open - [notations] algebra
open algebra
structure fin (n : nat) := (val : nat) (is_lt : val < n)
@ -103,11 +103,12 @@ theorem val_lt : ∀ i : fin n, val i < n
lemma max_lt (i j : fin n) : max i j < n :=
max_lt (is_lt i) (is_lt j)
definition lift : fin n → Π m, fin (n + m)
definition lift : fin n → Π m : nat, fin (n + m)
| (mk v h) m := mk v (lt_add_of_lt_right h m)
definition lift_succ (i : fin n) : fin (nat.succ n) :=
lift i 1
have r : fin (n+1), from lift i 1,
r
definition maxi [reducible] : fin (succ n) :=
mk n !lt_succ_self

View file

@ -93,7 +93,7 @@ finset.induction_on s dec_trivial
(suppose 0 ∈ s,
assert a ≠ 0, from suppose a = 0, by subst a; contradiction,
begin
cases a with a, contradiction,
cases a with a, exact absurd rfl `zero ≠ 0`,
have odd (2*2^a), by rewrite [pow_succ' at o, mul.comm]; exact o,
have even (2*2^a), from !even_two_mul,
exact absurd `even (2*2^a)` `odd (2*2^a)`
@ -123,7 +123,7 @@ begin
match n with
| 0 := iff.intro (λ h, !mem_insert) (λ h, by rewrite [hw at zmem]; exact zmem)
| succ m :=
assert d₁ : 1 div 2 = 0, from dec_trivial,
assert d₁ : 1 div 2 = (0:nat), from dec_trivial,
assert aux : _, from calc
succ m ∈ of_nat (2 * w + 1) ↔ m ∈ of_nat ((2*w+1) div 2) : succ_mem_of_nat
... ↔ m ∈ of_nat w : by rewrite [add.comm, add_mul_div_self_left _ _ (dec_trivial : 2 > 0), d₁, zero_add]
@ -248,12 +248,12 @@ begin
reflexivity,
cases a with a,
{ rewrite [predimage_insert_zero, ih, to_nat_insert nains, pow_zero],
have 0 ∉ of_nat (to_nat s), by rewrite of_nat_to_nat; assumption,
have 0 ∉ of_nat (to_nat s), begin rewrite of_nat_to_nat, exact nains end,
have even (to_nat s), from iff.mp !even_of_not_zero_mem this,
obtain w (hw : to_nat s = 2*w), from exists_of_even this,
obtain (w : nat) (hw : to_nat s = 2*w), from exists_of_even this,
begin
rewrite hw,
have d₁ : 1 div 2 = 0, from dec_trivial,
have d₁ : 1 div 2 = (0:nat), from dec_trivial,
show 2 * w div 2 = (1 + 2 * w) div 2, by
rewrite [add_mul_div_self_left _ _ (dec_trivial : 2 > 0), mul.comm,
mul_div_cancel _ (dec_trivial : 2 > 0), d₁, zero_add]

View file

@ -30,7 +30,7 @@ import algebra.relation algebra.binary algebra.ordered_ring
open eq.ops
open prod relation nat
open decidable binary
open - [notations] algebra
open algebra
/- the type of integers -/
@ -48,16 +48,28 @@ attribute int.of_nat [coercion]
notation `-[1+ ` n `]` := int.neg_succ_of_nat n -- for pretty-printing output
definition int_has_zero [reducible] [instance] : has_zero int :=
has_zero.mk (of_nat 0)
definition int_has_one [reducible] [instance] : has_one int :=
has_one.mk (of_nat 1)
theorem int_zero_eq_nat_zero : (0:int) = of_nat (0:nat) :=
rfl
theorem int_one_eq_nat_one : (1:int) = of_nat (1:nat) :=
rfl
/- definitions of basic functions -/
definition neg_of_nat :
| 0 := 0
| 0 := 0
| (succ m) := -[1+ m]
definition sub_nat_nat (m n : ) : :=
match (n - m : nat) with
| 0 := (m - n : nat) -- m ≥ n
| (succ k) := -[1+ k] -- m < n, and n - m = succ k
| 0 := of_nat (m - n) -- m ≥ n
| (succ k) := -[1+ k] -- m < n, and n - m = succ k
end
protected definition neg (a : ) : :=
@ -116,14 +128,13 @@ theorem of_nat_succ (n : ) : of_nat (succ n) = of_nat n + 1 := rfl
theorem of_nat_mul (n m : ) : of_nat (n * m) = of_nat n * of_nat m := rfl
theorem sub_nat_nat_of_ge {m n : } (H : m ≥ n) : sub_nat_nat m n = of_nat (m - n) :=
show sub_nat_nat m n = nat.cases_on 0 (m - n : nat) _, from (sub_eq_zero_of_le H) ▸ rfl
show sub_nat_nat m n = nat.cases_on 0 (m -[nat] n) _, from (sub_eq_zero_of_le H) ▸ rfl
section
local attribute sub_nat_nat [reducible]
theorem sub_nat_nat_of_lt {m n : } (H : m < n) :
sub_nat_nat m n = -[1+ pred (n - m)] :=
theorem sub_nat_nat_of_lt {m n : } (H : m < n) : sub_nat_nat m n = -[1+ pred (n - m)] :=
have H1 : n - m = succ (pred (n - m)), from eq.symm (succ_pred_of_pos (sub_pos_of_lt H)),
show sub_nat_nat m n = nat.cases_on (succ (nat.pred (n - m))) (m - n : nat) _, from H1 ▸ rfl
show sub_nat_nat m n = nat.cases_on (succ (nat.pred (n - m))) (m -[nat] n) _, from H1 ▸ rfl
end
definition nat_abs (a : ) : := int.cases_on a function.id succ
@ -562,7 +573,7 @@ theorem succ_neg_succ (a : ) : succ (-succ a) = -a :=
by rewrite [neg_succ,succ_pred]
theorem neg_pred (a : ) : -pred a = succ (-a) :=
by krewrite [↑pred,neg_sub,sub_eq_add_neg,add.comm]
by rewrite [↑pred,neg_sub,sub_eq_add_neg,add.comm]
theorem pred_neg_pred (a : ) : pred (-pred a) = -a :=
by rewrite [neg_pred,pred_succ]

View file

@ -20,14 +20,14 @@ namespace int
protected definition divide (a b : ) : :=
sign b *
(match a with
| of_nat m := (m div (nat_abs b) : nat)
| of_nat m := of_nat (m div (nat_abs b))
| -[1+m] := -[1+ ((m:nat) div (nat_abs b))]
end)
definition int_has_divide [reducible] [instance] [priority int.prio] : has_divide int :=
has_divide.mk int.divide
lemma of_nat_div_eq (m : nat) (b : ) : (of_nat m) div b = sign b * (m div (nat_abs b) : nat) :=
lemma of_nat_div_eq (m : nat) (b : ) : (of_nat m) div b = sign b * of_nat (m div (nat_abs b)) :=
rfl
lemma neg_succ_div_eq (m: nat) (b : ) : -[1+m] div b = sign b * -[1+ (m div (nat_abs b))] :=
@ -36,7 +36,7 @@ rfl
lemma divide.def (a b : ) : a div b =
sign b *
(match a with
| of_nat m := (m div (nat_abs b) : nat)
| of_nat m := of_nat (m div (nat_abs b))
| -[1+m] := -[1+ ((m:nat) div (nat_abs b))]
end) :=
rfl
@ -50,7 +50,7 @@ has_modulo.mk int.modulo
lemma modulo.def (a b : ) : a mod b = a - a div b * b :=
rfl
notation [priority int.prio] a ≡ b `[mod `:100 c `]`:0 := a mod c = b mod c
notation [priority int.prio] a ≡ b `[mod `:0 c:0 `]` := a mod c = b mod c
/- div -/
@ -105,11 +105,11 @@ by krewrite [of_nat_div_eq, nat.zero_div, mul_zero]
theorem div_zero (a : ) : a div 0 = 0 :=
by krewrite [divide.def, sign_zero, zero_mul]
theorem div_one (a : ) :a div 1 = a :=
theorem div_one (a : ) : a div 1 = a :=
assert (1 : int) > 0, from dec_trivial,
int.cases_on a
(take m, by rewrite [-of_nat_div, nat.div_one])
(take m, by rewrite [!neg_succ_of_nat_div this, -of_nat_div, nat.div_one])
(take m : nat, by rewrite [int_one_eq_nat_one, -of_nat_div, nat.div_one])
(take m : nat, by rewrite [!neg_succ_of_nat_div this, int_one_eq_nat_one, -of_nat_div, nat.div_one])
theorem eq_div_mul_add_mod (a b : ) : a = a div b * b + a mod b :=
!add.comm ▸ eq_add_of_sub_eq rfl
@ -231,7 +231,7 @@ theorem add_mul_div_self_left (a : ) {b : } (c : ) (H : b ≠ 0) :
theorem mul_div_cancel (a : ) {b : } (H : b ≠ 0) : a * b div b = a :=
calc
a * b div b = (0 + a * b) div b : zero_add
a * b div b = (0 + a * b) div b : algebra.zero_add
... = 0 div b + a : !add_mul_div_self H
... = a : by rewrite [zero_div, zero_add]

View file

@ -7,7 +7,7 @@ Definitions and properties of gcd, lcm, and coprime.
-/
import .div data.nat.gcd
open eq.ops
open - [notations] algebra
open algebra
namespace int
@ -264,8 +264,10 @@ dvd_of_coprime_of_dvd_mul_right H1 (!mul.comm ▸ H2)
theorem gcd_mul_left_cancel_of_coprime {c : } (a : ) {b : } (H : coprime c b) :
gcd (c * a) b = gcd a b :=
begin
revert H, rewrite [↑coprime, ↑gcd, *of_nat_eq_of_nat_iff, nat_abs_mul],
apply nat.gcd_mul_left_cancel_of_coprime
revert H, unfold [coprime, gcd],
rewrite [int_one_eq_nat_one],
rewrite [+of_nat_eq_of_nat_iff, nat_abs_mul],
apply nat.gcd_mul_left_cancel_of_coprime,
end
theorem gcd_mul_right_cancel_of_coprime (a : ) {c b : } (H : coprime c b) :

View file

@ -123,12 +123,12 @@ have a + of_nat (n + m) = a + 0, from
... = a + 0 : by rewrite add_zero,
have of_nat (n + m) = of_nat 0, from add.left_cancel this,
have n + m = 0, from of_nat.inj this,
have n = 0, from nat.eq_zero_of_add_eq_zero_right this,
assert n = 0, from nat.eq_zero_of_add_eq_zero_right this,
show a = b, from
calc
a = a + 0 : add_zero
... = a + n : this
... = b : Hn
a = a + 0 : add_zero
... = a + n : by rewrite this
... = b : Hn
theorem lt.irrefl (a : ) : ¬ a < a :=
(suppose a < a,

View file

@ -718,7 +718,7 @@ lemma count_gt_zero_of_mem : ∀ {a : A} {l : list A}, a ∈ l → count a l > 0
lemma count_eq_zero_of_not_mem {a : A} {l : list A} (h : a ∉ l) : count a l = 0 :=
match count a l with
| 0 := suppose count a l = 0, this
| zero := suppose count a l = zero, this
| (succ n) := suppose count a l = succ n, absurd (mem_of_count_gt_zero (begin rewrite this, exact dec_trivial end)) h
end rfl

View file

@ -309,12 +309,6 @@ protected definition comm_semiring [reducible] [trans_instance] : algebra.comm_s
zero_mul := zero_mul,
mul_zero := mul_zero,
mul_comm := mul.comm⦄
definition nat_has_zero [reducible] [instance] [priority nat.prio] : has_zero nat :=
has_zero.mk zero
definition nat_has_one [reducible] [instance] [priority nat.prio] : has_one nat :=
has_one.mk (succ zero)
end nat
section

View file

@ -51,9 +51,9 @@ theorem add_div_self_left {x : } (z : ) (H : x > 0) : (x + z) div x = succ
theorem add_mul_div_self {x y z : } (H : z > 0) : (x + y * z) div z = x div z + y :=
nat.induction_on y
(calc (x + zero * z) div z = (x + zero) div z : zero_mul
... = x div z : add_zero
... = x div z + zero : add_zero)
(calc (x + 0 * z) div z = (x + 0) div z : zero_mul
... = x div z : add_zero
... = x div z + 0 : add_zero)
(take y,
assume IH : (x + y * z) div z = x div z + y, calc
(x + succ y * z) div z = (x + (y * z + z)) div z : succ_mul
@ -84,7 +84,7 @@ protected definition modulo := fix mod.F
definition nat_has_modulo [reducible] [instance] [priority nat.prio] : has_modulo nat :=
has_modulo.mk nat.modulo
notation a ≡ b `[mod `:100 c `]`:0 := a mod c = b mod c
notation [priority nat.prio] a ≡ b `[mod `:0 c:0 `]` := a mod c = b mod c
theorem modulo_def (x y : nat) : modulo x y = if 0 < y ∧ y ≤ x then modulo (x - y) y else x :=
congr_fun (fix_eq mod.F x) y
@ -115,8 +115,8 @@ theorem add_mod_self_left (x z : ) : (x + z) mod x = z mod x :=
theorem add_mul_mod_self (x y z : ) : (x + y * z) mod z = x mod z :=
nat.induction_on y
(calc (x + zero * z) mod z = (x + zero) mod z : zero_mul
... = x mod z : add_zero)
(calc (x + 0 * z) mod z = (x + 0) mod z : zero_mul
... = x mod z : add_zero)
(take y,
assume IH : (x + y * z) mod z = x mod z,
calc
@ -156,7 +156,6 @@ have H1 : n mod 1 < 1, from !mod_lt !succ_pos,
eq_zero_of_le_zero (le_of_lt_succ H1)
/- properties of div and mod -/
set_option pp.all true
-- the quotient / remainder theorem
theorem eq_div_mul_add_mod (x y : ) : x = x div y * y + x mod y :=
@ -455,7 +454,7 @@ end
/- div and ordering -/
lemma le_of_dvd {m n} : n > 0 → m n → m ≤ n :=
lemma le_of_dvd {m n : nat} : n > 0 → m n → m ≤ n :=
assume (h₁ : n > 0) (h₂ : m n),
assert h₃ : n mod m = 0, from mod_eq_zero_of_dvd h₂,
by_contradiction
@ -550,7 +549,7 @@ begin
1 + k div m = succ (k div m) : add.comm
... ≤ n : succ_le_of_lt H1),
have H3 : n - k div m = n - k div m - 1 + 1, from (sub_add_cancel H2)⁻¹,
have H4 : m > 0, from pos_of_ne_zero (assume H': m = 0, not_lt_zero _ (!zero_mul ▸ H' ▸ H)),
have H4 : m > 0, from pos_of_ne_zero (assume H': m = 0, not_lt_zero k (begin rewrite [H' at H, zero_mul at H], exact H end)),
have H5 : k mod m + 1 ≤ m, from succ_le_of_lt (!mod_lt H4),
have H6 : m - (k mod m + 1) < m, from sub_lt_self H4 !succ_pos,
calc

View file

@ -35,9 +35,9 @@ lemma fib_fast_aux_lemma : ∀ n, (fib_fast_aux (succ n)).1 = (fib_fast_aux n).2
end
theorem fib_eq_fib_fast : ∀ n, fib_fast n = fib n
| 0 := rfl
| 1 := rfl
| (n+2) :=
| 0 := rfl
| 1 := rfl
| (succ (succ n)) :=
begin
have feq : fib_fast n = fib n, from fib_eq_fib_fast n,
have f1eq : fib_fast (succ n) = fib (succ n), from fib_eq_fib_fast (succ n),

View file

@ -29,5 +29,5 @@ lemma two_mul_partial_sum_eq : ∀ n, 2 * partial_sum n = (succ n) * n
theorem partial_sum_eq : ∀ n, partial_sum n = ((n + 1) * n) div 2 :=
take n,
assert h₁ : (2 * partial_sum n) div 2 = ((succ n) * n) div 2, by rewrite two_mul_partial_sum_eq,
assert h₂ : 2 > 0, from dec_trivial,
assert h₂ : (2:nat) > 0, from dec_trivial,
by rewrite [mul_div_cancel_left _ h₂ at h₁]; exact h₁

View file

@ -201,7 +201,8 @@ theorem pos_of_mul_pos_left {a b : } (H : 0 < a * b) : 0 < b :=
theorem pos_of_mul_pos_right {a b : } (H : 0 < a * b) : 0 < a :=
@algebra.pos_of_mul_pos_right _ _ a b H !zero_le
theorem zero_le_one : 0 ≤ 1 := dec_trivial
theorem zero_le_one : (0:nat) ≤ 1 :=
dec_trivial
/- properties specific to nat -/

View file

@ -9,7 +9,7 @@ import data.nat.power logic.identities
namespace nat
open decidable
open - [notations] algebra
open algebra
definition even (n : nat) := n mod 2 = 0
@ -63,10 +63,13 @@ iff.mp !odd_iff_not_even this
lemma odd_succ_of_even {n} : even n → odd (succ n) :=
suppose even n,
have n ≡ 0 [mod 2], from this,
have n+1 ≡ 0+1 [mod 2], from add_mod_eq_add_mod_right 1 this,
have h : n+1 ≡ 1 [mod 2], from this,
by_contradiction (suppose ¬ odd (succ n),
assert 0 = 1, from calc
0 = (n+1) mod 2 : even_of_not_odd this
... = 1 mod 2 : add_mod_eq_add_mod_right 1 `even n`,
have n+1 ≡ 0 [mod 2], from even_of_not_odd this,
have 1 ≡ 0 [mod 2], from eq.trans (eq.symm h) this,
assert 1 = 0, from this,
by contradiction)
lemma eq_1_of_ne_0_lt_2 : ∀ {n : nat}, n ≠ 0 → n < 2 → n = 1
@ -271,7 +274,7 @@ assume h₁ h₂,
(suppose odd n, or.elim (em (even m))
(suppose even m, absurd `odd n` (not_odd_of_even (iff.mpr h₂ `even m`)))
(suppose odd m,
assert d : 1 div 2 = 0, from dec_trivial,
assert d : 1 div 2 = (0:nat), from dec_trivial,
obtain w₁ (hw₁ : n = 2*w₁ + 1), from exists_of_odd `odd n`,
obtain w₂ (hw₂ : m = 2*w₂ + 1), from exists_of_odd `odd m`,
begin

View file

@ -70,7 +70,7 @@ private theorem le_squared : ∀ (n : nat), n ≤ n*n
assert aux₂ : 1 * succ n ≤ succ n * succ n, from mul_le_mul aux₁ !le.refl,
by rewrite [one_mul at aux₂]; exact aux₂
private theorem lt_squared : ∀ {n}, n > 1 → n < n * n
private theorem lt_squared : ∀ {n : nat}, n > 1 → n < n * n
| 0 h := absurd h dec_trivial
| 1 h := absurd h dec_trivial
| (succ (succ n)) h :=

View file

@ -48,6 +48,9 @@ has_le.mk pnat.le
definition pnat_has_lt [instance] [reducible] : has_lt pnat :=
has_lt.mk pnat.lt
definition pnat_has_one [instance] [reducible] : has_one pnat :=
has_one.mk (pos (1:nat) dec_trivial)
lemma mul.def (p q : +) : p * q = tag (p~ * q~) (mul_pos (pnat_pos p) (pnat_pos q)) :=
rfl

View file

@ -195,8 +195,8 @@ theorem inv_equiv_inv : ∀{a b : prerat}, a ≡ b → inv a ≡ inv b
have bn_zero : bn = 0, from num_eq_zero_of_equiv H an_zero,
eq.subst (calc
inv (mk an ad adp) = inv (mk 0 ad adp) : {an_zero}
... = zero : inv_zero
... = inv (mk 0 bd bdp) : inv_zero
... = zero : inv_zero adp
... = inv (mk 0 bd bdp) : inv_zero bdp
... = inv (mk bn bd bdp) : bn_zero) !equiv.refl)
(assume an_pos : an > 0,
have bn_pos : bn > 0, from num_pos_of_equiv H an_pos,
@ -356,6 +356,20 @@ definition of_int [coercion] (i : ) : := ⟦prerat.of_int i⟧
definition of_nat [coercion] (n : ) : := nat.to.rat n
definition of_num [coercion] [reducible] (n : num) : := num.to.rat n
protected definition prio := num.pred int.prio
definition rat_has_zero [reducible] [instance] [priority rat.prio] : has_zero rat :=
has_zero.mk (0:int)
definition rat_has_one [reducible] [instance] [priority rat.prio] : has_one rat :=
has_one.mk (1:int)
theorem rat_zero_eq_int_zero : (0:rat) = of_int (0:int) :=
rfl
theorem rat_one_eq_int_one : (1:rat) = of_int (1:int) :=
rfl
protected definition add : :=
quot.lift₂
(λ a b : prerat, ⟦prerat.add a b⟧)
@ -387,8 +401,6 @@ definition denom (a : ) : := prerat.denom (reduce a)
theorem denom_pos (a : ): denom a > 0 :=
prerat.denom_pos (reduce a)
protected definition prio := num.pred int.prio
definition rat_has_add [reducible] [instance] [priority rat.prio] : has_add rat :=
has_add.mk rat.add
@ -508,7 +520,7 @@ take a b, quot.rec_on_subsingleton₂ a b
then decidable.inl (quot.sound H)
else decidable.inr (assume H1, H (quot.exact H1)))
theorem inv_zero : inv 0 = 0 :=
theorem inv_zero : inv 0 = (0 : ) :=
quot.sound (prerat.inv_zero' ▸ !prerat.equiv.refl)
theorem quot_reduce (a : ) : ⟦reduce a⟧ = a :=

View file

@ -360,12 +360,15 @@ eq_of_sub_eq_zero this
section
open int
set_option pp.numerals false
set_option pp.implicit true
theorem num_nonneg_of_nonneg {q : } (H : q ≥ 0) : num q ≥ 0 :=
have of_int (num q) ≥ of_int 0,
begin
rewrite [-mul_denom],
apply mul_nonneg H,
rewrite [of_int_le_of_int_iff],
rewrite [rat_zero_eq_int_zero, of_int_le_of_int_iff],
exact int.le_of_lt !denom_pos
end,
show num q ≥ 0, from le_of_of_int_le_of_int this
@ -375,7 +378,7 @@ section
begin
rewrite [-mul_denom],
apply mul_pos H,
rewrite [of_int_lt_of_int_iff],
rewrite [rat_zero_eq_int_zero, of_int_lt_of_int_iff],
exact !denom_pos
end,
show num q > 0, from lt_of_of_int_lt_of_int this

View file

@ -65,8 +65,6 @@ theorem to_finset_insert (a : A) (s : set A) [fins : finite s] :
to_finset (insert a s) = finset.insert a (to_finset s) :=
by apply to_finset_eq_of_to_set_eq; rewrite [finset.to_set_insert, to_set_to_finset]
example : finite '{1, 2, 3} := _
theorem finite_union [instance] (s t : set A) [fins : finite s] [fint : finite t] :
finite (s t) :=
exists.intro (#finset to_finset s to_finset t)

View file

@ -478,7 +478,7 @@ theorem take_lemma (s₁ s₂ : stream A) : (∀ (n : nat), approx n s₁ = appr
begin
intro h, apply stream.ext, intro n,
induction n with n ih,
{injection (h 1), assumption},
{injection (h 1) with aux, exact aux},
{have h₁ : some (nth (succ n) s₁) = some (nth (succ n) s₂), by rewrite [-*nth_approx, h (succ (succ n))],
injection h₁, assumption}
end

View file

@ -5,6 +5,6 @@ Authors: Leonardo de Moura
-/
prelude
import init.datatypes init.reserved_notation init.tactic init.logic
import init.relation init.wf init.nat init.wf_k init.prod init.priority
import init.relation init.wf init.nat init.wf_k init.prod
import init.bool init.num init.sigma init.measurable init.setoid init.quot
import init.funext init.function init.subtype init.classical

View file

@ -91,10 +91,10 @@ namespace nat
theorem le_succ_of_pred_le {n m : } : pred n ≤ m → n ≤ succ m :=
nat.cases_on n le.step (λa, succ_le_succ)
theorem not_succ_le_zero (n : ) : ¬succ n ≤ zero :=
theorem not_succ_le_zero (n : ) : ¬succ n ≤ 0 :=
by intro H; cases H
theorem succ_le_zero_iff_false (n : ) : succ n ≤ zero ↔ false :=
theorem succ_le_zero_iff_false (n : ) : succ n ≤ 0 ↔ false :=
iff_false_intro !not_succ_le_zero
theorem not_succ_le_self : Π {n : }, ¬succ n ≤ n :=
@ -149,9 +149,9 @@ namespace nat
theorem lt.asymm {n m : } (H1 : n < m) : ¬ m < n :=
le_lt_antisymm (le_of_lt H1)
theorem not_lt_zero (a : ) : ¬ a < zero := !not_succ_le_zero
theorem not_lt_zero (a : ) : ¬ a < 0 := !not_succ_le_zero
theorem lt_zero_iff_false [simp] (a : ) : a < zero ↔ false :=
theorem lt_zero_iff_false [simp] (a : ) : a < 0 ↔ false :=
iff_false_intro (not_lt_zero a)
theorem eq_or_lt_of_le {a b : } (H : a ≤ b) : a = b a < b :=
@ -225,10 +225,10 @@ namespace nat
theorem sub_eq_succ_sub_succ (a b : ) : a - b = succ a - succ b :=
eq.symm !succ_sub_succ_eq_sub
theorem zero_sub_eq_zero [simp] (a : ) : zero - a = zero :=
theorem zero_sub_eq_zero [simp] (a : ) : 0 - a = 0 :=
nat.rec rfl (λ a, congr_arg pred) a
theorem zero_eq_zero_sub (a : ) : zero = zero - a :=
theorem zero_eq_zero_sub (a : ) : 0 = 0 - a :=
eq.symm !zero_sub_eq_zero
theorem sub_le (a b : ) : a - b ≤ a :=
@ -237,7 +237,7 @@ namespace nat
theorem sub_le_iff_true [simp] (a b : ) : a - b ≤ a ↔ true :=
iff_true_intro (sub_le a b)
theorem sub_lt {a b : } (H1 : zero < a) (H2 : zero < b) : a - b < a :=
theorem sub_lt {a b : } (H1 : 0 < a) (H2 : 0 < b) : a - b < a :=
!nat.cases_on (λh, absurd h !lt.irrefl)
(λa h, succ_le_succ (!nat.cases_on (λh, absurd h !lt.irrefl)
(λb c, eq.substr !succ_sub_succ_eq_sub !sub_le) H2)) H1

View file

@ -5,7 +5,7 @@ Authors: Leonardo de Moura
-/
prelude
import init.logic init.bool init.priority
import init.logic init.bool
open bool
definition pos_num.is_inhabited [instance] : inhabited pos_num :=
@ -89,26 +89,3 @@ end num
definition num_has_sub [instance] [reducible] : has_sub num :=
has_sub.mk num.sub
-- the coercion from num to nat is defined here,
-- so that it can already be used in init.tactic
namespace nat
protected definition prio := num.add std.priority.default 100
protected definition add (a b : nat) : nat :=
nat.rec_on b a (λ b₁ r, succ r)
definition nat_has_zero [reducible] [instance] : has_zero nat :=
has_zero.mk nat.zero
definition nat_has_one [reducible] [instance] : has_one nat :=
has_one.mk (nat.succ (nat.zero))
definition nat_has_add [reducible] [instance] [priority nat.prio] : has_add nat :=
has_add.mk nat.add
definition of_num [coercion] (n : num) : nat :=
num.rec zero
(λ n, pos_num.rec (succ zero) (λ n r, r + r + (succ zero)) (λ n r, r + r) n) n
end nat
attribute nat.of_num [reducible] -- of_num is also reducible if namespace "nat" is not opened

View file

@ -1,9 +0,0 @@
/-
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import init.reserved_notation
definition std.priority.default : num := 1000
definition std.priority.max : num := 4294967295

View file

@ -56,7 +56,7 @@ namespace pos_num
pos_num.rec_on a tt (λn r, ff) (λn r, ff)
definition pred (a : pos_num) : pos_num :=
pos_num.rec_on a one (λn r, bit0 n) (λn r, bool.rec_on (is_one n) one (bit1 r))
pos_num.rec_on a one (λn r, bit0 n) (λn r, bool.rec_on (is_one n) (bit1 r) one)
definition size (a : pos_num) : pos_num :=
pos_num.rec_on a one (λn r, succ r) (λn r, succ r)
@ -88,6 +88,29 @@ end num
definition num_has_add [reducible] [instance] : has_add num :=
has_add.mk num.add
definition std.priority.default : num := 1000
definition std.priority.max : num := 4294967295
namespace nat
protected definition prio := num.add std.priority.default 100
protected definition add (a b : nat) : nat :=
nat.rec_on b a (λ b₁ r, succ r)
definition nat_has_zero [reducible] [instance] : has_zero nat :=
has_zero.mk nat.zero
definition nat_has_one [reducible] [instance] : has_one nat :=
has_one.mk (nat.succ (nat.zero))
definition nat_has_add [reducible] [instance] [priority nat.prio] : has_add nat :=
has_add.mk nat.add
definition of_num (n : num) : nat :=
num.rec zero
(λ n, pos_num.rec (succ zero) (λ n r, add (add r r) (succ zero)) (λ n r, add r r) n) n
end nat
/-
Global declarations of right binding strength

View file

@ -7,7 +7,7 @@ Binomial coefficients, "n choose k".
-/
import data.nat.div data.nat.fact data.finset
open decidable
open - [notations] algebra
open algebra
namespace nat
@ -58,7 +58,8 @@ theorem choose_one_right (n : ) : choose n 1 = n :=
begin
induction n with [n, ih],
{apply rfl},
rewrite [choose_succ_succ, ih, choose_zero_right]
change choose (succ n) (succ 0) = succ n,
krewrite [choose_succ_succ, ih, choose_zero_right]
end
theorem choose_pos {n : } : ∀ {k : }, k ≤ n → choose n k > 0 :=

View file

@ -86,17 +86,19 @@ section
have nat_abs b = 1, from
by_cases
(suppose q = 0, by rewrite this)
(suppose q ≠ 0,
have ane0 : a ≠ 0, from
suppose aeq0 : a = 0,
have q = 0, begin rewrite eq_num_div_denom, rewrite aeq0, krewrite zero_div end,
absurd `q = 0` `q ≠ 0`,
have nat_abs a ≠ 0, from
suppose nat_abs a = 0,
have a = 0, from eq_zero_of_nat_abs_eq_zero this,
absurd `a = 0` `a ≠ 0`,
show nat_abs b = 1, from (root_irrational npos (pos_of_ne_zero this) H₂ H₁)),
show b = 1, using this, by rewrite [-of_nat_nat_abs_of_nonneg (le_of_lt !denom_pos), this]
(suppose qne0 : q ≠ 0,
using H₁ H₂, begin
have ane0 : a ≠ 0, from
suppose aeq0 : a = 0,
have qeq0 : q = 0, begin rewrite eq_num_div_denom, rewrite aeq0, krewrite algebra.zero_div end,
absurd qeq0 qne0,
have nat_abs a ≠ 0, from
suppose nat_abs a = 0,
have aeq0 : a = 0, from eq_zero_of_nat_abs_eq_zero this,
absurd aeq0 ane0,
show nat_abs b = 1, from (root_irrational npos (pos_of_ne_zero this) H₂ H₁)
end),
show b = 1, using this, begin rewrite [-of_nat_nat_abs_of_nonneg (le_of_lt !denom_pos), this] end
theorem eq_num_pow_of_pow_eq {q : } {n : } {c : } (npos : n > 0) (H : q^n = c) :
c = (num q)^n :=

View file

@ -191,8 +191,8 @@ lemma dvd_or_dvd_of_prime_of_dvd_mul {p m n : nat} : prime p → p m * n →
lemma dvd_of_prime_of_dvd_pow {p m : nat} : ∀ {n}, prime p → p m^n → p m
| 0 hp hd :=
assert p = 1, from eq_one_of_dvd_one hd,
have 1 ≥ 2, by rewrite -this; apply ge_two_of_prime hp,
assert p = 1, from eq_one_of_dvd_one hd,
have (1:nat) ≥ 2, begin rewrite -this at {1}, apply ge_two_of_prime hp end,
absurd this dec_trivial
| (succ n) hp hd :=
have p (m^n)*m, by rewrite [pow_succ' at hd]; exact hd,

View file

@ -489,6 +489,8 @@ auto pretty_fn::pp_child(expr const & e, unsigned bp, bool ignore_hide) -> resul
if (is_app(e)) {
if (auto r = pp_local_ref(e))
return add_paren_if_needed(*r, bp);
if (m_numerals)
if (auto n = to_num(e)) return pp_num(*n);
expr const & f = app_fn(e);
if (auto it = is_abbreviated(f)) {
return pp_abbreviation(e, *it, true, bp, ignore_hide);
@ -982,6 +984,8 @@ auto pretty_fn::pp_notation_child(expr const & e, unsigned lbp, unsigned rbp) ->
if (auto it = is_abbreviated(e))
return pp_abbreviation(e, *it, false, rbp);
if (is_app(e)) {
if (m_numerals)
if (auto n = to_num(e)) return pp_num(*n);
expr const & f = app_fn(e);
if (auto it = is_abbreviated(f)) {
return pp_abbreviation(e, *it, true, rbp);
@ -1233,9 +1237,10 @@ auto pretty_fn::pp(expr const & e, bool ignore_hide) -> result {
flet<unsigned> let_d(m_depth, m_depth+1);
m_num_steps++;
if (m_numerals)
if (auto n = to_num(e)) return pp_num(*n);
if (auto n = is_abbreviated(e))
return pp_abbreviation(e, *n, false);
if (auto r = pp_notation(e))
return *r;
@ -1245,8 +1250,6 @@ auto pretty_fn::pp(expr const & e, bool ignore_hide) -> result {
if (is_let(e)) return pp_let(e);
if (is_typed_expr(e)) return pp(get_typed_expr_expr(e));
if (is_let_value(e)) return pp(get_let_value_expr(e));
if (m_numerals)
if (auto n = to_num(e)) return pp_num(*n);
if (m_num_nat_coe)
if (auto k = to_unsigned(e))
return format(*k);

View file

@ -6,6 +6,7 @@ Author: Leonardo de Moura
*/
#include "kernel/type_checker.h"
#include "library/num.h"
#include "library/util.h"
#include "library/constants.h"
namespace lean {
@ -42,6 +43,18 @@ optional<expr> is_bit1(expr const & e) {
return some_expr(app_arg(e));
}
optional<expr> unfold_num_app(environment const & env, expr const & e) {
if (is_zero(e) || is_one(e) || is_bit0(e) || is_bit1(e)) {
return unfold_app(env, e);
} else {
return none_expr();
}
}
bool is_numeral_const_name(name const & n) {
return n == get_zero_name() || n == get_one_name() || n == get_bit0_name() || n == get_bit1_name();
}
static bool is_num(expr const & e, bool first) {
if (is_zero(e))
return first;

View file

@ -20,6 +20,12 @@ bool is_one(expr const & e);
optional<expr> is_bit0(expr const & e);
optional<expr> is_bit1(expr const & e);
/** \brief Return true iff \c n is zero, one, bit0 or bit1 */
bool is_numeral_const_name(name const & n);
/** Unfold \c e it is is_zero, is_one, is_bit0 or is_bit1 application */
optional<expr> unfold_num_app(environment const & env, expr const & e);
/** \brief If the given expression encodes a numeral, then convert it back to mpz numeral.
\see from_num */
optional<mpz> to_num(expr const & e);

View file

@ -33,6 +33,7 @@ Author: Leonardo de Moura
#include "library/unfold_macros.h"
#include "library/generic_exception.h"
#include "library/class_instance_synth.h"
#include "library/num.h"
#include "library/tactic/clear_tactic.h"
#include "library/tactic/trace_tactic.h"
#include "library/tactic/rewrite_tactic.h"
@ -489,7 +490,19 @@ public:
constraint_seq cs;
expr p1 = m_tc.whnf(p, cs);
expr t1 = m_tc.whnf(t, cs);
return !cs && (p1 != p || t1 != t) && ctx.match(p1, t1);
if (!cs && (p1 != p || t1 != t) && ctx.match(p1, t1)) {
return true;
} else if (!has_expr_metavar(p1)) {
// special support for numerals
if (auto p2 = unfold_num_app(m_tc.env(), p1)) {
// unfold nested projection
if (auto p3 = unfold_app(m_tc.env(), *p2)) {
p3 = m_tc.whnf(*p3, cs);
return !cs && p1 != *p3 && ctx.match(*p3, t1);
}
}
}
return false;
} catch (exception&) {
return false;
}
@ -1592,6 +1605,13 @@ class rewrite_fn {
}
}
type_checker_ptr mk_tc(bool full) {
auto aux_pred = full ? mk_irreducible_pred(m_env) : mk_not_quasireducible_pred(m_env);
return mk_type_checker(m_env, m_ngen.mk_child(),[=](name const & n) {
return aux_pred(n) && !is_numeral_const_name(n);
});
}
void process_failure(expr const & elem, bool type_error, kernel_exception * ex = nullptr) {
std::shared_ptr<kernel_exception> saved_ex;
if (ex)
@ -1641,7 +1661,7 @@ public:
rewrite_fn(environment const & env, io_state const & ios, elaborate_fn const & elab, proof_state const & ps,
bool full, bool keyed):
m_env(env), m_ios(ios), m_elab(elab), m_ps(ps), m_ngen(ps.get_ngen()),
m_tc(mk_type_checker(m_env, m_ngen.mk_child(), full ? UnfoldSemireducible : UnfoldQuasireducible)),
m_tc(mk_tc(full)),
m_matcher_tc(mk_matcher_tc(full)),
m_relaxed_tc(mk_type_checker(m_env, m_ngen.mk_child())),
m_mplugin(m_ios, *m_matcher_tc) {