# HG changeset patch # User paulson # Date 1519378094 0 # Node ID db202a00a29caeb9c34426ad764d98676a4b52ad # Parent 3c02b0522e230e951869308638d6fd3a041e4ead fixing ennreal using add_mono1; shifting results from linordered_semidom to linordered_nonzero_semiring diff -r 3c02b0522e23 -r db202a00a29c src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy Thu Feb 22 22:58:38 2018 +0000 +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Fri Feb 23 09:28:14 2018 +0000 @@ -11,7 +11,7 @@ lemma ereal_ineq_diff_add: assumes "b \ (-\::ereal)" "a \ b" shows "a = b + (a-b)" -by (metis add.commute assms(1) assms(2) ereal_eq_minus_iff ereal_minus_le_iff ereal_plus_eq_PInfty) +by (metis add.commute assms ereal_eq_minus_iff ereal_minus_le_iff ereal_plus_eq_PInfty) lemma Limsup_const_add: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" @@ -345,7 +345,11 @@ (transfer ; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+ instance ennreal :: linordered_nonzero_semiring - proof qed (transfer; simp) +proof + fix a b::ennreal + show "a < b \ a + 1 < b + 1" + by transfer (simp add: add_right_mono ereal_add_cancel_right less_le) +qed (transfer; simp) instance ennreal :: strict_ordered_ab_semigroup_add proof diff -r 3c02b0522e23 -r db202a00a29c src/HOL/Nat.thy --- a/src/HOL/Nat.thy Thu Feb 22 22:58:38 2018 +0000 +++ b/src/HOL/Nat.thy Fri Feb 23 09:28:14 2018 +0000 @@ -1774,7 +1774,7 @@ class ring_char_0 = ring_1 + semiring_char_0 -context linordered_semidom +context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" @@ -1783,8 +1783,21 @@ lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) +lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" + by (auto simp: le_iff_add intro!: add_increasing2) + lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" - by (induct m n rule: diff_induct) (simp_all add: add_pos_nonneg) +proof(induct m n rule: diff_induct) + case (1 m) then show ?case + by auto +next + case (2 n) then show ?case + by (simp add: add_pos_nonneg) +next + case (3 m n) + then show ?case + by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) +qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) @@ -1795,7 +1808,7 @@ lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp -text \Every \linordered_semidom\ has characteristic zero.\ +text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: eq_iff) @@ -1810,6 +1823,14 @@ end + +context linordered_semidom +begin +subclass linordered_nonzero_semiring .. +subclass semiring_char_0 .. +end + + context linordered_idom begin