feat(library/data/real): add more theorems to reals
This commit is contained in:
parent
f929f82cb4
commit
954e30b59a
2 changed files with 33 additions and 0 deletions
|
@ -292,6 +292,26 @@ theorem canon_bound {s : seq} (Hs : regular s) (n : ℕ+) : abs (s n) ≤ rat_of
|
||||||
... = of_nat (ubound (abs (s pone)) + (1 + 1)) : of_nat_add
|
... = of_nat (ubound (abs (s pone)) + (1 + 1)) : of_nat_add
|
||||||
... = of_nat (ubound (abs (s pone)) + 1 + 1) : nat.add.assoc
|
... = of_nat (ubound (abs (s pone)) + 1 + 1) : nat.add.assoc
|
||||||
|
|
||||||
|
theorem bdd_of_regular {s : seq} (H : regular s) : ∃ b : ℚ, ∀ n : ℕ+, s n ≤ b :=
|
||||||
|
begin
|
||||||
|
existsi rat_of_pnat (K s),
|
||||||
|
intros,
|
||||||
|
apply rat.le.trans,
|
||||||
|
apply le_abs_self,
|
||||||
|
apply canon_bound H
|
||||||
|
end
|
||||||
|
|
||||||
|
theorem bdd_of_regular_strict {s : seq} (H : regular s) : ∃ b : ℚ, ∀ n : ℕ+, s n < b :=
|
||||||
|
begin
|
||||||
|
cases bdd_of_regular H with [b, Hb],
|
||||||
|
existsi b + 1,
|
||||||
|
intro n,
|
||||||
|
apply rat.lt_of_le_of_lt,
|
||||||
|
apply Hb,
|
||||||
|
apply rat.lt_add_of_pos_right,
|
||||||
|
apply rat.zero_lt_one
|
||||||
|
end
|
||||||
|
|
||||||
definition K₂ (s t : seq) := max (K s) (K t)
|
definition K₂ (s t : seq) := max (K s) (K t)
|
||||||
|
|
||||||
theorem K₂_symm (s t : seq) : K₂ s t = K₂ t s :=
|
theorem K₂_symm (s t : seq) : K₂ s t = K₂ t s :=
|
||||||
|
|
|
@ -990,6 +990,19 @@ theorem lt_of_const_lt_const {a b : ℚ} (H : s_lt (const a) (const b)) : a < b
|
||||||
apply pnat.inv_pos
|
apply pnat.inv_pos
|
||||||
end
|
end
|
||||||
|
|
||||||
|
theorem s_le_of_le_pointwise {s t : seq} (Hs : regular s) (Ht : s.regular t)
|
||||||
|
(H : ∀ n : ℕ+, s n ≤ t n) : s_le s t :=
|
||||||
|
begin
|
||||||
|
rewrite [↑s_le, ↑nonneg, ↑sadd, ↑sneg],
|
||||||
|
intros,
|
||||||
|
apply rat.le.trans,
|
||||||
|
apply iff.mpr !neg_nonpos_iff_nonneg,
|
||||||
|
apply le_of_lt,
|
||||||
|
apply inv_pos,
|
||||||
|
apply iff.mpr !sub_nonneg_iff_le,
|
||||||
|
apply H
|
||||||
|
end
|
||||||
|
|
||||||
-------- lift to reg_seqs
|
-------- lift to reg_seqs
|
||||||
definition r_lt (s t : reg_seq) := s_lt (reg_seq.sq s) (reg_seq.sq t)
|
definition r_lt (s t : reg_seq) := s_lt (reg_seq.sq s) (reg_seq.sq t)
|
||||||
definition r_le (s t : reg_seq) := s_le (reg_seq.sq s) (reg_seq.sq t)
|
definition r_le (s t : reg_seq) := s_le (reg_seq.sq s) (reg_seq.sq t)
|
||||||
|
|
Loading…
Reference in a new issue