feat(library): define custom recursors for nat, and minimize the use of krewrite
This commit is contained in:
parent
1f5d950b46
commit
ce21996635
16 changed files with 106 additions and 66 deletions
|
@ -58,7 +58,7 @@ iff.intro
|
|||
|
||||
private lemma odd_of_zero_mem (s : nat) : 0 ∈ of_nat s ↔ odd s :=
|
||||
begin
|
||||
unfold of_nat, krewrite [mem_sep_eq, pow_zero, div_one, mem_upto_eq],
|
||||
unfold of_nat, rewrite [mem_sep_eq, pow_zero, div_one, mem_upto_eq],
|
||||
show 0 < succ s ∧ odd s ↔ odd s, from
|
||||
iff.intro
|
||||
(assume h, and.right h)
|
||||
|
@ -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, exact absurd rfl `zero ≠ 0`,
|
||||
cases a with a, exact absurd rfl `0 ≠ 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)`
|
||||
|
|
|
@ -145,6 +145,12 @@ theorem eq_zero_of_nat_abs_eq_zero : Π {a : ℤ}, nat_abs a = 0 → a = 0
|
|||
| (of_nat m) H := congr_arg of_nat H
|
||||
| -[1+ m'] H := absurd H !succ_ne_zero
|
||||
|
||||
theorem nat_abs_zero : nat_abs (0:int) = (0:nat) :=
|
||||
rfl
|
||||
|
||||
theorem nat_abs_one : nat_abs (1:int) = (1:nat) :=
|
||||
rfl
|
||||
|
||||
/- int is a quotient of ordered pairs of natural numbers -/
|
||||
|
||||
protected definition equiv (p q : ℕ × ℕ) : Prop := pr1 p + pr2 q = pr2 p + pr1 q
|
||||
|
@ -526,7 +532,7 @@ protected definition integral_domain [reducible] [trans_instance] : algebra.inte
|
|||
⦃algebra.integral_domain,
|
||||
add := int.add,
|
||||
add_assoc := add.assoc,
|
||||
zero := zero,
|
||||
zero := 0,
|
||||
zero_add := zero_add,
|
||||
add_zero := add_zero,
|
||||
neg := int.neg,
|
||||
|
|
|
@ -56,15 +56,15 @@ notation [priority int.prio] a ≡ b `[mod `:0 c:0 `]` := a mod c = b mod c
|
|||
|
||||
theorem of_nat_div (m n : nat) : of_nat (m div n) = (of_nat m) div (of_nat n) :=
|
||||
nat.cases_on n
|
||||
(begin krewrite [of_nat_div_eq, sign_zero, zero_mul, nat.div_zero] end)
|
||||
(take (n : nat), by krewrite [of_nat_div_eq, sign_of_succ, one_mul])
|
||||
(begin rewrite [of_nat_div_eq, of_nat_zero, sign_zero, zero_mul, nat.div_zero] end)
|
||||
(take (n : nat), by rewrite [of_nat_div_eq, sign_of_succ, one_mul])
|
||||
|
||||
theorem neg_succ_of_nat_div (m : nat) {b : ℤ} (H : b > 0) :
|
||||
-[1+m] div b = -(m div b + 1) :=
|
||||
calc
|
||||
-[1+m] div b = sign b * _ : rfl
|
||||
... = -[1+(m div (nat_abs b))] : by krewrite [sign_of_pos H, one_mul]
|
||||
... = -(m div b + 1) : by krewrite [of_nat_div_eq, sign_of_pos H, one_mul]
|
||||
... = -[1+(m div (nat_abs b))] : by rewrite [sign_of_pos H, one_mul]
|
||||
... = -(m div b + 1) : by rewrite [of_nat_div_eq, sign_of_pos H, one_mul]
|
||||
|
||||
theorem div_neg (a b : ℤ) : a div -b = -(a div b) :=
|
||||
begin
|
||||
|
@ -100,10 +100,10 @@ calc
|
|||
... < 0 : neg_neg_of_pos this
|
||||
|
||||
theorem zero_div (b : ℤ) : 0 div b = 0 :=
|
||||
by krewrite [of_nat_div_eq, nat.zero_div, mul_zero]
|
||||
by rewrite [of_nat_div_eq, nat.zero_div, of_nat_zero, mul_zero]
|
||||
|
||||
theorem div_zero (a : ℤ) : a div 0 = 0 :=
|
||||
by krewrite [divide.def, sign_zero, zero_mul]
|
||||
by rewrite [divide.def, sign_zero, zero_mul]
|
||||
|
||||
theorem div_one (a : ℤ) : a div 1 = a :=
|
||||
assert (1 : int) > 0, from dec_trivial,
|
||||
|
@ -137,7 +137,7 @@ lt.by_cases
|
|||
assert a < -b, from abs_of_neg this ▸ H2,
|
||||
calc
|
||||
a div b = - (a div -b) : by rewrite [div_neg, neg_neg]
|
||||
... = 0 : by krewrite [div_eq_zero_of_lt H1 this, neg_zero])
|
||||
... = 0 : by rewrite [div_eq_zero_of_lt H1 this, neg_zero])
|
||||
(suppose b = 0, this⁻¹ ▸ !div_zero)
|
||||
(suppose b > 0,
|
||||
have a < b, from abs_of_pos this ▸ H2,
|
||||
|
@ -277,10 +277,10 @@ theorem mod_abs (a b : ℤ) : a mod (abs b) = a mod b :=
|
|||
abs.by_cases rfl !mod_neg
|
||||
|
||||
theorem zero_mod (b : ℤ) : 0 mod b = 0 :=
|
||||
by krewrite [(modulo.def), zero_div, zero_mul, sub_zero]
|
||||
by rewrite [(modulo.def), zero_div, zero_mul, sub_zero]
|
||||
|
||||
theorem mod_zero (a : ℤ) : a mod 0 = a :=
|
||||
by krewrite [(modulo.def), mul_zero, sub_zero]
|
||||
by rewrite [(modulo.def), mul_zero, sub_zero]
|
||||
|
||||
theorem mod_one (a : ℤ) : a mod 1 = 0 :=
|
||||
calc
|
||||
|
@ -340,7 +340,7 @@ have H2 : a mod (abs b) < abs b, from
|
|||
|
||||
theorem add_mul_mod_self {a b c : ℤ} : (a + b * c) mod c = a mod c :=
|
||||
decidable.by_cases
|
||||
(assume cz : c = 0, by krewrite [cz, mul_zero, add_zero])
|
||||
(assume cz : c = 0, by rewrite [cz, mul_zero, add_zero])
|
||||
(assume cnz, by rewrite [(modulo.def), !add_mul_div_self cnz, mul.right_distrib,
|
||||
sub_add_eq_sub_sub_swap, add_sub_cancel])
|
||||
|
||||
|
@ -421,7 +421,7 @@ lt.by_cases
|
|||
... = b div c : by rewrite [div_neg, neg_neg])
|
||||
(assume H1 : c = 0,
|
||||
calc
|
||||
a * b div (a * c) = 0 : by krewrite [H1, mul_zero, div_zero]
|
||||
a * b div (a * c) = 0 : by rewrite [H1, mul_zero, div_zero]
|
||||
... = b div c : by rewrite [H1, div_zero])
|
||||
(assume H1 : c > 0,
|
||||
mul_div_mul_of_pos_aux _ H H1)
|
||||
|
@ -476,7 +476,7 @@ lt.by_cases
|
|||
... ≤ abs a : H _ _ (neg_pos_of_neg H1))
|
||||
(assume H1 : b = 0,
|
||||
calc
|
||||
abs (a div b) = 0 : by krewrite [H1, div_zero, abs_zero]
|
||||
abs (a div b) = 0 : by rewrite [H1, div_zero, abs_zero]
|
||||
... ≤ abs a : abs_nonneg)
|
||||
(assume H1 : b > 0, H _ _ H1)
|
||||
|
||||
|
@ -538,7 +538,7 @@ theorem mul_div_cancel' {a b : ℤ} (H : a ∣ b) : a * (b div a) = b :=
|
|||
|
||||
theorem mul_div_assoc (a : ℤ) {b c : ℤ} (H : c ∣ b) : (a * b) div c = a * (b div c) :=
|
||||
decidable.by_cases
|
||||
(assume cz : c = 0, by krewrite [cz, *div_zero, mul_zero])
|
||||
(assume cz : c = 0, by rewrite [cz, *div_zero, mul_zero])
|
||||
(assume cnz : c ≠ 0,
|
||||
obtain d (H' : b = d * c), from exists_eq_mul_left_of_dvd H,
|
||||
by rewrite [H', -mul.assoc, *(!mul_div_cancel cnz)])
|
||||
|
@ -585,7 +585,7 @@ div_eq_of_eq_mul_right H1 (!mul.comm ▸ H2)
|
|||
|
||||
theorem neg_div_of_dvd {a b : ℤ} (H : b ∣ a) : -a div b = -(a div b) :=
|
||||
decidable.by_cases
|
||||
(assume H1 : b = 0, by krewrite [H1, *div_zero, neg_zero])
|
||||
(assume H1 : b = 0, by rewrite [H1, *div_zero, neg_zero])
|
||||
(assume H1 : b ≠ 0,
|
||||
dvd.elim H
|
||||
(take c, assume H' : a = b * c,
|
||||
|
|
|
@ -21,14 +21,17 @@ of_nat_nonneg (nat.gcd (nat_abs a) (nat_abs b))
|
|||
theorem gcd.comm (a b : ℤ) : gcd a b = gcd b a :=
|
||||
by rewrite [↑gcd, nat.gcd.comm]
|
||||
|
||||
set_option pp.all true
|
||||
set_option pp.numerals false
|
||||
|
||||
theorem gcd_zero_right (a : ℤ) : gcd a 0 = abs a :=
|
||||
by krewrite [↑gcd, nat.gcd_zero_right, of_nat_nat_abs]
|
||||
by rewrite [↑gcd, nat_abs_zero, nat.gcd_zero_right, of_nat_nat_abs]
|
||||
|
||||
theorem gcd_zero_left (a : ℤ) : gcd 0 a = abs a :=
|
||||
by rewrite [gcd.comm, gcd_zero_right]
|
||||
|
||||
theorem gcd_one_right (a : ℤ) : gcd a 1 = 1 :=
|
||||
by krewrite [↑gcd, nat.gcd_one_right]
|
||||
by rewrite [↑gcd, nat_abs_one, nat.gcd_one_right]
|
||||
|
||||
theorem gcd_one_left (a : ℤ) : gcd 1 a = 1 :=
|
||||
by rewrite [gcd.comm, gcd_one_right]
|
||||
|
@ -191,13 +194,13 @@ theorem lcm.comm (a b : ℤ) : lcm a b = lcm b a :=
|
|||
by rewrite [↑lcm, nat.lcm.comm]
|
||||
|
||||
theorem lcm_zero_left (a : ℤ) : lcm 0 a = 0 :=
|
||||
by krewrite [↑lcm, nat.lcm_zero_left]
|
||||
by rewrite [↑lcm, nat_abs_zero, nat.lcm_zero_left]
|
||||
|
||||
theorem lcm_zero_right (a : ℤ) : lcm a 0 = 0 :=
|
||||
!lcm.comm ▸ !lcm_zero_left
|
||||
|
||||
theorem lcm_one_left (a : ℤ) : lcm 1 a = abs a :=
|
||||
by krewrite [↑lcm, nat.lcm_one_left, of_nat_nat_abs]
|
||||
by rewrite [↑lcm, nat_abs_one, nat.lcm_one_left, of_nat_nat_abs]
|
||||
|
||||
theorem lcm_one_right (a : ℤ) : lcm a 1 = abs a :=
|
||||
!lcm.comm ▸ !lcm_one_left
|
||||
|
@ -212,7 +215,7 @@ theorem lcm_abs_abs (a b : ℤ) : lcm (abs a) (abs b) = lcm a b :=
|
|||
by rewrite [↑lcm, *nat_abs_abs]
|
||||
|
||||
theorem lcm_self (a : ℤ) : lcm a a = abs a :=
|
||||
by krewrite [↑lcm, nat.lcm_self, of_nat_nat_abs]
|
||||
by rewrite [↑lcm, nat.lcm_self, of_nat_nat_abs]
|
||||
|
||||
theorem dvd_lcm_left (a b : ℤ) : a ∣ lcm a b :=
|
||||
by rewrite [↑lcm, -abs_dvd_iff, -of_nat_nat_abs, of_nat_dvd_of_nat_iff]; apply nat.dvd_lcm_left
|
||||
|
|
|
@ -29,14 +29,14 @@ nat.induction_on x
|
|||
(λ y, nat.induction_on y
|
||||
rfl
|
||||
(λ y₁ ih, calc
|
||||
zero + succ y₁ = succ (zero + y₁) : rfl
|
||||
... = succ (zero ⊕ y₁) : {ih}
|
||||
... = zero ⊕ (succ y₁) : rfl))
|
||||
0 + succ y₁ = succ (0 + y₁) : rfl
|
||||
... = succ (0 ⊕ y₁) : {ih}
|
||||
... = 0 ⊕ (succ y₁) : rfl))
|
||||
(λ x₁ ih₁ y, nat.induction_on y
|
||||
(calc
|
||||
succ x₁ + zero = succ (x₁ + zero) : rfl
|
||||
... = succ (x₁ ⊕ zero) : {ih₁ zero}
|
||||
... = succ x₁ ⊕ zero : rfl)
|
||||
succ x₁ + 0 = succ (x₁ + 0) : rfl
|
||||
... = succ (x₁ ⊕ 0) : {ih₁ 0}
|
||||
... = succ x₁ ⊕ 0 : rfl)
|
||||
(λ y₁ ih₂, calc
|
||||
succ x₁ + succ y₁ = succ (succ x₁ + y₁) : rfl
|
||||
... = succ (succ x₁ ⊕ y₁) : {ih₂}
|
||||
|
|
|
@ -19,13 +19,13 @@ algebra.pow_le_pow_of_le i !zero_le H
|
|||
theorem eq_zero_of_pow_eq_zero {a m : ℕ} (H : a^m = 0) : a = 0 :=
|
||||
or.elim (eq_zero_or_pos m)
|
||||
(suppose m = 0,
|
||||
by krewrite [`m = 0` at H, pow_zero at H]; contradiction)
|
||||
by rewrite [`m = 0` at H, pow_zero at H]; contradiction)
|
||||
(suppose m > 0,
|
||||
have h₁ : ∀ m, a^succ m = 0 → a = 0,
|
||||
begin
|
||||
intro m,
|
||||
induction m with m ih,
|
||||
{krewrite pow_one; intros; assumption},
|
||||
{rewrite pow_one; intros; assumption},
|
||||
rewrite pow_succ,
|
||||
intro H,
|
||||
cases eq_zero_or_eq_zero_of_mul_eq_zero H with h₃ h₄,
|
||||
|
@ -54,7 +54,7 @@ theorem le_pow_self {x : ℕ} (H : x > 1) : ∀ i, i ≤ x^i
|
|||
|
||||
theorem mul_self_eq_pow_2 (a : nat) : a * a = a ^ 2 :=
|
||||
show a * a = a ^ (succ (succ zero)), from
|
||||
by krewrite [*pow_succ, *pow_zero, mul_one] -- TODO(Leo): krewrite -> rewrite after new numeral encoding
|
||||
by rewrite [*pow_succ, *pow_zero, mul_one]
|
||||
|
||||
theorem pow_cancel_left : ∀ {a b c : nat}, a > 1 → a^b = a^c → b = c
|
||||
| a 0 0 h₁ h₂ := rfl
|
||||
|
@ -72,8 +72,7 @@ theorem pow_cancel_left : ∀ {a b c : nat}, a > 1 → a^b = a^c → b = c
|
|||
by rewrite [pow_cancel_left h₁ this]
|
||||
|
||||
theorem pow_div_cancel : ∀ {a b : nat}, a ≠ 0 → (a ^ succ b) div a = a ^ b
|
||||
-- TODO(Leo): krewrite -> rewrite after new numeral encoding
|
||||
| a 0 h := by krewrite [pow_succ, pow_zero, mul_one, div_self (pos_of_ne_zero h)]
|
||||
| a 0 h := by rewrite [pow_succ, pow_zero, mul_one, div_self (pos_of_ne_zero h)]
|
||||
| a (succ b) h := by rewrite [pow_succ, mul_div_cancel_left _ (pos_of_ne_zero h)]
|
||||
|
||||
lemma dvd_pow : ∀ (i : nat) {n : nat}, n > 0 → i ∣ i^n
|
||||
|
|
|
@ -234,7 +234,7 @@ have hsimp [visible] : ∀ a b : ℚ, (a + a) + (b + b) = (a + b) + (a + b),
|
|||
by rewrite [-add_halves m, -add_halves n, hsimp]
|
||||
|
||||
theorem inv_mul_eq_mul_inv {p q : ℕ+} : (p * q)⁻¹ = p⁻¹ * q⁻¹ :=
|
||||
begin krewrite [↑inv, pnat_to_rat_mul, algebra.one_div_mul_one_div] end
|
||||
begin rewrite [↑inv, pnat_to_rat_mul, algebra.one_div_mul_one_div] end
|
||||
|
||||
theorem inv_mul_le_inv (p q : ℕ+) : (p * q)⁻¹ ≤ q⁻¹ :=
|
||||
begin
|
||||
|
@ -331,12 +331,12 @@ assert (pceil (a / b))⁻¹ ≤ 1 / (1 / (b / a)),
|
|||
show 1 / (b / a) ≤ rat_of_pnat (pceil (a / b)),
|
||||
begin
|
||||
rewrite div_div_eq_mul_div,
|
||||
krewrite algebra.one_mul,
|
||||
rewrite algebra.one_mul,
|
||||
apply ubound_ge
|
||||
end
|
||||
end,
|
||||
begin
|
||||
krewrite one_div_one_div at this,
|
||||
rewrite one_div_one_div at this,
|
||||
exact this
|
||||
end
|
||||
|
||||
|
|
|
@ -47,7 +47,7 @@ have -num b > 0, from pos_of_mul_pos_right (pos_of_neg_neg this) (le_of_lt (deno
|
|||
neg_of_neg_pos this
|
||||
|
||||
theorem equiv_of_num_eq_zero {a b : prerat} (H1 : num a = 0) (H2 : num b = 0) : a ≡ b :=
|
||||
by krewrite [↑equiv, H1, H2, *zero_mul]
|
||||
by rewrite [↑equiv, H1, H2, *zero_mul]
|
||||
|
||||
theorem equiv.trans [trans] {a b c : prerat} (H1 : a ≡ b) (H2 : b ≡ c) : a ≡ c :=
|
||||
decidable.by_cases
|
||||
|
@ -115,10 +115,10 @@ definition inv : prerat → prerat
|
|||
| inv (prerat.mk -[1+n] d dp) := prerat.mk (-d) (nat.succ n) !of_nat_succ_pos
|
||||
|
||||
theorem equiv_zero_of_num_eq_zero {a : prerat} (H : num a = 0) : a ≡ zero :=
|
||||
by krewrite [↑equiv, H, ↑zero, ↑num, ↑of_int, *zero_mul]
|
||||
by rewrite [↑equiv, H, ↑zero, ↑num, ↑of_int, *zero_mul]
|
||||
|
||||
theorem num_eq_zero_of_equiv_zero {a : prerat} : a ≡ zero → num a = 0 :=
|
||||
by krewrite [↑equiv, ↑zero, ↑of_int, mul_one, zero_mul]; intro H; exact H
|
||||
by rewrite [↑equiv, ↑zero, ↑of_int, mul_one, zero_mul]; intro H; exact H
|
||||
|
||||
theorem inv_zero {d : int} (dp : d > 0) : inv (mk nat.zero d dp) = zero :=
|
||||
begin rewrite [↑inv, ▸*] end
|
||||
|
@ -219,10 +219,10 @@ by rewrite [↑add, ↑equiv, ▸*, *(mul.comm (num c)), *(λy, mul.comm y (deno
|
|||
*mul.right_distrib, *mul.assoc, *add.assoc]
|
||||
|
||||
theorem add_zero (a : prerat) : add a zero ≡ a :=
|
||||
by krewrite [↑add, ↑equiv, ↑zero, ↑of_int, ▸*, *mul_one, zero_mul, add_zero]
|
||||
by rewrite [↑add, ↑equiv, ↑zero, ↑of_int, ▸*, *mul_one, zero_mul, add_zero]
|
||||
|
||||
theorem add.left_inv (a : prerat) : add (neg a) a ≡ zero :=
|
||||
by krewrite [↑add, ↑equiv, ↑neg, ↑zero, ↑of_int, ▸*, -neg_mul_eq_neg_mul, add.left_inv, *zero_mul]
|
||||
by rewrite [↑add, ↑equiv, ↑neg, ↑zero, ↑of_int, ▸*, -neg_mul_eq_neg_mul, add.left_inv, *zero_mul]
|
||||
|
||||
theorem mul.comm (a b : prerat) : mul a b ≡ mul b a :=
|
||||
by rewrite [↑mul, ↑equiv, mul.comm (num a), mul.comm (denom a)]
|
||||
|
@ -273,7 +273,7 @@ theorem mul_inv_cancel : ∀{a : prerat}, ¬ a ≡ zero → mul a (inv a) ≡ on
|
|||
theorem zero_not_equiv_one : ¬ zero ≡ one :=
|
||||
begin
|
||||
esimp [equiv, zero, one, of_int],
|
||||
krewrite [zero_mul, int.mul_one],
|
||||
rewrite [zero_mul, int.mul_one],
|
||||
exact zero_ne_one
|
||||
end
|
||||
|
||||
|
@ -299,28 +299,29 @@ begin
|
|||
intros, apply rfl
|
||||
end
|
||||
|
||||
set_option pp.full_names true
|
||||
theorem reduce_equiv : ∀ a : prerat, reduce a ≡ a
|
||||
| (mk an ad adpos) :=
|
||||
decidable.by_cases
|
||||
(assume anz : an = 0,
|
||||
by krewrite [↑reduce, if_pos anz, ↑equiv, anz, *zero_mul])
|
||||
begin rewrite [↑reduce, if_pos anz, ↑equiv, anz], krewrite zero_mul end)
|
||||
(assume annz : an ≠ 0,
|
||||
by rewrite [↑reduce, if_neg annz, ↑equiv, int.mul.comm, -!mul_div_assoc !gcd_dvd_left,
|
||||
-!mul_div_assoc !gcd_dvd_right, int.mul.comm])
|
||||
-!mul_div_assoc !gcd_dvd_right, int.mul.comm])
|
||||
|
||||
theorem reduce_eq_reduce : ∀{a b : prerat}, a ≡ b → reduce a = reduce b
|
||||
| (mk an ad adpos) (mk bn bd bdpos) :=
|
||||
assume H : an * bd = bn * ad,
|
||||
decidable.by_cases
|
||||
(assume anz : an = 0,
|
||||
have H' : bn * ad = 0, by krewrite [-H, anz, zero_mul],
|
||||
have H' : bn * ad = 0, by rewrite [-H, anz, zero_mul],
|
||||
assert bnz : bn = 0,
|
||||
from or_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero H') (ne_of_gt adpos),
|
||||
by rewrite [↑reduce, if_pos anz, if_pos bnz])
|
||||
(assume annz : an ≠ 0,
|
||||
assert bnnz : bn ≠ 0, from
|
||||
assume bnz,
|
||||
have H' : an * bd = 0, by krewrite [H, bnz, zero_mul],
|
||||
have H' : an * bd = 0, by rewrite [H, bnz, zero_mul],
|
||||
have anz : an = 0,
|
||||
from or_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero H') (ne_of_gt bdpos),
|
||||
show false, from annz anz,
|
||||
|
@ -584,7 +585,7 @@ iff.mpr (!eq_div_iff_mul_eq H) (mul_denom a)
|
|||
theorem of_int_div {a b : ℤ} (H : b ∣ a) : of_int (a div b) = of_int a / of_int b :=
|
||||
decidable.by_cases
|
||||
(assume bz : b = 0,
|
||||
by krewrite [bz, div_zero, algebra.div_zero])
|
||||
by rewrite [bz, div_zero, of_int_zero, algebra.div_zero])
|
||||
(assume bnz : b ≠ 0,
|
||||
have bnz' : of_int b ≠ 0, from assume oibz, bnz (of_int.inj oibz),
|
||||
have H' : of_int (a div b) * of_int b = of_int a, from
|
||||
|
@ -601,7 +602,7 @@ theorem of_int_pow (a : ℤ) (n : ℕ) : of_int (a^n) = (of_int a)^n :=
|
|||
begin
|
||||
induction n with n ih,
|
||||
apply eq.refl,
|
||||
krewrite [pow_succ, pow_succ, of_int_mul, ih]
|
||||
rewrite [pow_succ, pow_succ, of_int_mul, ih]
|
||||
end
|
||||
|
||||
theorem of_nat_pow (a : ℕ) (n : ℕ) : of_nat (a^n) = (of_nat a)^n :=
|
||||
|
|
|
@ -209,7 +209,7 @@ theorem le.trans (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c :=
|
|||
assert H3 : nonneg (c - b + (b - a)), from nonneg_add H2 H1,
|
||||
begin
|
||||
revert H3,
|
||||
krewrite [rat.sub.def, add.assoc, neg_add_cancel_left],
|
||||
rewrite [rat.sub.def, add.assoc, sub_eq_add_neg, neg_add_cancel_left],
|
||||
intro H3, apply H3
|
||||
end
|
||||
|
||||
|
|
|
@ -114,10 +114,14 @@ inductive string : Type :=
|
|||
| empty : string
|
||||
| str : char → string → string
|
||||
|
||||
inductive nat :=
|
||||
| zero : nat
|
||||
| succ : nat → nat
|
||||
|
||||
inductive option (A : Type) : Type :=
|
||||
| none {} : option A
|
||||
| some : A → option A
|
||||
|
||||
-- Remark: we manually generate the nat.rec_on, nat.induction_on, nat.cases_on and nat.no_confusion.
|
||||
-- We do that because we want 0 instead of nat.zero in these eliminators.
|
||||
set_option inductive.rec_on false
|
||||
set_option inductive.cases_on false
|
||||
inductive nat :=
|
||||
| zero : nat
|
||||
| succ : nat → nat
|
||||
|
|
|
@ -10,6 +10,33 @@ open eq.ops decidable or
|
|||
notation `ℕ` := nat
|
||||
|
||||
namespace nat
|
||||
protected definition rec_on [reducible] [recursor] [unfold 2]
|
||||
{C : ℕ → Type} (n : ℕ) (H₁ : C 0) (H₂ : Π (a : ℕ), C a → C (succ a)) : C n :=
|
||||
nat.rec H₁ H₂ n
|
||||
|
||||
protected definition induction_on [recursor]
|
||||
{C : ℕ → Prop} (n : ℕ) (H₁ : C 0) (H₂ : Π (a : ℕ), C a → C (succ a)) : C n :=
|
||||
nat.rec H₁ H₂ n
|
||||
|
||||
protected definition cases_on [reducible] [recursor] [unfold 2]
|
||||
{C : ℕ → Type} (n : ℕ) (H₁ : C 0) (H₂ : Π (a : ℕ), C (succ a)) : C n :=
|
||||
nat.rec H₁ (λ a ih, H₂ a) n
|
||||
|
||||
protected definition no_confusion_type [reducible] (P : Type) (v₁ v₂ : ℕ) : Type :=
|
||||
nat.rec
|
||||
(nat.rec
|
||||
(P → P)
|
||||
(λ a₂ ih, P)
|
||||
v₂)
|
||||
(λ a₁ ih, nat.rec
|
||||
P
|
||||
(λ a₂ ih, (a₁ = a₂ → P) → P)
|
||||
v₂)
|
||||
v₁
|
||||
|
||||
protected definition no_confusion [reducible] [unfold 4]
|
||||
{P : Type} {v₁ v₂ : ℕ} (H : v₁ = v₂) : nat.no_confusion_type P v₁ v₂ :=
|
||||
eq.rec (λ H₁ : v₁ = v₁, nat.rec (λ h, h) (λ a ih h, h (eq.refl a)) v₁) H H
|
||||
|
||||
/- basic definitions on natural numbers -/
|
||||
inductive le (a : ℕ) : ℕ → Prop :=
|
||||
|
|
|
@ -95,7 +95,7 @@ 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)
|
||||
nat.rec a (λ b₁ r, succ r) b
|
||||
|
||||
definition nat_has_zero [reducible] [instance] : has_zero nat :=
|
||||
has_zero.mk nat.zero
|
||||
|
|
|
@ -59,7 +59,7 @@ begin
|
|||
induction n with [n, ih],
|
||||
{apply rfl},
|
||||
change choose (succ n) (succ 0) = succ n,
|
||||
krewrite [choose_succ_succ, ih, choose_zero_right]
|
||||
rewrite [choose_succ_succ, ih, choose_zero_right]
|
||||
end
|
||||
|
||||
theorem choose_pos {n : ℕ} : ∀ {k : ℕ}, k ≤ n → choose n k > 0 :=
|
||||
|
|
|
@ -290,7 +290,7 @@ lemma rotl_rotr : ∀ {n : nat} (m : nat), (@rotl n m) ∘ (rotr m) = id
|
|||
| (nat.succ n) := take m, funext take i, calc (mk_mod n (n*m)) + (-(mk_mod n (n*m)) + i) = i : add_neg_cancel_left
|
||||
|
||||
lemma rotl_succ {n : nat} : (rotl 1) ∘ (@succ n) = lift_succ :=
|
||||
funext (take i, eq_of_veq (begin krewrite [↑compose, ↑rotl, ↑madd, mul_one n, ↑mk_mod, mod_add_mod, ↑lift_succ, val_succ, -succ_add_eq_succ_add, add_mod_self_left, mod_eq_of_lt (lt.trans (is_lt i) !lt_succ_self), -val_lift] end))
|
||||
funext (take i, eq_of_veq (begin rewrite [↑compose, ↑rotl, ↑madd, mul_one n, ↑mk_mod, mod_add_mod, ↑lift_succ, val_succ, -succ_add_eq_succ_add, add_mod_self_left, mod_eq_of_lt (lt.trans (is_lt i) !lt_succ_self), -val_lift] end))
|
||||
|
||||
definition list.rotl {A : Type} : ∀ l : list A, list A
|
||||
| [] := []
|
||||
|
|
|
@ -41,7 +41,7 @@ section
|
|||
(H : a^n = c * b^n) : b = 1 :=
|
||||
have bpos : b > 0, from pos_of_ne_zero
|
||||
(suppose b = 0,
|
||||
have a^n = 0, by krewrite [H, this, zero_pow npos],
|
||||
have a^n = 0, by rewrite [H, this, zero_pow npos],
|
||||
assert a = 0, from eq_zero_of_pow_eq_zero this,
|
||||
show false, from ne_of_lt `0 < a` this⁻¹),
|
||||
have H₁ : ∀ p, prime p → ¬ p ∣ b, from
|
||||
|
@ -90,7 +90,7 @@ section
|
|||
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,
|
||||
have qeq0 : q = 0, by rewrite [eq_num_div_denom, aeq0, of_int_zero, algebra.zero_div],
|
||||
absurd qeq0 qne0,
|
||||
have nat_abs a ≠ 0, from
|
||||
suppose nat_abs a = 0,
|
||||
|
@ -104,7 +104,7 @@ section
|
|||
c = (num q)^n :=
|
||||
have denom q = 1, from denom_eq_one_of_pow_eq npos H,
|
||||
have of_int c = of_int ((num q)^n), using this,
|
||||
by krewrite [-H, eq_num_div_denom q at {1}, this, div_one, of_int_pow],
|
||||
by rewrite [-H, eq_num_div_denom q at {1}, this, of_int_one, div_one, of_int_pow],
|
||||
show c = (num q)^n , from of_int.inj this
|
||||
end
|
||||
|
||||
|
@ -122,9 +122,9 @@ section
|
|||
by let H := (pos_of_prime primep); rewrite this at H; exfalso; exact !lt.irrefl H),
|
||||
assert agtz : a > 0, from pos_of_ne_zero
|
||||
(suppose a = 0,
|
||||
show false, using npos pnez, by revert peq; krewrite [this, zero_pow npos]; exact pnez),
|
||||
show false, using npos pnez, by revert peq; rewrite [this, zero_pow npos]; exact pnez),
|
||||
have n * mult p a = 1, from calc
|
||||
n * mult p a = mult p (a^n) : begin krewrite [mult_pow n agtz primep] end
|
||||
n * mult p a = mult p (a^n) : begin rewrite [mult_pow n agtz primep] end
|
||||
... = mult p p : peq
|
||||
... = 1 : mult_self (gt_one_of_prime primep),
|
||||
have n ∣ 1, from dvd_of_mul_right_eq this,
|
||||
|
@ -162,7 +162,7 @@ section
|
|||
have a * a = c * b * b, by rewrite -mul.assoc at H; apply H,
|
||||
have a div (gcd a b) = c * b div gcd (c * b) a, from div_gcd_eq_div_gcd this bpos apos,
|
||||
have a = c * b div gcd c a,
|
||||
using this, by revert this; krewrite [↑coprime at co, co, int.div_one, H₁]; intros; assumption,
|
||||
using this, by revert this; rewrite [↑coprime at co, co, int.div_one, H₁]; intros; assumption,
|
||||
have a = b * (c div gcd c a),
|
||||
using this,
|
||||
by revert this; rewrite [mul.comm, !mul_div_assoc !gcd_dvd_left]; intros; assumption,
|
||||
|
|
|
@ -120,7 +120,7 @@ private theorem mult_pow_mul {p n : ℕ} (i : ℕ) (pgt1 : p > 1) (npos : n > 0)
|
|||
mult p (p^i * n) = i + mult p n :=
|
||||
begin
|
||||
induction i with [i, ih],
|
||||
{krewrite [pow_zero, one_mul, zero_add]},
|
||||
{rewrite [pow_zero, one_mul, zero_add]},
|
||||
have p > 0, from lt.trans zero_lt_one pgt1,
|
||||
have psin_pos : p^(succ i) * n > 0, from mul_pos (!pow_pos_of_pos this) npos,
|
||||
have p ∣ p^(succ i) * n, by rewrite [pow_succ, mul.assoc]; apply dvd_mul_right,
|
||||
|
@ -129,7 +129,7 @@ begin
|
|||
end
|
||||
|
||||
theorem mult_pow_self {p : ℕ} (i : ℕ) (pgt1 : p > 1) : mult p (p^i) = i :=
|
||||
by krewrite [-(mul_one (p^i)), mult_pow_mul i pgt1 zero_lt_one, mult_one_right]
|
||||
by rewrite [-(mul_one (p^i)), mult_pow_mul i pgt1 zero_lt_one, mult_one_right]
|
||||
|
||||
theorem mult_self {p : ℕ} (pgt1 : p > 1) : mult p p = 1 :=
|
||||
by rewrite [-pow_one p at {2}]; apply mult_pow_self 1 pgt1
|
||||
|
@ -175,7 +175,7 @@ calc
|
|||
theorem mult_pow {p m : ℕ} (n : ℕ) (mpos : m > 0) (primep : prime p) : mult p (m^n) = n * mult p m :=
|
||||
begin
|
||||
induction n with n ih,
|
||||
krewrite [pow_zero, mult_one_right, zero_mul],
|
||||
rewrite [pow_zero, mult_one_right, zero_mul],
|
||||
rewrite [pow_succ, mult_mul primep mpos (!pow_pos_of_pos mpos), ih, succ_mul, add.comm]
|
||||
end
|
||||
|
||||
|
@ -247,7 +247,7 @@ theorem mult_pow_eq_zero_of_prime_of_ne {p q : ℕ} (primep : prime p) (primeq :
|
|||
(pneq : p ≠ q) (i : ℕ) : mult p (q^i) = 0 :=
|
||||
begin
|
||||
induction i with i ih,
|
||||
{krewrite [pow_zero, mult_one_right]},
|
||||
{rewrite [pow_zero, mult_one_right]},
|
||||
have qpos : q > 0, from pos_of_prime primeq,
|
||||
have qipos : q^i > 0, from !pow_pos_of_pos qpos,
|
||||
rewrite [pow_succ', mult_mul primep qipos qpos, ih, mult_eq_zero_of_prime_of_ne primep primeq pneq]
|
||||
|
|
Loading…
Reference in a new issue