2014-12-05 22:03:24 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Floris van Doorn, Jeremy Avigad
|
2014-08-23 00:56:25 +00:00
|
|
|
|
|
2014-12-17 18:32:38 +00:00
|
|
|
|
The integers, with addition, multiplication, and subtraction. The representation of the integers is
|
2014-12-26 21:25:05 +00:00
|
|
|
|
chosen to compute efficiently.
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
To faciliate proving things about these operations, we show that the integers are a quotient of
|
2014-12-26 21:25:05 +00:00
|
|
|
|
ℕ × ℕ with the usual equivalence relation, ≡, and functions
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
abstr : ℕ × ℕ → ℤ
|
|
|
|
|
repr : ℤ → ℕ × ℕ
|
|
|
|
|
|
2014-12-26 21:25:05 +00:00
|
|
|
|
satisfying:
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
abstr_repr (a : ℤ) : abstr (repr a) = a
|
|
|
|
|
repr_abstr (p : ℕ × ℕ) : repr (abstr p) ≡ p
|
|
|
|
|
abstr_eq (p q : ℕ × ℕ) : p ≡ q → abstr p = abstr q
|
|
|
|
|
|
|
|
|
|
For example, to "lift" statements about add to statements about padd, we need to prove the
|
|
|
|
|
following:
|
|
|
|
|
|
|
|
|
|
repr_add (a b : ℤ) : repr (a + b) = padd (repr a) (repr b)
|
|
|
|
|
padd_congr (p p' q q' : ℕ × ℕ) (H1 : p ≡ p') (H2 : q ≡ q') : padd p q ≡ p' q'
|
2014-12-26 21:25:05 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
-/
|
2014-12-17 18:32:38 +00:00
|
|
|
|
import data.nat.basic data.nat.order data.nat.sub data.prod
|
|
|
|
|
import algebra.relation algebra.binary algebra.ordered_ring
|
2014-10-02 00:51:17 +00:00
|
|
|
|
open eq.ops
|
2014-12-17 18:32:38 +00:00
|
|
|
|
open prod relation nat
|
2015-07-14 13:54:53 +00:00
|
|
|
|
open decidable binary
|
2015-10-12 03:29:31 +00:00
|
|
|
|
open algebra
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- the type of integers -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
inductive int : Type :=
|
2015-02-26 01:00:10 +00:00
|
|
|
|
| of_nat : nat → int
|
|
|
|
|
| neg_succ_of_nat : nat → int
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
notation `ℤ` := int
|
2015-05-14 02:01:48 +00:00
|
|
|
|
definition int.of_num [coercion] [reducible] [constructor] (n : num) : ℤ :=
|
|
|
|
|
int.of_nat (nat.of_num n)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
namespace int
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-06-29 11:00:42 +00:00
|
|
|
|
attribute int.of_nat [coercion]
|
|
|
|
|
|
2015-09-30 15:06:31 +00:00
|
|
|
|
notation `-[1+ ` n `]` := int.neg_succ_of_nat n -- for pretty-printing output
|
2015-07-24 15:56:18 +00:00
|
|
|
|
|
2015-10-22 22:11:42 +00:00
|
|
|
|
protected definition prio : num := num.pred nat.prio
|
|
|
|
|
|
|
|
|
|
definition int_has_zero [reducible] [instance] [priority int.prio] : has_zero int :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
has_zero.mk (of_nat 0)
|
|
|
|
|
|
2015-10-22 22:11:42 +00:00
|
|
|
|
definition int_has_one [reducible] [instance] [priority int.prio] : has_one int :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
has_one.mk (of_nat 1)
|
|
|
|
|
|
2015-10-13 00:18:10 +00:00
|
|
|
|
theorem of_nat_zero : of_nat (0:nat) = (0:int) :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
rfl
|
|
|
|
|
|
2015-10-13 00:18:10 +00:00
|
|
|
|
theorem of_nat_one : of_nat (1:nat) = (1:int) :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
rfl
|
|
|
|
|
|
2014-12-22 19:21:22 +00:00
|
|
|
|
/- definitions of basic functions -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
definition neg_of_nat : ℕ → ℤ
|
2015-10-12 03:29:31 +00:00
|
|
|
|
| 0 := 0
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (succ m) := -[1+ m]
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
definition sub_nat_nat (m n : ℕ) : ℤ :=
|
2015-10-07 23:44:47 +00:00
|
|
|
|
match (n - m : nat) with
|
2015-10-12 03:29:31 +00:00
|
|
|
|
| 0 := of_nat (m - n) -- m ≥ n
|
|
|
|
|
| (succ k) := -[1+ k] -- m < n, and n - m = succ k
|
2015-07-24 15:56:18 +00:00
|
|
|
|
end
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-07 23:44:47 +00:00
|
|
|
|
protected definition neg (a : ℤ) : ℤ :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
int.cases_on a neg_of_nat succ
|
|
|
|
|
|
2015-10-07 23:44:47 +00:00
|
|
|
|
protected definition add : ℤ → ℤ → ℤ
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (of_nat m) (of_nat n) := m + n
|
|
|
|
|
| (of_nat m) -[1+ n] := sub_nat_nat m (succ n)
|
|
|
|
|
| -[1+ m] (of_nat n) := sub_nat_nat n (succ m)
|
|
|
|
|
| -[1+ m] -[1+ n] := neg_of_nat (succ m + succ n)
|
|
|
|
|
|
2015-10-07 23:44:47 +00:00
|
|
|
|
protected definition mul : ℤ → ℤ → ℤ
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (of_nat m) (of_nat n) := m * n
|
|
|
|
|
| (of_nat m) -[1+ n] := neg_of_nat (m * succ n)
|
|
|
|
|
| -[1+ m] (of_nat n) := neg_of_nat (succ m * n)
|
|
|
|
|
| -[1+ m] -[1+ n] := succ m * succ n
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- notation -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-07 23:44:47 +00:00
|
|
|
|
definition int_has_add [reducible] [instance] [priority int.prio] : has_add int := has_add.mk int.add
|
|
|
|
|
definition int_has_neg [reducible] [instance] [priority int.prio] : has_neg int := has_neg.mk int.neg
|
|
|
|
|
definition int_has_mul [reducible] [instance] [priority int.prio] : has_mul int := has_mul.mk int.mul
|
2014-08-22 00:54:50 +00:00
|
|
|
|
|
2015-10-22 20:45:20 +00:00
|
|
|
|
lemma mul_of_nat_of_nat (m n : nat) : of_nat m * of_nat n = of_nat (m * n) :=
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma mul_of_nat_neg_succ_of_nat (m n : nat) : of_nat m * -[1+ n] = neg_of_nat (m * succ n) :=
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma mul_neg_succ_of_nat_of_nat (m n : nat) : -[1+ m] * of_nat n = neg_of_nat (succ m * n) :=
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma mul_neg_succ_of_nat_neg_succ_of_nat (m n : nat) : -[1+ m] * -[1+ n] = succ m * succ n :=
|
|
|
|
|
rfl
|
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- some basic functions and properties -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-02-25 19:14:57 +00:00
|
|
|
|
theorem of_nat.inj {m n : ℕ} (H : of_nat m = of_nat n) : m = n :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
int.no_confusion H imp.id
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-09-12 14:00:34 +00:00
|
|
|
|
theorem eq_of_of_nat_eq_of_nat {m n : ℕ} (H : of_nat m = of_nat n) : m = n :=
|
|
|
|
|
of_nat.inj H
|
|
|
|
|
|
|
|
|
|
theorem of_nat_eq_of_nat_iff (m n : ℕ) : of_nat m = of_nat n ↔ m = n :=
|
2015-05-30 12:09:34 +00:00
|
|
|
|
iff.intro of_nat.inj !congr_arg
|
|
|
|
|
|
2015-02-25 19:14:57 +00:00
|
|
|
|
theorem neg_succ_of_nat.inj {m n : ℕ} (H : neg_succ_of_nat m = neg_succ_of_nat n) : m = n :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
int.no_confusion H imp.id
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-06-29 05:16:58 +00:00
|
|
|
|
theorem neg_succ_of_nat_eq (n : ℕ) : -[1+ n] = -(n + 1) := rfl
|
2015-02-25 19:14:57 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
private definition has_decidable_eq₂ : Π (a b : ℤ), decidable (a = b)
|
|
|
|
|
| (of_nat m) (of_nat n) := decidable_of_decidable_of_iff
|
2015-09-12 14:00:34 +00:00
|
|
|
|
(nat.has_decidable_eq m n) (iff.symm (of_nat_eq_of_nat_iff m n))
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (of_nat m) -[1+ n] := inr (by contradiction)
|
|
|
|
|
| -[1+ m] (of_nat n) := inr (by contradiction)
|
|
|
|
|
| -[1+ m] -[1+ n] := if H : m = n then
|
|
|
|
|
inl (congr_arg neg_succ_of_nat H) else inr (not.mto neg_succ_of_nat.inj H)
|
|
|
|
|
|
2015-10-09 21:50:17 +00:00
|
|
|
|
definition has_decidable_eq [instance] [priority int.prio] : decidable_eq ℤ := has_decidable_eq₂
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-05-25 10:00:41 +00:00
|
|
|
|
theorem of_nat_add (n m : nat) : of_nat (n + m) = of_nat n + of_nat m := rfl
|
2014-12-22 19:21:22 +00:00
|
|
|
|
|
|
|
|
|
theorem of_nat_succ (n : ℕ) : of_nat (succ n) = of_nat n + 1 := rfl
|
|
|
|
|
|
2015-05-25 10:00:41 +00:00
|
|
|
|
theorem of_nat_mul (n m : ℕ) : of_nat (n * m) = of_nat n * of_nat m := rfl
|
2014-12-22 19:21:22 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem sub_nat_nat_of_ge {m n : ℕ} (H : m ≥ n) : sub_nat_nat m n = of_nat (m - n) :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
show sub_nat_nat m n = nat.cases_on 0 (m -[nat] n) _, from (sub_eq_zero_of_le H) ▸ rfl
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
|
section
|
|
|
|
|
local attribute sub_nat_nat [reducible]
|
2015-10-12 03:29:31 +00:00
|
|
|
|
theorem sub_nat_nat_of_lt {m n : ℕ} (H : m < n) : sub_nat_nat m n = -[1+ pred (n - m)] :=
|
2015-10-23 14:06:20 +00:00
|
|
|
|
have H1 : n - m = succ (pred (n - m)), from eq.symm (succ_pred_of_pos (nat.sub_pos_of_lt H)),
|
2015-10-12 03:29:31 +00:00
|
|
|
|
show sub_nat_nat m n = nat.cases_on (succ (nat.pred (n - m))) (m -[nat] n) _, from H1 ▸ rfl
|
2015-01-09 02:47:44 +00:00
|
|
|
|
end
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
definition nat_abs (a : ℤ) : ℕ := int.cases_on a function.id succ
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem nat_abs_of_nat (n : ℕ) : nat_abs n = n := rfl
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-08-17 03:23:03 +00:00
|
|
|
|
theorem eq_zero_of_nat_abs_eq_zero : Π {a : ℤ}, nat_abs a = 0 → a = 0
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (of_nat m) H := congr_arg of_nat H
|
|
|
|
|
| -[1+ m'] H := absurd H !succ_ne_zero
|
2014-12-17 18:32:38 +00:00
|
|
|
|
|
2015-10-13 20:16:19 +00:00
|
|
|
|
theorem nat_abs_zero : nat_abs (0:int) = (0:nat) :=
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
theorem nat_abs_one : nat_abs (1:int) = (1:nat) :=
|
|
|
|
|
rfl
|
|
|
|
|
|
2014-12-26 21:25:05 +00:00
|
|
|
|
/- int is a quotient of ordered pairs of natural numbers -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-04-16 02:25:02 +00:00
|
|
|
|
protected definition equiv (p q : ℕ × ℕ) : Prop := pr1 p + pr2 q = pr2 p + pr1 q
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-09-30 15:06:31 +00:00
|
|
|
|
local infix ≡ := int.equiv
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-05-02 22:15:35 +00:00
|
|
|
|
protected theorem equiv.refl [refl] {p : ℕ × ℕ} : p ≡ p := !add.comm
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-05-02 22:15:35 +00:00
|
|
|
|
protected theorem equiv.symm [symm] {p q : ℕ × ℕ} (H : p ≡ q) : q ≡ p :=
|
2014-08-20 02:32:44 +00:00
|
|
|
|
calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
pr1 q + pr2 p = pr2 p + pr1 q : by rewrite add.comm
|
2014-12-05 22:03:24 +00:00
|
|
|
|
... = pr1 p + pr2 q : H⁻¹
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = pr2 q + pr1 p : by rewrite add.comm
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-05-02 22:15:35 +00:00
|
|
|
|
protected theorem equiv.trans [trans] {p q r : ℕ × ℕ} (H1 : p ≡ q) (H2 : q ≡ r) : p ≡ r :=
|
2015-10-22 20:09:26 +00:00
|
|
|
|
add.right_cancel (calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
pr1 p + pr2 r + pr2 q = pr1 p + pr2 q + pr2 r : by rewrite add.right_comm
|
2014-12-05 22:03:24 +00:00
|
|
|
|
... = pr2 p + pr1 q + pr2 r : {H1}
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = pr2 p + (pr1 q + pr2 r) : by rewrite add.assoc
|
2014-12-05 22:03:24 +00:00
|
|
|
|
... = pr2 p + (pr2 q + pr1 r) : {H2}
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = pr2 p + pr2 q + pr1 r : by rewrite add.assoc
|
|
|
|
|
... = pr2 p + pr1 r + pr2 q : by rewrite add.right_comm)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-05-19 05:35:18 +00:00
|
|
|
|
protected theorem equiv_equiv : is_equivalence int.equiv :=
|
2014-12-17 18:32:38 +00:00
|
|
|
|
is_equivalence.mk @equiv.refl @equiv.symm @equiv.trans
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
protected theorem equiv_cases {p q : ℕ × ℕ} (H : p ≡ q) :
|
2014-12-05 22:03:24 +00:00
|
|
|
|
(pr1 p ≥ pr2 p ∧ pr1 q ≥ pr2 q) ∨ (pr1 p < pr2 p ∧ pr1 q < pr2 q) :=
|
2015-10-07 23:44:47 +00:00
|
|
|
|
or.elim (@le_or_gt _ _ (pr2 p) (pr1 p))
|
2015-07-21 16:10:56 +00:00
|
|
|
|
(suppose pr1 p ≥ pr2 p,
|
|
|
|
|
have pr2 p + pr1 q ≥ pr2 p + pr2 q, from H ▸ add_le_add_right this (pr2 q),
|
|
|
|
|
or.inl (and.intro `pr1 p ≥ pr2 p` (le_of_add_le_add_left this)))
|
2015-10-07 23:44:47 +00:00
|
|
|
|
(suppose H₁ : pr1 p < pr2 p,
|
|
|
|
|
have pr2 p + pr1 q < pr2 p + pr2 q, from H ▸ add_lt_add_right H₁ (pr2 q),
|
|
|
|
|
or.inr (and.intro H₁ (lt_of_add_lt_add_left this)))
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-04-16 02:25:02 +00:00
|
|
|
|
protected theorem equiv_of_eq {p q : ℕ × ℕ} (H : p = q) : p ≡ q := H ▸ equiv.refl
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- the representation and abstraction functions -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
definition abstr (a : ℕ × ℕ) : ℤ := sub_nat_nat (pr1 a) (pr2 a)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem abstr_of_ge {p : ℕ × ℕ} (H : pr1 p ≥ pr2 p) : abstr p = of_nat (pr1 p - pr2 p) :=
|
|
|
|
|
sub_nat_nat_of_ge H
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem abstr_of_lt {p : ℕ × ℕ} (H : pr1 p < pr2 p) :
|
2015-07-24 15:56:18 +00:00
|
|
|
|
abstr p = -[1+ pred (pr2 p - pr1 p)] :=
|
2014-12-05 22:03:24 +00:00
|
|
|
|
sub_nat_nat_of_lt H
|
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
definition repr : ℤ → ℕ × ℕ
|
|
|
|
|
| (of_nat m) := (m, 0)
|
|
|
|
|
| -[1+ m] := (0, succ m)
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem abstr_repr : Π (a : ℤ), abstr (repr a) = a
|
|
|
|
|
| (of_nat m) := (sub_nat_nat_of_ge (zero_le m))
|
|
|
|
|
| -[1+ m] := rfl
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem repr_sub_nat_nat (m n : ℕ) : repr (sub_nat_nat m n) ≡ (m, n) :=
|
2015-10-22 20:09:26 +00:00
|
|
|
|
nat.lt_ge_by_cases
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(take H : m < n,
|
|
|
|
|
have H1 : repr (sub_nat_nat m n) = (0, n - m), by
|
2015-10-23 14:06:20 +00:00
|
|
|
|
rewrite [sub_nat_nat_of_lt H, -(succ_pred_of_pos (nat.sub_pos_of_lt H))],
|
|
|
|
|
H1⁻¹ ▸ (!zero_add ⬝ (nat.sub_add_cancel (le_of_lt H))⁻¹))
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(take H : m ≥ n,
|
|
|
|
|
have H1 : repr (sub_nat_nat m n) = (m - n, 0), from sub_nat_nat_of_ge H ▸ rfl,
|
2015-10-23 14:06:20 +00:00
|
|
|
|
H1⁻¹ ▸ ((nat.sub_add_cancel H) ⬝ !zero_add⁻¹))
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem repr_abstr (p : ℕ × ℕ) : repr (abstr p) ≡ p :=
|
|
|
|
|
!prod.eta ▸ !repr_sub_nat_nat
|
|
|
|
|
|
|
|
|
|
theorem abstr_eq {p q : ℕ × ℕ} (Hequiv : p ≡ q) : abstr p = abstr q :=
|
2015-05-19 05:35:18 +00:00
|
|
|
|
or.elim (int.equiv_cases Hequiv)
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(and.rec (assume (Hp : pr1 p ≥ pr2 p) (Hq : pr1 q ≥ pr2 q),
|
|
|
|
|
have H : pr1 p - pr2 p = pr1 q - pr2 q, from
|
|
|
|
|
calc pr1 p - pr2 p
|
2015-10-23 14:06:20 +00:00
|
|
|
|
= pr1 p + pr2 q - pr2 q - pr2 p : by rewrite nat.add_sub_cancel
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = pr2 p + pr1 q - pr2 q - pr2 p : Hequiv
|
2015-10-23 14:06:20 +00:00
|
|
|
|
... = pr2 p + (pr1 q - pr2 q) - pr2 p : nat.add_sub_assoc Hq
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = pr1 q - pr2 q + pr2 p - pr2 p : by rewrite add.comm
|
2015-10-23 14:06:20 +00:00
|
|
|
|
... = pr1 q - pr2 q : by rewrite nat.add_sub_cancel,
|
2015-07-24 15:56:18 +00:00
|
|
|
|
abstr_of_ge Hp ⬝ (H ▸ rfl) ⬝ (abstr_of_ge Hq)⁻¹))
|
|
|
|
|
(and.rec (assume (Hp : pr1 p < pr2 p) (Hq : pr1 q < pr2 q),
|
|
|
|
|
have H : pr2 p - pr1 p = pr2 q - pr1 q, from
|
|
|
|
|
calc pr2 p - pr1 p
|
2015-10-23 14:06:20 +00:00
|
|
|
|
= pr2 p + pr1 q - pr1 q - pr1 p : by rewrite nat.add_sub_cancel
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = pr1 p + pr2 q - pr1 q - pr1 p : Hequiv
|
2015-10-23 14:06:20 +00:00
|
|
|
|
... = pr1 p + (pr2 q - pr1 q) - pr1 p : nat.add_sub_assoc (le_of_lt Hq)
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = pr2 q - pr1 q + pr1 p - pr1 p : by rewrite add.comm
|
2015-10-23 14:06:20 +00:00
|
|
|
|
... = pr2 q - pr1 q : by rewrite nat.add_sub_cancel,
|
2015-07-24 15:56:18 +00:00
|
|
|
|
abstr_of_lt Hp ⬝ (H ▸ rfl) ⬝ (abstr_of_lt Hq)⁻¹))
|
|
|
|
|
|
|
|
|
|
theorem equiv_iff (p q : ℕ × ℕ) : (p ≡ q) ↔ (abstr p = abstr q) :=
|
|
|
|
|
iff.intro abstr_eq (assume H, equiv.trans (H ▸ equiv.symm (repr_abstr p)) (repr_abstr q))
|
|
|
|
|
|
|
|
|
|
theorem equiv_iff3 (p q : ℕ × ℕ) : (p ≡ q) ↔ ((p ≡ p) ∧ (q ≡ q) ∧ (abstr p = abstr q)) :=
|
|
|
|
|
iff.trans !equiv_iff (iff.symm
|
|
|
|
|
(iff.trans (and_iff_right !equiv.refl) (and_iff_right !equiv.refl)))
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem eq_abstr_of_equiv_repr {a : ℤ} {p : ℕ × ℕ} (Hequiv : repr a ≡ p) : a = abstr p :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
!abstr_repr⁻¹ ⬝ abstr_eq Hequiv
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem eq_of_repr_equiv_repr {a b : ℤ} (H : repr a ≡ repr b) : a = b :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
eq_abstr_of_equiv_repr H ⬝ !abstr_repr
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
|
section
|
|
|
|
|
local attribute abstr [reducible]
|
|
|
|
|
local attribute dist [reducible]
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem nat_abs_abstr : Π (p : ℕ × ℕ), nat_abs (abstr p) = dist (pr1 p) (pr2 p)
|
2015-10-22 20:09:26 +00:00
|
|
|
|
| (m, n) := nat.lt_ge_by_cases
|
2014-12-05 22:03:24 +00:00
|
|
|
|
(assume H : m < n,
|
|
|
|
|
calc
|
2015-07-24 15:56:18 +00:00
|
|
|
|
nat_abs (abstr (m, n)) = nat_abs (-[1+ pred (n - m)]) : int.abstr_of_lt H
|
2015-10-23 14:06:20 +00:00
|
|
|
|
... = n - m : succ_pred_of_pos (nat.sub_pos_of_lt H)
|
2015-01-12 21:28:42 +00:00
|
|
|
|
... = dist m n : dist_eq_sub_of_le (le_of_lt H))
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(assume H : m ≥ n, (abstr_of_ge H)⁻¹ ▸ (dist_eq_sub_of_ge H)⁻¹)
|
2015-01-09 02:47:44 +00:00
|
|
|
|
end
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2014-12-22 19:21:22 +00:00
|
|
|
|
theorem cases_of_nat_succ (a : ℤ) : (∃n : ℕ, a = of_nat n) ∨ (∃n : ℕ, a = - (of_nat (succ n))) :=
|
|
|
|
|
int.cases_on a (take m, or.inl (exists.intro _ rfl)) (take m, or.inr (exists.intro _ rfl))
|
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem cases_of_nat (a : ℤ) : (∃n : ℕ, a = of_nat n) ∨ (∃n : ℕ, a = - of_nat n) :=
|
|
|
|
|
or.imp_right (Exists.rec (take n, (exists.intro _))) !cases_of_nat_succ
|
|
|
|
|
|
2014-12-22 19:21:22 +00:00
|
|
|
|
theorem by_cases_of_nat {P : ℤ → Prop} (a : ℤ)
|
|
|
|
|
(H1 : ∀n : ℕ, P (of_nat n)) (H2 : ∀n : ℕ, P (- of_nat n)) :
|
2014-08-22 00:54:50 +00:00
|
|
|
|
P a :=
|
2014-12-22 19:21:22 +00:00
|
|
|
|
or.elim (cases_of_nat a)
|
2014-09-05 01:41:06 +00:00
|
|
|
|
(assume H, obtain (n : ℕ) (H3 : a = n), from H, H3⁻¹ ▸ H1 n)
|
|
|
|
|
(assume H, obtain (n : ℕ) (H3 : a = -n), from H, H3⁻¹ ▸ H2 n)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-22 19:21:22 +00:00
|
|
|
|
theorem by_cases_of_nat_succ {P : ℤ → Prop} (a : ℤ)
|
|
|
|
|
(H1 : ∀n : ℕ, P (of_nat n)) (H2 : ∀n : ℕ, P (- of_nat (succ n))) :
|
|
|
|
|
P a :=
|
|
|
|
|
or.elim (cases_of_nat_succ a)
|
|
|
|
|
(assume H, obtain (n : ℕ) (H3 : a = n), from H, H3⁻¹ ▸ H1 n)
|
|
|
|
|
(assume H, obtain (n : ℕ) (H3 : a = -(succ n)), from H, H3⁻¹ ▸ H2 n)
|
|
|
|
|
|
2014-12-17 18:32:38 +00:00
|
|
|
|
/-
|
2014-12-26 21:25:05 +00:00
|
|
|
|
int is a ring
|
2014-12-17 18:32:38 +00:00
|
|
|
|
-/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- addition -/
|
|
|
|
|
|
2015-02-01 16:52:13 +00:00
|
|
|
|
definition padd (p q : ℕ × ℕ) : ℕ × ℕ := (pr1 p + pr1 q, pr2 p + pr2 q)
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem repr_add : Π (a b : ℤ), repr (add a b) ≡ padd (repr a) (repr b)
|
|
|
|
|
| (of_nat m) (of_nat n) := !equiv.refl
|
2015-10-11 19:00:57 +00:00
|
|
|
|
| (of_nat m) -[1+ n] :=
|
|
|
|
|
begin
|
|
|
|
|
change repr (sub_nat_nat m (succ n)) ≡ (m + 0, 0 + succ n),
|
|
|
|
|
rewrite [zero_add, add_zero],
|
|
|
|
|
apply repr_sub_nat_nat
|
|
|
|
|
end
|
|
|
|
|
| -[1+ m] (of_nat n) :=
|
|
|
|
|
begin
|
|
|
|
|
change repr (-[1+ m] + n) ≡ (0 + n, succ m + 0),
|
|
|
|
|
rewrite [zero_add, add_zero],
|
|
|
|
|
apply repr_sub_nat_nat
|
|
|
|
|
end
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| -[1+ m] -[1+ n] := !repr_sub_nat_nat
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem padd_congr {p p' q q' : ℕ × ℕ} (Ha : p ≡ p') (Hb : q ≡ q') : padd p q ≡ padd p' q' :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
calc pr1 p + pr1 q + (pr2 p' + pr2 q')
|
|
|
|
|
= pr1 p + pr2 p' + (pr1 q + pr2 q') : add.comm4
|
2014-12-05 22:03:24 +00:00
|
|
|
|
... = pr2 p + pr1 p' + (pr1 q + pr2 q') : {Ha}
|
|
|
|
|
... = pr2 p + pr1 p' + (pr2 q + pr1 q') : {Hb}
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = pr2 p + pr2 q + (pr1 p' + pr1 q') : add.comm4
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem padd_comm (p q : ℕ × ℕ) : padd p q = padd q p :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
calc (pr1 p + pr1 q, pr2 p + pr2 q)
|
2015-10-07 23:44:47 +00:00
|
|
|
|
= (pr1 q + pr1 p, pr2 p + pr2 q) : by rewrite add.comm
|
|
|
|
|
... = (pr1 q + pr1 p, pr2 q + pr2 p) : by rewrite (add.comm (pr2 p) (pr2 q))
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem padd_assoc (p q r : ℕ × ℕ) : padd (padd p q) r = padd p (padd q r) :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
calc (pr1 p + pr1 q + pr1 r, pr2 p + pr2 q + pr2 r)
|
2015-10-07 23:44:47 +00:00
|
|
|
|
= (pr1 p + (pr1 q + pr1 r), pr2 p + pr2 q + pr2 r) : by rewrite add.assoc
|
|
|
|
|
... = (pr1 p + (pr1 q + pr1 r), pr2 p + (pr2 q + pr2 r)) : by rewrite add.assoc
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem add_comm (a b : ℤ) : a + b = b + a :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
eq_of_repr_equiv_repr (equiv.trans !repr_add
|
|
|
|
|
(equiv.symm (!padd_comm ▸ !repr_add)))
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem add_assoc (a b c : ℤ) : a + b + c = a + (b + c) :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
eq_of_repr_equiv_repr (calc
|
|
|
|
|
repr (a + b + c)
|
|
|
|
|
≡ padd (repr (a + b)) (repr c) : repr_add
|
|
|
|
|
... ≡ padd (padd (repr a) (repr b)) (repr c) : padd_congr !repr_add !equiv.refl
|
|
|
|
|
... = padd (repr a) (padd (repr b) (repr c)) : !padd_assoc
|
|
|
|
|
... ≡ padd (repr a) (repr (b + c)) : padd_congr !equiv.refl !repr_add
|
|
|
|
|
... ≡ repr (a + (b + c)) : repr_add)
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem add_zero : Π (a : ℤ), a + 0 = a := int.rec (λm, rfl) (λm, rfl)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem zero_add (a : ℤ) : 0 + a = a := !int.add_comm ▸ !int.add_zero
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-17 18:32:38 +00:00
|
|
|
|
/- negation -/
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-17 18:32:38 +00:00
|
|
|
|
definition pneg (p : ℕ × ℕ) : ℕ × ℕ := (pr2 p, pr1 p)
|
|
|
|
|
|
|
|
|
|
-- note: this is =, not just ≡
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem repr_neg : Π (a : ℤ), repr (- a) = pneg (repr a)
|
|
|
|
|
| 0 := rfl
|
|
|
|
|
| (succ m) := rfl
|
|
|
|
|
| -[1+ m] := rfl
|
2014-12-17 18:32:38 +00:00
|
|
|
|
|
|
|
|
|
theorem pneg_congr {p p' : ℕ × ℕ} (H : p ≡ p') : pneg p ≡ pneg p' := eq.symm H
|
|
|
|
|
|
|
|
|
|
theorem pneg_pneg (p : ℕ × ℕ) : pneg (pneg p) = p := !prod.eta
|
|
|
|
|
|
|
|
|
|
theorem nat_abs_neg (a : ℤ) : nat_abs (-a) = nat_abs a :=
|
|
|
|
|
calc
|
|
|
|
|
nat_abs (-a) = nat_abs (abstr (repr (-a))) : abstr_repr
|
|
|
|
|
... = nat_abs (abstr (pneg (repr a))) : repr_neg
|
|
|
|
|
... = dist (pr1 (pneg (repr a))) (pr2 (pneg (repr a))) : nat_abs_abstr
|
2015-01-12 21:28:42 +00:00
|
|
|
|
... = dist (pr2 (pneg (repr a))) (pr1 (pneg (repr a))) : dist.comm
|
2014-12-17 18:32:38 +00:00
|
|
|
|
... = nat_abs (abstr (repr a)) : nat_abs_abstr
|
|
|
|
|
... = nat_abs a : abstr_repr
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem padd_pneg (p : ℕ × ℕ) : padd p (pneg p) ≡ (0, 0) :=
|
2015-10-14 19:27:09 +00:00
|
|
|
|
show pr1 p + pr2 p + 0 = pr2 p + pr1 p + 0, from !nat.add_comm ▸ rfl
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem padd_padd_pneg (p q : ℕ × ℕ) : padd (padd p q) (pneg q) ≡ p :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
calc pr1 p + pr1 q + pr2 q + pr2 p
|
2015-10-14 19:27:09 +00:00
|
|
|
|
= pr1 p + (pr1 q + pr2 q) + pr2 p : algebra.add.assoc
|
|
|
|
|
... = pr1 p + (pr1 q + pr2 q + pr2 p) : algebra.add.assoc
|
|
|
|
|
... = pr1 p + (pr2 q + pr1 q + pr2 p) : algebra.add.comm
|
|
|
|
|
... = pr1 p + (pr2 q + pr2 p + pr1 q) : algebra.add.right_comm
|
|
|
|
|
... = pr1 p + (pr2 p + pr2 q + pr1 q) : algebra.add.comm
|
|
|
|
|
... = pr2 p + pr2 q + pr1 q + pr1 p : algebra.add.comm
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem add_left_inv (a : ℤ) : -a + a = 0 :=
|
2014-12-17 18:32:38 +00:00
|
|
|
|
have H : repr (-a + a) ≡ repr 0, from
|
2014-12-05 22:03:24 +00:00
|
|
|
|
calc
|
2014-12-17 18:32:38 +00:00
|
|
|
|
repr (-a + a) ≡ padd (repr (neg a)) (repr a) : repr_add
|
|
|
|
|
... = padd (pneg (repr a)) (repr a) : repr_neg
|
2014-12-05 22:03:24 +00:00
|
|
|
|
... ≡ repr 0 : padd_pneg,
|
|
|
|
|
eq_of_repr_equiv_repr H
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-26 21:25:05 +00:00
|
|
|
|
/- nat abs -/
|
2014-12-22 19:21:22 +00:00
|
|
|
|
|
|
|
|
|
definition pabs (p : ℕ × ℕ) : ℕ := dist (pr1 p) (pr2 p)
|
|
|
|
|
|
|
|
|
|
theorem pabs_congr {p q : ℕ × ℕ} (H : p ≡ q) : pabs p = pabs q :=
|
|
|
|
|
calc
|
|
|
|
|
pabs p = nat_abs (abstr p) : nat_abs_abstr
|
|
|
|
|
... = nat_abs (abstr q) : abstr_eq H
|
|
|
|
|
... = pabs q : nat_abs_abstr
|
|
|
|
|
|
|
|
|
|
theorem nat_abs_eq_pabs_repr (a : ℤ) : nat_abs a = pabs (repr a) :=
|
|
|
|
|
calc
|
|
|
|
|
nat_abs a = nat_abs (abstr (repr a)) : abstr_repr
|
|
|
|
|
... = pabs (repr a) : nat_abs_abstr
|
|
|
|
|
|
|
|
|
|
theorem nat_abs_add_le (a b : ℤ) : nat_abs (a + b) ≤ nat_abs a + nat_abs b :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
calc
|
2015-08-17 03:23:03 +00:00
|
|
|
|
nat_abs (a + b) = pabs (repr (a + b)) : nat_abs_eq_pabs_repr
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = pabs (padd (repr a) (repr b)) : pabs_congr !repr_add
|
|
|
|
|
... ≤ pabs (repr a) + pabs (repr b) : dist_add_add_le_add_dist_dist
|
|
|
|
|
... = pabs (repr a) + nat_abs b : nat_abs_eq_pabs_repr
|
|
|
|
|
... = nat_abs a + nat_abs b : nat_abs_eq_pabs_repr
|
2014-12-22 19:21:22 +00:00
|
|
|
|
|
2015-10-22 20:45:20 +00:00
|
|
|
|
theorem nat_abs_neg_of_nat (n : nat) : nat_abs (neg_of_nat n) = n :=
|
|
|
|
|
begin cases n, reflexivity, reflexivity end
|
|
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
|
section
|
|
|
|
|
local attribute nat_abs [reducible]
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem nat_abs_mul : Π (a b : ℤ), nat_abs (a * b) = (nat_abs a) * (nat_abs b)
|
|
|
|
|
| (of_nat m) (of_nat n) := rfl
|
2015-10-22 20:45:20 +00:00
|
|
|
|
| (of_nat m) -[1+ n] := by rewrite [mul_of_nat_neg_succ_of_nat, nat_abs_neg_of_nat]
|
|
|
|
|
| -[1+ m] (of_nat n) := by rewrite [mul_neg_succ_of_nat_of_nat, nat_abs_neg_of_nat]
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| -[1+ m] -[1+ n] := rfl
|
2015-01-09 02:47:44 +00:00
|
|
|
|
end
|
2014-12-22 19:21:22 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
/- multiplication -/
|
|
|
|
|
|
|
|
|
|
definition pmul (p q : ℕ × ℕ) : ℕ × ℕ :=
|
|
|
|
|
(pr1 p * pr1 q + pr2 p * pr2 q, pr1 p * pr2 q + pr2 p * pr1 q)
|
|
|
|
|
|
|
|
|
|
theorem repr_neg_of_nat (m : ℕ) : repr (neg_of_nat m) = (0, m) :=
|
|
|
|
|
nat.cases_on m rfl (take m', rfl)
|
|
|
|
|
|
|
|
|
|
-- note: we have =, not just ≡
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem repr_mul : Π (a b : ℤ), repr (a * b) = pmul (repr a) (repr b)
|
|
|
|
|
| (of_nat m) (of_nat n) := calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
(m * n + 0 * 0, m * 0 + 0) = (m * n + 0 * 0, m * 0 + 0 * n) : by rewrite *zero_mul
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| (of_nat m) -[1+ n] := calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
repr ((m : int) * -[1+ n]) = (m * 0 + 0, m * succ n + 0 * 0) : repr_neg_of_nat
|
|
|
|
|
... = (m * 0 + 0 * succ n, m * succ n + 0 * 0) : by rewrite *zero_mul
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| -[1+ m] (of_nat n) := calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
repr (-[1+ m] * (n:int)) = (0 + succ m * 0, succ m * n) : repr_neg_of_nat
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = (0 + succ m * 0, 0 + succ m * n) : nat.zero_add
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = (0 * n + succ m * 0, 0 + succ m * n) : by rewrite zero_mul
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| -[1+ m] -[1+ n] := calc
|
2015-10-07 23:44:47 +00:00
|
|
|
|
(succ m * succ n, 0) = (succ m * succ n, 0 * succ n) : by rewrite zero_mul
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = (0 + succ m * succ n, 0 * succ n) : nat.zero_add
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
|
|
|
|
theorem equiv_mul_prep {xa ya xb yb xn yn xm ym : ℕ}
|
2014-08-20 02:32:44 +00:00
|
|
|
|
(H1 : xa + yb = ya + xb) (H2 : xn + ym = yn + xm)
|
2015-07-24 15:56:18 +00:00
|
|
|
|
: xa*xn+ya*yn+(xb*ym+yb*xm) = xa*yn+ya*xn+(xb*xm+yb*ym) :=
|
2015-10-22 21:35:27 +00:00
|
|
|
|
nat.add_right_cancel (calc
|
2015-07-24 15:56:18 +00:00
|
|
|
|
xa*xn+ya*yn + (xb*ym+yb*xm) + (yb*xn+xb*yn + (xb*xn+yb*yn))
|
2015-10-07 23:44:47 +00:00
|
|
|
|
= xa*xn+ya*yn + (yb*xn+xb*yn) + (xb*ym+yb*xm + (xb*xn+yb*yn)) : by rewrite add.comm4
|
2015-10-14 19:27:09 +00:00
|
|
|
|
... = xa*xn+ya*yn + (yb*xn+xb*yn) + (xb*xn+yb*yn + (xb*ym+yb*xm)) : by rewrite {xb*ym+yb*xm +_}nat.add_comm
|
|
|
|
|
... = xa*xn+yb*xn + (ya*yn+xb*yn) + (xb*xn+xb*ym + (yb*yn+yb*xm)) : by exact !congr_arg2 !add.comm4 !add.comm4
|
|
|
|
|
... = ya*xn+xb*xn + (xa*yn+yb*yn) + (xb*yn+xb*xm + (yb*xn+yb*ym)) : by rewrite[-+left_distrib,-+right_distrib]; exact H1 ▸ H2 ▸ rfl
|
|
|
|
|
... = ya*xn+xa*yn + (xb*xn+yb*yn) + (xb*yn+yb*xn + (xb*xm+yb*ym)) : by exact !congr_arg2 !add.comm4 !add.comm4
|
|
|
|
|
... = xa*yn+ya*xn + (xb*xn+yb*yn) + (xb*yn+yb*xn + (xb*xm+yb*ym)) : by rewrite {xa*yn + _}nat.add_comm
|
|
|
|
|
... = xa*yn+ya*xn + (xb*xn+yb*yn) + (yb*xn+xb*yn + (xb*xm+yb*ym)) : by rewrite {xb*yn + _}nat.add_comm
|
|
|
|
|
... = xa*yn+ya*xn + (yb*xn+xb*yn) + (xb*xn+yb*yn + (xb*xm+yb*ym)) : by rewrite (!add.comm4)
|
|
|
|
|
... = xa*yn+ya*xn + (yb*xn+xb*yn) + (xb*xm+yb*ym + (xb*xn+yb*yn)) : by rewrite {xb*xn+yb*yn + _}nat.add_comm
|
2015-10-07 23:44:47 +00:00
|
|
|
|
... = xa*yn+ya*xn + (xb*xm+yb*ym) + (yb*xn+xb*yn + (xb*xn+yb*yn)) : by rewrite add.comm4)
|
2015-07-24 15:56:18 +00:00
|
|
|
|
|
|
|
|
|
theorem pmul_congr {p p' q q' : ℕ × ℕ} : p ≡ p' → q ≡ q' → pmul p q ≡ pmul p' q' := equiv_mul_prep
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-12-05 22:03:24 +00:00
|
|
|
|
theorem pmul_comm (p q : ℕ × ℕ) : pmul p q = pmul q p :=
|
2015-10-14 19:27:09 +00:00
|
|
|
|
show (_,_) = (_,_),
|
|
|
|
|
begin
|
|
|
|
|
congruence,
|
|
|
|
|
{ congruence, repeat rewrite mul.comm },
|
|
|
|
|
{ rewrite algebra.add.comm, congruence, repeat rewrite mul.comm }
|
|
|
|
|
end
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem mul_comm (a b : ℤ) : a * b = b * a :=
|
2014-12-05 22:03:24 +00:00
|
|
|
|
eq_of_repr_equiv_repr
|
|
|
|
|
((calc
|
|
|
|
|
repr (a * b) = pmul (repr a) (repr b) : repr_mul
|
|
|
|
|
... = pmul (repr b) (repr a) : pmul_comm
|
2014-12-17 18:32:38 +00:00
|
|
|
|
... = repr (b * a) : repr_mul) ▸ !equiv.refl)
|
2014-12-05 22:03:24 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
private theorem pmul_assoc_prep {p1 p2 q1 q2 r1 r2 : ℕ} :
|
|
|
|
|
((p1*q1+p2*q2)*r1+(p1*q2+p2*q1)*r2, (p1*q1+p2*q2)*r2+(p1*q2+p2*q1)*r1) =
|
|
|
|
|
(p1*(q1*r1+q2*r2)+p2*(q1*r2+q2*r1), p1*(q1*r2+q2*r1)+p2*(q1*r1+q2*r2)) :=
|
2015-10-14 19:27:09 +00:00
|
|
|
|
begin
|
|
|
|
|
rewrite[+left_distrib,+right_distrib,*algebra.mul.assoc],
|
|
|
|
|
exact (congr_arg2 pair (!add.comm4 ⬝ (!congr_arg !nat.add_comm))
|
|
|
|
|
(!add.comm4 ⬝ (!congr_arg !nat.add_comm)))
|
|
|
|
|
end
|
2015-07-24 15:56:18 +00:00
|
|
|
|
|
|
|
|
|
theorem pmul_assoc (p q r: ℕ × ℕ) : pmul (pmul p q) r = pmul p (pmul q r) := pmul_assoc_prep
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem mul_assoc (a b c : ℤ) : (a * b) * c = a * (b * c) :=
|
2014-12-05 22:03:24 +00:00
|
|
|
|
eq_of_repr_equiv_repr
|
|
|
|
|
((calc
|
|
|
|
|
repr (a * b * c) = pmul (repr (a * b)) (repr c) : repr_mul
|
|
|
|
|
... = pmul (pmul (repr a) (repr b)) (repr c) : repr_mul
|
|
|
|
|
... = pmul (repr a) (pmul (repr b) (repr c)) : pmul_assoc
|
|
|
|
|
... = pmul (repr a) (repr (b * c)) : repr_mul
|
2014-12-17 18:32:38 +00:00
|
|
|
|
... = repr (a * (b * c)) : repr_mul) ▸ !equiv.refl)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem mul_one : Π (a : ℤ), a * 1 = a
|
|
|
|
|
| (of_nat m) := !int.zero_add -- zero_add happens to be def. = to this thm
|
2015-07-24 15:56:18 +00:00
|
|
|
|
| -[1+ m] := !nat.zero_add ▸ rfl
|
2015-07-22 17:13:19 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem one_mul (a : ℤ) : 1 * a = a :=
|
|
|
|
|
int.mul_comm a 1 ▸ int.mul_one a
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
private theorem mul_distrib_prep {a1 a2 b1 b2 c1 c2 : ℕ} :
|
2015-10-14 19:27:09 +00:00
|
|
|
|
((a1+b1)*c1+(a2+b2)*c2, (a1+b1)*c2+(a2+b2)*c1) =
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(a1*c1+a2*c2+(b1*c1+b2*c2), a1*c2+a2*c1+(b1*c2+b2*c1)) :=
|
2015-10-14 19:27:09 +00:00
|
|
|
|
begin
|
|
|
|
|
rewrite +right_distrib, congruence,
|
|
|
|
|
{rewrite add.comm4},
|
|
|
|
|
{rewrite add.comm4}
|
|
|
|
|
end
|
2015-07-24 15:56:18 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem mul_right_distrib (a b c : ℤ) : (a + b) * c = a * c + b * c :=
|
2014-12-05 22:03:24 +00:00
|
|
|
|
eq_of_repr_equiv_repr
|
|
|
|
|
(calc
|
|
|
|
|
repr ((a + b) * c) = pmul (repr (a + b)) (repr c) : repr_mul
|
2015-10-14 19:27:09 +00:00
|
|
|
|
... ≡ pmul (padd (repr a) (repr b)) (repr c) : pmul_congr !repr_add equiv.refl
|
2015-07-24 15:56:18 +00:00
|
|
|
|
... = padd (pmul (repr a) (repr c)) (pmul (repr b) (repr c)) : mul_distrib_prep
|
2015-10-14 19:27:09 +00:00
|
|
|
|
... = padd (repr (a * c)) (pmul (repr b) (repr c)) : repr_mul
|
|
|
|
|
... = padd (repr (a * c)) (repr (b * c)) : repr_mul
|
|
|
|
|
... ≡ repr (a * c + b * c) : repr_add)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem mul_left_distrib (a b c : ℤ) : a * (b + c) = a * b + a * c :=
|
2014-08-20 02:32:44 +00:00
|
|
|
|
calc
|
2015-10-22 20:36:45 +00:00
|
|
|
|
a * (b + c) = (b + c) * a : int.mul_comm
|
|
|
|
|
... = b * a + c * a : int.mul_right_distrib
|
|
|
|
|
... = a * b + c * a : int.mul_comm
|
|
|
|
|
... = a * b + a * c : int.mul_comm
|
2014-12-17 18:32:38 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem zero_ne_one : (0 : int) ≠ 1 :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
assume H : 0 = 1, !succ_ne_zero (of_nat.inj H)⁻¹
|
2014-12-17 18:32:38 +00:00
|
|
|
|
|
2015-10-22 20:36:45 +00:00
|
|
|
|
protected theorem eq_zero_or_eq_zero_of_mul_eq_zero {a b : ℤ} (H : a * b = 0) : a = 0 ∨ b = 0 :=
|
2015-08-17 03:23:03 +00:00
|
|
|
|
or.imp eq_zero_of_nat_abs_eq_zero eq_zero_of_nat_abs_eq_zero
|
2015-10-07 23:44:47 +00:00
|
|
|
|
(eq_zero_or_eq_zero_of_mul_eq_zero (by rewrite [-nat_abs_mul, H]))
|
|
|
|
|
|
2015-10-09 21:50:17 +00:00
|
|
|
|
protected definition integral_domain [reducible] [trans_instance] : algebra.integral_domain int :=
|
2015-10-07 23:44:47 +00:00
|
|
|
|
⦃algebra.integral_domain,
|
|
|
|
|
add := int.add,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
add_assoc := int.add_assoc,
|
2015-10-13 20:16:19 +00:00
|
|
|
|
zero := 0,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
zero_add := int.zero_add,
|
|
|
|
|
add_zero := int.add_zero,
|
2015-10-07 23:44:47 +00:00
|
|
|
|
neg := int.neg,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
add_left_inv := int.add_left_inv,
|
|
|
|
|
add_comm := int.add_comm,
|
2015-10-07 23:44:47 +00:00
|
|
|
|
mul := int.mul,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
mul_assoc := int.mul_assoc,
|
2015-10-07 23:44:47 +00:00
|
|
|
|
one := 1,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
one_mul := int.one_mul,
|
|
|
|
|
mul_one := int.mul_one,
|
|
|
|
|
left_distrib := int.mul_left_distrib,
|
|
|
|
|
right_distrib := int.mul_right_distrib,
|
|
|
|
|
mul_comm := int.mul_comm,
|
|
|
|
|
zero_ne_one := int.zero_ne_one,
|
|
|
|
|
eq_zero_or_eq_zero_of_mul_eq_zero := @int.eq_zero_or_eq_zero_of_mul_eq_zero⦄
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
|
|
|
|
definition int_has_sub [reducible] [instance] [priority int.prio] : has_sub int :=
|
2015-10-09 19:47:55 +00:00
|
|
|
|
has_sub.mk has_sub.sub
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
|
|
|
|
definition int_has_dvd [reducible] [instance] [priority int.prio] : has_dvd int :=
|
2015-10-09 19:47:55 +00:00
|
|
|
|
has_dvd.mk has_dvd.dvd
|
2014-12-17 18:32:38 +00:00
|
|
|
|
|
2015-02-25 19:14:57 +00:00
|
|
|
|
/- additional properties -/
|
2015-10-09 01:35:37 +00:00
|
|
|
|
theorem of_nat_sub {m n : ℕ} (H : m ≥ n) : of_nat (m - n) = of_nat m - of_nat n :=
|
2015-09-01 02:53:11 +00:00
|
|
|
|
assert m - n + n = m, from nat.sub_add_cancel H,
|
|
|
|
|
begin
|
|
|
|
|
symmetry,
|
2015-10-09 01:35:37 +00:00
|
|
|
|
apply algebra.sub_eq_of_eq_add,
|
2015-10-09 21:17:09 +00:00
|
|
|
|
rewrite [-of_nat_add, this]
|
2015-09-01 02:53:11 +00:00
|
|
|
|
end
|
|
|
|
|
|
2015-06-29 05:16:58 +00:00
|
|
|
|
theorem neg_succ_of_nat_eq' (m : ℕ) : -[1+ m] = -m - 1 :=
|
2015-10-07 23:44:47 +00:00
|
|
|
|
by rewrite [neg_succ_of_nat_eq, neg_add]
|
2015-03-13 03:43:19 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
definition succ (a : ℤ) := a + (succ zero)
|
|
|
|
|
definition pred (a : ℤ) := a - (succ zero)
|
2015-05-15 00:18:29 +00:00
|
|
|
|
theorem pred_succ (a : ℤ) : pred (succ a) = a := !sub_add_cancel
|
|
|
|
|
theorem succ_pred (a : ℤ) : succ (pred a) = a := !add_sub_cancel
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
2015-05-15 00:18:29 +00:00
|
|
|
|
theorem neg_succ (a : ℤ) : -succ a = pred (-a) :=
|
|
|
|
|
by rewrite [↑succ,neg_add]
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
2015-05-15 00:18:29 +00:00
|
|
|
|
theorem succ_neg_succ (a : ℤ) : succ (-succ a) = -a :=
|
|
|
|
|
by rewrite [neg_succ,succ_pred]
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
2015-05-15 00:18:29 +00:00
|
|
|
|
theorem neg_pred (a : ℤ) : -pred a = succ (-a) :=
|
2015-10-12 03:29:31 +00:00
|
|
|
|
by rewrite [↑pred,neg_sub,sub_eq_add_neg,add.comm]
|
2015-10-07 23:44:47 +00:00
|
|
|
|
|
2015-05-15 00:18:29 +00:00
|
|
|
|
theorem pred_neg_pred (a : ℤ) : pred (-pred a) = -a :=
|
|
|
|
|
by rewrite [neg_pred,pred_succ]
|
|
|
|
|
|
|
|
|
|
theorem pred_nat_succ (n : ℕ) : pred (nat.succ n) = n := pred_succ n
|
|
|
|
|
theorem neg_nat_succ (n : ℕ) : -nat.succ n = pred (-n) := !neg_succ
|
|
|
|
|
theorem succ_neg_nat_succ (n : ℕ) : succ (-nat.succ n) = -n := !succ_neg_succ
|
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition rec_nat_on [unfold 2] {P : ℤ → Type} (z : ℤ) (H0 : P 0)
|
2015-05-15 00:18:29 +00:00
|
|
|
|
(Hsucc : Π⦃n : ℕ⦄, P n → P (succ n)) (Hpred : Π⦃n : ℕ⦄, P (-n) → P (-nat.succ n)) : P z :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
int.rec (nat.rec H0 Hsucc) (λn, nat.rec H0 Hpred (nat.succ n)) z
|
2015-05-15 00:18:29 +00:00
|
|
|
|
|
|
|
|
|
--the only computation rule of rec_nat_on which is not definitional
|
|
|
|
|
theorem rec_nat_on_neg {P : ℤ → Type} (n : nat) (H0 : P zero)
|
|
|
|
|
(Hsucc : Π⦃n : nat⦄, P n → P (succ n)) (Hpred : Π⦃n : nat⦄, P (-n) → P (-nat.succ n))
|
|
|
|
|
: rec_nat_on (-nat.succ n) H0 Hsucc Hpred = Hpred (rec_nat_on (-n) H0 Hsucc Hpred) :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
nat.rec rfl (λn H, rfl) n
|
2015-05-15 00:18:29 +00:00
|
|
|
|
|
2014-09-17 21:39:05 +00:00
|
|
|
|
end int
|