feat(library/data/real): fill in lots of sorrys about pnats

This commit is contained in:
Rob Lewis 2015-06-11 17:01:55 +10:00 committed by Leonardo de Moura
parent 4b38e14586
commit cf7c85e5fd
4 changed files with 12 additions and 13 deletions

View file

@ -11,8 +11,8 @@ To do:
o Rename things and possibly make theorems private
-/
import data.nat data.rat.order
open nat eq eq.ops
import data.nat data.rat.order data.nat.pnat
open nat eq eq.ops pnat
open -[coercions] rat
local notation 0 := rat.of_num 0
local notation 1 := rat.of_num 1
@ -21,7 +21,7 @@ local notation 1 := rat.of_num 1
-----------------------------------------------
-- positive naturals
inductive pnat : Type :=
/-inductive pnat : Type :=
pos : Π n : nat, n > 0 → pnat
notation `+` := pnat
@ -169,7 +169,7 @@ theorem pnat.lt_of_not_le {p q : +} (H : ¬ p ≤ q) : q < p := sorry
theorem pnat.inv_cancel (p : +) : pnat.to_rat p * p⁻¹ = (1 : ) := sorry
theorem pnat.inv_cancel_right (p : +) : p⁻¹ * pnat.to_rat p = (1 : ) := sorry
theorem pnat.inv_cancel_right (p : +) : p⁻¹ * pnat.to_rat p = (1 : ) := sorry-/
-------------------------------------
-- theorems to add to (ordered) field and/or rat
@ -220,6 +220,9 @@ theorem s_mul_assoc_lemma_4 {n : +} {ε q : } (Hε : ε > 0) (Hq : q > 0)
-------------------------------------
-- small helper lemmas
theorem s_mul_assoc_lemma_3 (a b n : +) (p : ) :
p * ((a * n)⁻¹ + (b * n)⁻¹) = p * (a⁻¹ + b⁻¹) * n⁻¹ := sorry
theorem find_thirds (a b : ) : ∃ n : +, a + n⁻¹ + n⁻¹ + n⁻¹ < a + b := sorry
theorem squeeze {a b : } (H : ∀ j : +, a ≤ b + j⁻¹ + j⁻¹ + j⁻¹) : a ≤ b :=
@ -1073,7 +1076,7 @@ theorem zero_nequiv_one : ¬ zero ≡ one :=
... = abs 1 : abs_of_pos zero_lt_one
... ≤ 2⁻¹ : H,
let H'' := ge_of_inv_le H',
apply absurd (one_lt_two) (pnat.not_lt_of_le (pnat.le_of_lt H''))
apply absurd (one_lt_two) (pnat.not_lt_of_le H'')
end
---------------------------------------------

View file

@ -237,6 +237,8 @@ end s
namespace real
theorem p_add_fractions (n : +) : (2 * n)⁻¹ + (2 * 3 * n)⁻¹ + (3 * n)⁻¹ = n⁻¹ := sorry
theorem rewrite_helper9 (a b c : ) : b - c = (b - a) - (c - a) := sorry
theorem rewrite_helper10 (a b c d : ) : c - d = (c - a) + (a - b) + (b - d) := sorry

View file

@ -14,7 +14,7 @@ open -[coercions] rat
open -[coercions] nat
local notation 0 := rat.of_num 0
local notation 1 := rat.of_num 1
open eq.ops
open eq.ops pnat
local notation 2 := pnat.pos (nat.of_num 2) dec_trivial
@ -44,9 +44,6 @@ theorem abs_abs_sub_abs_le_abs_sub (a b : ) : abs (abs a - abs b) ≤ abs (a
theorem abs_one_div (q : ) : abs (1 / q) = 1 / abs q := sorry
theorem div_le_pnat (q : ) (n : +) (H : q ≥ n⁻¹) : 1 / q ≤ pnat.to_rat n := sorry
theorem pnat_cancel' (n m : +) : (n * n * m)⁻¹ * (pnat.to_rat n * pnat.to_rat n) = m⁻¹ := sorry
-- does this not exist already??
theorem forall_of_not_exists {A : Type} {P : A → Prop} (H : ¬ ∃ a : A, P a) : ∀ a : A, ¬ P a :=

View file

@ -14,16 +14,13 @@ To do:
import data.real.basic data.rat data.nat
open -[coercions] rat
open -[coercions] nat
open eq eq.ops
open eq eq.ops pnat
local notation 0 := rat.of_num 0
local notation 1 := rat.of_num 1
----------------------------------------------------------------------------------------------------
-- pnat theorems
theorem inv_add_lt_left (p q : +) : (p + q)⁻¹ < p⁻¹ := sorry
theorem pnat_lt_add_left (p q : +) : p < p + q := sorry
notation 2 := pnat.pos (of_num 2) dec_trivial