feat(library/data/list/comb): show that (list A) is isomorphic to A if A is isomorphic to nat
This commit is contained in:
parent
303e9031d6
commit
de3d0e4162
3 changed files with 124 additions and 0 deletions
|
@ -472,6 +472,49 @@ lemma list_equiv_of_equiv {A B : Type} : A ≃ B → list A ≃ list B
|
|||
mk (map f) (map g)
|
||||
begin intros, rewrite [map_map, id_of_left_inverse l, map_id] end
|
||||
begin intros, rewrite [map_map, id_of_righ_inverse r, map_id] end
|
||||
|
||||
private definition to_nat : list nat → nat
|
||||
| [] := 0
|
||||
| (x::xs) := succ (mkpair (to_nat xs) x)
|
||||
|
||||
open prod.ops
|
||||
|
||||
private definition of_nat.F : Π (n : nat), (Π m, m < n → list nat) → list nat
|
||||
| 0 f := []
|
||||
| (succ n) f := (unpair n).2 :: f (unpair n).1 (unpair_lt n)
|
||||
|
||||
private definition of_nat : nat → list nat :=
|
||||
well_founded.fix of_nat.F
|
||||
|
||||
private lemma of_nat_zero : of_nat 0 = [] :=
|
||||
well_founded.fix_eq of_nat.F 0
|
||||
|
||||
private lemma of_nat_succ (n : nat)
|
||||
: of_nat (succ n) = (unpair n).2 :: of_nat (unpair n).1 :=
|
||||
well_founded.fix_eq of_nat.F (succ n)
|
||||
|
||||
private lemma to_nat_of_nat (n : nat) : to_nat (of_nat n) = n :=
|
||||
nat.case_strong_induction_on n
|
||||
_
|
||||
(λ n ih,
|
||||
begin
|
||||
rewrite of_nat_succ, unfold to_nat,
|
||||
have to_nat (of_nat (unpair n).1) = (unpair n).1, from ih _ (le_of_lt_succ (unpair_lt n)),
|
||||
rewrite this, rewrite mkpair_unpair
|
||||
end)
|
||||
|
||||
private lemma of_nat_to_nat : ∀ (l : list nat), of_nat (to_nat l) = l
|
||||
| [] := rfl
|
||||
| (x::xs) := begin unfold to_nat, rewrite of_nat_succ, rewrite *unpair_mkpair, esimp, congruence, apply of_nat_to_nat end
|
||||
|
||||
lemma list_nat_equiv_nat : list nat ≃ nat :=
|
||||
mk to_nat of_nat of_nat_to_nat to_nat_of_nat
|
||||
|
||||
lemma list_equiv_self_of_equiv_nat {A : Type} : A ≃ nat → list A ≃ A :=
|
||||
suppose A ≃ nat, calc
|
||||
list A ≃ list nat : list_equiv_of_equiv this
|
||||
... ≃ nat : list_nat_equiv_nat
|
||||
... ≃ A : this
|
||||
end
|
||||
end list
|
||||
|
||||
|
|
|
@ -69,4 +69,26 @@ by_cases
|
|||
esimp [unpair],
|
||||
rewrite [add.assoc (a * a) a b, sqrt_offset_eq `a + b ≤ a + a`, *add_sub_cancel_left, if_neg `¬ a + b < a`]
|
||||
end)
|
||||
|
||||
open prod.ops
|
||||
|
||||
theorem unpair_lt_aux {n : nat} : n ≥ 1 → (unpair n).1 < n :=
|
||||
suppose n ≥ 1,
|
||||
or.elim (eq_or_lt_of_le this)
|
||||
(suppose 1 = n, by subst n; exact dec_trivial)
|
||||
(suppose n > 1,
|
||||
let s := sqrt n in
|
||||
by_cases
|
||||
(suppose h : n - s*s < s,
|
||||
assert n > 0, from lt_of_succ_lt `n > 1`,
|
||||
assert sqrt n > 0, from sqrt_pos_of_pos this,
|
||||
assert sqrt n * sqrt n > 0, from mul_pos this this,
|
||||
begin unfold unpair, rewrite [if_pos h], esimp, exact sub_lt `n > 0` `sqrt n * sqrt n > 0` end)
|
||||
(suppose ¬ n - s*s < s, begin unfold unpair, rewrite [if_neg this], esimp, apply sqrt_lt `n > 1` end))
|
||||
|
||||
theorem unpair_lt : ∀ (n : nat), (unpair n).1 < succ n
|
||||
| 0 := dec_trivial
|
||||
| (succ n) :=
|
||||
have (unpair (succ n)).1 < succ n, from unpair_lt_aux dec_trivial,
|
||||
lt.step this
|
||||
end nat
|
||||
|
|
|
@ -26,6 +26,14 @@ theorem sqrt_aux_of_le : ∀ {s n : nat}, s * s ≤ n → sqrt_aux s n = s
|
|||
| 0 n h := rfl
|
||||
| (succ s) n h := by rewrite [sqrt_aux_succ_of_pos h]
|
||||
|
||||
theorem sqrt_aux_le : ∀ (s n), sqrt_aux s n ≤ s
|
||||
| 0 n := !zero_le
|
||||
| (succ s) n := or.elim (em ((succ s)*(succ s) ≤ n))
|
||||
(λ h, begin unfold sqrt_aux, rewrite [if_pos h] end)
|
||||
(λ h,
|
||||
assert sqrt_aux s n ≤ succ s, from le.step (sqrt_aux_le s n),
|
||||
begin unfold sqrt_aux, rewrite [if_neg h], assumption end)
|
||||
|
||||
definition sqrt (n : nat) : nat :=
|
||||
sqrt_aux n n
|
||||
|
||||
|
@ -61,6 +69,57 @@ 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
|
||||
| 0 h := absurd h dec_trivial
|
||||
| 1 h := absurd h dec_trivial
|
||||
| (succ (succ n)) h :=
|
||||
have 1 < succ (succ n), from dec_trivial,
|
||||
assert succ (succ n) * 1 < succ (succ n) * succ (succ n), from mul_lt_mul_of_pos_left this dec_trivial,
|
||||
by rewrite [mul_one at this]; exact this
|
||||
|
||||
theorem sqrt_le (n : nat) : sqrt n ≤ n :=
|
||||
calc sqrt n ≤ sqrt n * sqrt n : le_squared
|
||||
... ≤ n : sqrt_lower
|
||||
|
||||
theorem eq_zero_of_sqrt_eq_zero {n : nat} : sqrt n = 0 → n = 0 :=
|
||||
suppose sqrt n = 0,
|
||||
assert n ≤ sqrt n * sqrt n + sqrt n + sqrt n, from !sqrt_upper,
|
||||
have n ≤ 0, by rewrite [*`sqrt n = 0` at this]; exact this,
|
||||
eq_zero_of_le_zero this
|
||||
|
||||
theorem le_three_of_sqrt_eq_one {n : nat} : sqrt n = 1 → n ≤ 3 :=
|
||||
suppose sqrt n = 1,
|
||||
assert n ≤ sqrt n * sqrt n + sqrt n + sqrt n, from !sqrt_upper,
|
||||
show n ≤ 3, by rewrite [*`sqrt n = 1` at this]; exact this
|
||||
|
||||
theorem sqrt_lt : ∀ {n : nat}, n > 1 → sqrt n < n
|
||||
| 0 h := absurd h dec_trivial
|
||||
| 1 h := absurd h dec_trivial
|
||||
| 2 h := dec_trivial
|
||||
| 3 h := dec_trivial
|
||||
| (n+4) h :=
|
||||
have sqrt (n+4) > 1, from by_contradiction
|
||||
(suppose ¬ sqrt (n+4) > 1,
|
||||
have sqrt (n+4) ≤ 1, from le_of_not_gt this,
|
||||
or.elim (eq_or_lt_of_le this)
|
||||
(suppose sqrt (n+4) = 1,
|
||||
have n+4 ≤ 3, from le_three_of_sqrt_eq_one this,
|
||||
absurd this dec_trivial)
|
||||
(suppose sqrt (n+4) < 1,
|
||||
have sqrt (n+4) = 0, from eq_zero_of_le_zero (le_of_lt_succ this),
|
||||
have n + 4 = 0, from eq_zero_of_sqrt_eq_zero this,
|
||||
absurd this dec_trivial)),
|
||||
calc sqrt (n+4) < sqrt (n+4) * sqrt (n+4) : lt_squared this
|
||||
... ≤ n+4 : sqrt_lower
|
||||
|
||||
theorem sqrt_pos_of_pos {n : nat} : n > 0 → sqrt n > 0 :=
|
||||
suppose n > 0,
|
||||
have sqrt n ≠ 0, from
|
||||
suppose sqrt n = 0,
|
||||
assert n = 0, from eq_zero_of_sqrt_eq_zero this,
|
||||
by subst n; exact absurd `0 > 0` !lt.irrefl,
|
||||
pos_of_ne_zero this
|
||||
|
||||
theorem sqrt_aux_offset_eq {n k : nat} (h₁ : k ≤ n + n) : ∀ {s}, s ≥ n → sqrt_aux s (n*n + k) = n
|
||||
| 0 h₂ :=
|
||||
assert neqz : n = 0, from eq_zero_of_le_zero h₂,
|
||||
|
|
Loading…
Reference in a new issue