changes made due to new Ring_and_Field theory
authorobua
Tue, 11 May 2004 20:11:08 +0200
changeset 1473883f1a514dcb4
parent 14737 77ea79aed99d
child 14739 86c6f272ef79
changes made due to new Ring_and_Field theory
src/HOL/Algebra/abstract/Ring.ML
src/HOL/Algebra/abstract/Ring.thy
src/HOL/Algebra/poly/LongDiv.thy
src/HOL/Complex/CLim.thy
src/HOL/Complex/NSComplex.thy
src/HOL/Finite_Set.thy
src/HOL/Hyperreal/HyperDef.thy
src/HOL/Hyperreal/HyperNat.thy
src/HOL/Hyperreal/MacLaurin.ML
src/HOL/Hyperreal/MacLaurin.thy
src/HOL/Hyperreal/MacLaurin_lemmas.ML
src/HOL/IMP/Compiler.thy
src/HOL/Integ/Bin.thy
src/HOL/Integ/IntArith.thy
src/HOL/Integ/IntDef.thy
src/HOL/Integ/NatBin.thy
src/HOL/Integ/Parity.thy
src/HOL/Integ/Presburger.thy
src/HOL/Integ/int_arith1.ML
src/HOL/Integ/int_factor_simprocs.ML
src/HOL/IsaMakefile
src/HOL/LOrder.thy
src/HOL/Library/Multiset.thy
src/HOL/Matrix/Matrix.thy
src/HOL/Nat.thy
src/HOL/NumberTheory/WilsonBij.thy
src/HOL/NumberTheory/WilsonRuss.thy
src/HOL/OrderedGroup.thy
src/HOL/Power.thy
src/HOL/Presburger.thy
src/HOL/Real/PReal.thy
src/HOL/Real/RealDef.thy
src/HOL/Ring_and_Field.thy
src/HOL/UNITY/Comp/AllocBase.thy
src/HOL/arith_data.ML
src/HOL/ex/Lagrange.thy
     1.1 --- a/src/HOL/Algebra/abstract/Ring.ML	Tue May 11 14:00:02 2004 +0200
     1.2 +++ b/src/HOL/Algebra/abstract/Ring.ML	Tue May 11 20:11:08 2004 +0200
     1.3 @@ -5,9 +5,9 @@
     1.4  *)
     1.5  
     1.6  (*
     1.7 -val a_assoc = thm "plus_ac0.assoc";
     1.8 -val l_zero = thm "plus_ac0.zero";
     1.9 -val a_comm = thm "plus_ac0.commute";
    1.10 +val a_assoc = thm "semigroup_add.add_assoc";
    1.11 +val l_zero = thm "comm_monoid_add.add_0";
    1.12 +val a_comm = thm "ab_semigroup_add.add_commute";
    1.13  
    1.14  section "Rings";
    1.15  
     2.1 --- a/src/HOL/Algebra/abstract/Ring.thy	Tue May 11 14:00:02 2004 +0200
     2.2 +++ b/src/HOL/Algebra/abstract/Ring.thy	Tue May 11 20:11:08 2004 +0200
     2.3 @@ -117,37 +117,37 @@
     2.4  
     2.5  (* Basic facts --- move to HOL!!! *)
     2.6  
     2.7 -lemma natsum_0 [simp]: "setsum f {..(0::nat)} = (f 0::'a::plus_ac0)"
     2.8 +lemma natsum_0 [simp]: "setsum f {..(0::nat)} = (f 0::'a::comm_monoid_add)"
     2.9  by simp
    2.10  
    2.11  lemma natsum_Suc [simp]:
    2.12 -  "setsum f {..Suc n} = (f (Suc n) + setsum f {..n}::'a::plus_ac0)"
    2.13 +  "setsum f {..Suc n} = (f (Suc n) + setsum f {..n}::'a::comm_monoid_add)"
    2.14  by (simp add: atMost_Suc)
    2.15  
    2.16  lemma natsum_Suc2:
    2.17 -  "setsum f {..Suc n} = (setsum (%i. f (Suc i)) {..n} + f 0::'a::plus_ac0)"
    2.18 +  "setsum f {..Suc n} = (setsum (%i. f (Suc i)) {..n} + f 0::'a::comm_monoid_add)"
    2.19  proof (induct n)
    2.20    case 0 show ?case by simp
    2.21  next
    2.22 -  case Suc thus ?case by (simp add: plus_ac0.assoc) 
    2.23 +  case Suc thus ?case by (simp add: semigroup_add.add_assoc) 
    2.24  qed
    2.25  
    2.26  lemma natsum_cong [cong]:
    2.27 -  "!!k. [| j = k; !!i::nat. i <= k ==> f i = (g i::'a::plus_ac0) |] ==>
    2.28 +  "!!k. [| j = k; !!i::nat. i <= k ==> f i = (g i::'a::comm_monoid_add) |] ==>
    2.29          setsum f {..j} = setsum g {..k}"
    2.30  by (induct j) auto
    2.31  
    2.32 -lemma natsum_zero [simp]: "setsum (%i. 0) {..n::nat} = (0::'a::plus_ac0)"
    2.33 +lemma natsum_zero [simp]: "setsum (%i. 0) {..n::nat} = (0::'a::comm_monoid_add)"
    2.34  by (induct n) simp_all
    2.35  
    2.36  lemma natsum_add [simp]:
    2.37 -  "!!f::nat=>'a::plus_ac0.
    2.38 +  "!!f::nat=>'a::comm_monoid_add.
    2.39     setsum (%i. f i + g i) {..n::nat} = setsum f {..n} + setsum g {..n}"
    2.40 -by (induct n) (simp_all add: plus_ac0)
    2.41 +by (induct n) (simp_all add: add_ac)
    2.42  
    2.43  (* Facts specific to rings *)
    2.44  
    2.45 -instance ring < plus_ac0
    2.46 +instance ring < comm_monoid_add
    2.47  proof
    2.48    fix x y z
    2.49    show "(x::'a::ring) + y = y + x" by (rule a_comm)
     3.1 --- a/src/HOL/Algebra/poly/LongDiv.thy	Tue May 11 14:00:02 2004 +0200
     3.2 +++ b/src/HOL/Algebra/poly/LongDiv.thy	Tue May 11 20:11:08 2004 +0200
     3.3 @@ -22,7 +22,7 @@
     3.4    apply (simp)
     3.5    apply (force)
     3.6    apply (simp)
     3.7 -  apply (subst plus_ac0.commute[of m])
     3.8 +  apply (subst ab_semigroup_add.add_commute[of m])
     3.9    apply (simp)
    3.10    done
    3.11  
     4.1 --- a/src/HOL/Complex/CLim.thy	Tue May 11 14:00:02 2004 +0200
     4.2 +++ b/src/HOL/Complex/CLim.thy	Tue May 11 20:11:08 2004 +0200
     4.3 @@ -17,7 +17,7 @@
     4.4  by (simp add: numeral_2_eq_2)
     4.5  
     4.6  text{*Changing the quantified variable. Install earlier?*}
     4.7 -lemma all_shift: "(\<forall>x::'a::ring. P x) = (\<forall>x. P (x-a))";
     4.8 +lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
     4.9  apply auto 
    4.10  apply (drule_tac x="x+a" in spec) 
    4.11  apply (simp add: diff_minus add_assoc) 
     5.1 --- a/src/HOL/Complex/NSComplex.thy	Tue May 11 14:00:02 2004 +0200
     5.2 +++ b/src/HOL/Complex/NSComplex.thy	Tue May 11 20:11:08 2004 +0200
     5.3 @@ -412,7 +412,7 @@
     5.4  
     5.5  lemma hcomplex_add_minus_eq_minus:
     5.6        "x + y = (0::hcomplex) ==> x = -y"
     5.7 -apply (drule Ring_and_Field.equals_zero_I)
     5.8 +apply (drule OrderedGroup.equals_zero_I)
     5.9  apply (simp add: minus_equation_iff [of x y])
    5.10  done
    5.11  
    5.12 @@ -429,7 +429,7 @@
    5.13  subsection{*More Multiplication Laws*}
    5.14  
    5.15  lemma hcomplex_mult_one_right: "z * (1::hcomplex) = z"
    5.16 -by (rule Ring_and_Field.mult_1_right)
    5.17 +by (rule OrderedGroup.mult_1_right)
    5.18  
    5.19  lemma hcomplex_mult_minus_one [simp]: "- 1 * (z::hcomplex) = -z"
    5.20  by simp
    5.21 @@ -454,7 +454,7 @@
    5.22  by (simp add: hcomplex_diff_def hcomplex_minus hcomplex_add complex_diff_def)
    5.23  
    5.24  lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)"
    5.25 -by (rule Ring_and_Field.diff_eq_eq)
    5.26 +by (rule OrderedGroup.diff_eq_eq)
    5.27  
    5.28  lemma hcomplex_add_divide_distrib: "(x+y)/(z::hcomplex) = x/z + y/z"
    5.29  by (rule Ring_and_Field.add_divide_distrib)
     6.1 --- a/src/HOL/Finite_Set.thy	Tue May 11 14:00:02 2004 +0200
     6.2 +++ b/src/HOL/Finite_Set.thy	Tue May 11 20:11:08 2004 +0200
     6.3 @@ -742,15 +742,15 @@
     6.4  subsection {* Generalized summation over a set *}
     6.5  
     6.6  constdefs
     6.7 -  setsum :: "('a => 'b) => 'a set => 'b::plus_ac0"
     6.8 +  setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
     6.9    "setsum f A == if finite A then fold (op + o f) 0 A else 0"
    6.10  
    6.11  syntax
    6.12 -  "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("(3\<Sum>_:_. _)" [0, 51, 10] 10)
    6.13 +  "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_:_. _)" [0, 51, 10] 10)
    6.14  syntax (xsymbols)
    6.15 -  "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
    6.16 +  "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
    6.17  syntax (HTML output)
    6.18 -  "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
    6.19 +  "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
    6.20  translations
    6.21    "\<Sum>i:A. b" == "setsum (%i. b) A"  -- {* Beware of argument permutation! *}
    6.22  
    6.23 @@ -761,7 +761,7 @@
    6.24  lemma setsum_insert [simp]:
    6.25      "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
    6.26    by (simp add: setsum_def
    6.27 -    LC.fold_insert [OF LC.intro] plus_ac0_left_commute)
    6.28 +    LC.fold_insert [OF LC.intro] add_left_commute)
    6.29  
    6.30  lemma setsum_reindex [rule_format]: "finite B ==>
    6.31                    inj_on f B --> setsum h (f ` B) = setsum (h \<circ> f) B"
    6.32 @@ -820,7 +820,7 @@
    6.33      ==> setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
    6.34    -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
    6.35    apply (induct set: Finites, simp)
    6.36 -  apply (simp add: plus_ac0 Int_insert_left insert_absorb)
    6.37 +  apply (simp add: add_ac Int_insert_left insert_absorb)
    6.38    done
    6.39  
    6.40  lemma setsum_Un_disjoint: "finite A ==> finite B
    6.41 @@ -874,7 +874,7 @@
    6.42    apply (case_tac "finite A")
    6.43     prefer 2 apply (simp add: setsum_def)
    6.44    apply (erule finite_induct, auto)
    6.45 -  apply (simp add: plus_ac0)
    6.46 +  apply (simp add: add_ac)
    6.47    done
    6.48  
    6.49  subsubsection {* Properties in more restricted classes of structures *}
    6.50 @@ -892,18 +892,18 @@
    6.51  
    6.52  lemma setsum_constant_nat [simp]:
    6.53      "finite A ==> (\<Sum>x: A. y) = (card A) * y"
    6.54 -  -- {* Later generalized to any semiring. *}
    6.55 +  -- {* Later generalized to any comm_semiring_1_cancel. *}
    6.56    by (erule finite_induct, auto)
    6.57  
    6.58  lemma setsum_Un: "finite A ==> finite B ==>
    6.59      (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
    6.60    -- {* For the natural numbers, we have subtraction. *}
    6.61 -  by (subst setsum_Un_Int [symmetric], auto)
    6.62 +  by (subst setsum_Un_Int [symmetric], auto simp add: ring_eq_simps)
    6.63  
    6.64  lemma setsum_Un_ring: "finite A ==> finite B ==>
    6.65 -    (setsum f (A Un B) :: 'a :: ring) =
    6.66 +    (setsum f (A Un B) :: 'a :: comm_ring_1) =
    6.67        setsum f A + setsum f B - setsum f (A Int B)"
    6.68 -  by (subst setsum_Un_Int [symmetric], auto)
    6.69 +  by (subst setsum_Un_Int [symmetric], auto simp add: ring_eq_simps)
    6.70  
    6.71  lemma setsum_diff1: "(setsum f (A - {a}) :: nat) =
    6.72      (if a:A then setsum f A - f a else setsum f A)"
    6.73 @@ -914,16 +914,16 @@
    6.74    apply (drule_tac a = a in mk_disjoint_insert, auto)
    6.75    done
    6.76  
    6.77 -lemma setsum_negf: "finite A ==> setsum (%x. - (f x)::'a::ring) A =
    6.78 +lemma setsum_negf: "finite A ==> setsum (%x. - (f x)::'a::comm_ring_1) A =
    6.79    - setsum f A"
    6.80    by (induct set: Finites, auto)
    6.81  
    6.82 -lemma setsum_subtractf: "finite A ==> setsum (%x. ((f x)::'a::ring) - g x) A =
    6.83 +lemma setsum_subtractf: "finite A ==> setsum (%x. ((f x)::'a::comm_ring_1) - g x) A =
    6.84    setsum f A - setsum g A"
    6.85    by (simp add: diff_minus setsum_addf setsum_negf)
    6.86  
    6.87  lemma setsum_nonneg: "[| finite A;
    6.88 -    \<forall>x \<in> A. (0::'a::ordered_semiring) \<le> f x |] ==>
    6.89 +    \<forall>x \<in> A. (0::'a::ordered_semidom) \<le> f x |] ==>
    6.90      0 \<le>  setsum f A";
    6.91    apply (induct set: Finites, auto)
    6.92    apply (subgoal_tac "0 + 0 \<le> f x + setsum f F", simp)
    6.93 @@ -983,16 +983,16 @@
    6.94  subsection {* Generalized product over a set *}
    6.95  
    6.96  constdefs
    6.97 -  setprod :: "('a => 'b) => 'a set => 'b::times_ac1"
    6.98 +  setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult"
    6.99    "setprod f A == if finite A then fold (op * o f) 1 A else 1"
   6.100  
   6.101  syntax
   6.102 -  "_setprod" :: "idt => 'a set => 'b => 'b::times_ac1"  ("(3\<Prod>_:_. _)" [0, 51, 10] 10)
   6.103 +  "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_:_. _)" [0, 51, 10] 10)
   6.104  
   6.105  syntax (xsymbols)
   6.106 -  "_setprod" :: "idt => 'a set => 'b => 'b::times_ac1"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
   6.107 +  "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
   6.108  syntax (HTML output)
   6.109 -  "_setprod" :: "idt => 'a set => 'b => 'b::times_ac1"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
   6.110 +  "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
   6.111  translations
   6.112    "\<Prod>i:A. b" == "setprod (%i. b) A"  -- {* Beware of argument permutation! *}
   6.113  
   6.114 @@ -1002,7 +1002,7 @@
   6.115  lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
   6.116      setprod f (insert a A) = f a * setprod f A"
   6.117    by (auto simp add: setprod_def LC_def LC.fold_insert
   6.118 -      times_ac1_left_commute)
   6.119 +      mult_left_commute)
   6.120  
   6.121  lemma setprod_reindex [rule_format]: "finite B ==>
   6.122                    inj_on f B --> setprod h (f ` B) = setprod (h \<circ> f) B"
   6.123 @@ -1043,7 +1043,7 @@
   6.124  
   6.125  lemma setprod_1: "setprod (%i. 1) A = 1"
   6.126    apply (case_tac "finite A")
   6.127 -  apply (erule finite_induct, auto simp add: times_ac1)
   6.128 +  apply (erule finite_induct, auto simp add: mult_ac)
   6.129    apply (simp add: setprod_def)
   6.130    done
   6.131  
   6.132 @@ -1056,13 +1056,13 @@
   6.133  lemma setprod_Un_Int: "finite A ==> finite B
   6.134      ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
   6.135    apply (induct set: Finites, simp)
   6.136 -  apply (simp add: times_ac1 insert_absorb)
   6.137 -  apply (simp add: times_ac1 Int_insert_left insert_absorb)
   6.138 +  apply (simp add: mult_ac insert_absorb)
   6.139 +  apply (simp add: mult_ac Int_insert_left insert_absorb)
   6.140    done
   6.141  
   6.142  lemma setprod_Un_disjoint: "finite A ==> finite B
   6.143    ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
   6.144 -  apply (subst setprod_Un_Int [symmetric], auto simp add: times_ac1)
   6.145 +  apply (subst setprod_Un_Int [symmetric], auto simp add: mult_ac)
   6.146    done
   6.147  
   6.148  lemma setprod_UN_disjoint:
   6.149 @@ -1110,9 +1110,9 @@
   6.150  lemma setprod_timesf: "setprod (%x. f x * g x) A =
   6.151      (setprod f A * setprod g A)"
   6.152    apply (case_tac "finite A")
   6.153 -   prefer 2 apply (simp add: setprod_def times_ac1)
   6.154 +   prefer 2 apply (simp add: setprod_def mult_ac)
   6.155    apply (erule finite_induct, auto)
   6.156 -  apply (simp add: times_ac1)
   6.157 +  apply (simp add: mult_ac)
   6.158    done
   6.159  
   6.160  subsubsection {* Properties in more restricted classes of structures *}
   6.161 @@ -1127,13 +1127,13 @@
   6.162    apply (auto simp add: power_Suc)
   6.163    done
   6.164  
   6.165 -lemma setprod_zero: "finite A ==> EX x: A. f x = (0::'a::semiring) ==>
   6.166 +lemma setprod_zero: "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1_cancel) ==>
   6.167      setprod f A = 0"
   6.168    apply (induct set: Finites, force, clarsimp)
   6.169    apply (erule disjE, auto)
   6.170    done
   6.171  
   6.172 -lemma setprod_nonneg [rule_format]: "(ALL x: A. (0::'a::ordered_ring) \<le> f x)
   6.173 +lemma setprod_nonneg [rule_format]: "(ALL x: A. (0::'a::ordered_idom) \<le> f x)
   6.174       --> 0 \<le> setprod f A"
   6.175    apply (case_tac "finite A")
   6.176    apply (induct set: Finites, force, clarsimp)
   6.177 @@ -1142,7 +1142,7 @@
   6.178    apply (auto simp add: setprod_def)
   6.179    done
   6.180  
   6.181 -lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_ring) < f x)
   6.182 +lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_idom) < f x)
   6.183       --> 0 < setprod f A"
   6.184    apply (case_tac "finite A")
   6.185    apply (induct set: Finites, force, clarsimp)
   6.186 @@ -1152,13 +1152,13 @@
   6.187    done
   6.188  
   6.189  lemma setprod_nonzero [rule_format]:
   6.190 -    "(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==>
   6.191 +    "(ALL x y. (x::'a::comm_semiring_1_cancel) * y = 0 --> x = 0 | y = 0) ==>
   6.192        finite A ==> (ALL x: A. f x \<noteq> (0::'a)) --> setprod f A \<noteq> 0"
   6.193    apply (erule finite_induct, auto)
   6.194    done
   6.195  
   6.196  lemma setprod_zero_eq:
   6.197 -    "(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==>
   6.198 +    "(ALL x y. (x::'a::comm_semiring_1_cancel) * y = 0 --> x = 0 | y = 0) ==>
   6.199       finite A ==> (setprod f A = (0::'a)) = (EX x: A. f x = 0)"
   6.200    apply (insert setprod_zero [of A f] setprod_nonzero [of A f], blast)
   6.201    done
     7.1 --- a/src/HOL/Hyperreal/HyperDef.thy	Tue May 11 14:00:02 2004 +0200
     7.2 +++ b/src/HOL/Hyperreal/HyperDef.thy	Tue May 11 20:11:08 2004 +0200
     7.3 @@ -332,7 +332,7 @@
     7.4  lemma hypreal_add_zero_left: "(0::hypreal) + z = z"
     7.5  by (cases z, simp add: hypreal_zero_def hypreal_add)
     7.6  
     7.7 -instance hypreal :: plus_ac0
     7.8 +instance hypreal :: comm_monoid_add
     7.9    by intro_classes
    7.10      (assumption | 
    7.11        rule hypreal_add_commute hypreal_add_assoc hypreal_add_zero_left)+
    7.12 @@ -423,9 +423,6 @@
    7.13  instance hypreal :: field
    7.14  proof
    7.15    fix x y z :: hypreal
    7.16 -  show "(x + y) + z = x + (y + z)" by (rule hypreal_add_assoc)
    7.17 -  show "x + y = y + x" by (rule hypreal_add_commute)
    7.18 -  show "0 + x = x" by simp
    7.19    show "- x + x = 0" by (simp add: hypreal_add_minus_left)
    7.20    show "x - y = x + (-y)" by (simp add: hypreal_diff_def)
    7.21    show "(x * y) * z = x * (y * z)" by (rule hypreal_mult_assoc)
    7.22 @@ -512,7 +509,7 @@
    7.23  
    7.24  lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)"
    7.25  apply auto
    7.26 -apply (rule Ring_and_Field.add_right_cancel [of _ "-y", THEN iffD1], auto)
    7.27 +apply (rule OrderedGroup.add_right_cancel [of _ "-y", THEN iffD1], auto)
    7.28  done
    7.29  
    7.30  lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
     8.1 --- a/src/HOL/Hyperreal/HyperNat.thy	Tue May 11 14:00:02 2004 +0200
     8.2 +++ b/src/HOL/Hyperreal/HyperNat.thy	Tue May 11 20:11:08 2004 +0200
     8.3 @@ -158,7 +158,7 @@
     8.4  apply (simp add: hypnat_zero_def hypnat_add)
     8.5  done
     8.6  
     8.7 -instance hypnat :: plus_ac0
     8.8 +instance hypnat :: comm_monoid_add
     8.9    by intro_classes
    8.10      (assumption |
    8.11        rule hypnat_add_commute hypnat_add_assoc hypnat_add_zero_left)+
    8.12 @@ -280,13 +280,10 @@
    8.13  declare hypnat_zero_not_eq_one [THEN not_sym, simp]
    8.14  
    8.15  
    8.16 -text{*The Hypernaturals Form A Semiring*}
    8.17 -instance hypnat :: semiring
    8.18 +text{*The Hypernaturals Form A comm_semiring_1_cancel*}
    8.19 +instance hypnat :: comm_semiring_1_cancel
    8.20  proof
    8.21    fix i j k :: hypnat
    8.22 -  show "(i + j) + k = i + (j + k)" by (rule hypnat_add_assoc)
    8.23 -  show "i + j = j + i" by (rule hypnat_add_commute)
    8.24 -  show "0 + i = i" by simp
    8.25    show "(i * j) * k = i * (j * k)" by (rule hypnat_mult_assoc)
    8.26    show "i * j = j * i" by (rule hypnat_mult_commute)
    8.27    show "1 * i = i" by (rule hypnat_mult_1)
    8.28 @@ -352,9 +349,9 @@
    8.29  done
    8.30  
    8.31  
    8.32 -subsection{*The Hypernaturals Form an Ordered Semiring*}
    8.33 +subsection{*The Hypernaturals Form an Ordered comm_semiring_1_cancel*}
    8.34  
    8.35 -instance hypnat :: ordered_semiring
    8.36 +instance hypnat :: ordered_semidom
    8.37  proof
    8.38    fix x y z :: hypnat
    8.39    show "0 < (1::hypnat)"
    8.40 @@ -456,7 +453,7 @@
    8.41  qed
    8.42  
    8.43  
    8.44 -subsection{*The Embedding @{term hypnat_of_nat} Preserves Ring and 
    8.45 +subsection{*The Embedding @{term hypnat_of_nat} Preserves comm_ring_1 and 
    8.46        Order Properties*}
    8.47  
    8.48  constdefs
     9.1 --- a/src/HOL/Hyperreal/MacLaurin.ML	Tue May 11 14:00:02 2004 +0200
     9.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     9.3 @@ -1,717 +0,0 @@
     9.4 -(*  Title       : MacLaurin.thy
     9.5 -    Author      : Jacques D. Fleuriot
     9.6 -    Copyright   : 2001 University of Edinburgh
     9.7 -    Description : MacLaurin series
     9.8 -*)
     9.9 -
    9.10 -Goal "sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
    9.11 -by (induct_tac "n" 1);
    9.12 -by Auto_tac;
    9.13 -qed "sumr_offset";
    9.14 -
    9.15 -Goal "ALL f. sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
    9.16 -by (induct_tac "n" 1);
    9.17 -by Auto_tac;
    9.18 -qed "sumr_offset2";
    9.19 -
    9.20 -Goal "sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
    9.21 -by (simp_tac (simpset() addsimps [sumr_offset]) 1);
    9.22 -qed "sumr_offset3";
    9.23 -
    9.24 -Goal "ALL n f. sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
    9.25 -by (simp_tac (simpset() addsimps [sumr_offset]) 1);
    9.26 -qed "sumr_offset4";
    9.27 -
    9.28 -Goal "0 < n ==> \
    9.29 -\     sumr (Suc 0) (Suc n) (%n. (if even(n) then 0 else \
    9.30 -\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n) = \
    9.31 -\     sumr 0 (Suc n) (%n. (if even(n) then 0 else \
    9.32 -\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n)";
    9.33 -by (res_inst_tac [("n1","1")] (sumr_split_add RS subst) 1);
    9.34 -by Auto_tac;
    9.35 -qed "sumr_from_1_from_0";
    9.36 -
    9.37 -(*---------------------------------------------------------------------------*)
    9.38 -(* Maclaurin's theorem with Lagrange form of remainder                       *)
    9.39 -(*---------------------------------------------------------------------------*)
    9.40 -
    9.41 -(* Annoying: Proof is now even longer due mostly to 
    9.42 -   change in behaviour of simplifier  since Isabelle99 *)
    9.43 -Goal " [| 0 < h; 0 < n; diff 0 = f; \
    9.44 -\      ALL m t. \
    9.45 -\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
    9.46 -\   ==> EX t. 0 < t & \
    9.47 -\             t < h & \
    9.48 -\             f h = \
    9.49 -\             sumr 0 n (%m. (diff m 0 / real (fact m)) * h ^ m) + \
    9.50 -\             (diff n t / real (fact n)) * h ^ n";
    9.51 -by (case_tac "n = 0" 1);
    9.52 -by (Force_tac 1);
    9.53 -by (dtac not0_implies_Suc 1);
    9.54 -by (etac exE 1);
    9.55 -by (subgoal_tac 
    9.56 -     "EX B. f h = sumr 0 n (%m. (diff m 0 / real (fact m)) * (h ^ m)) \
    9.57 -\                  + (B * ((h ^ n) / real (fact n)))" 1);
    9.58 -
    9.59 -by (simp_tac (HOL_ss addsimps [real_add_commute, real_divide_def,
    9.60 -    ARITH_PROVE "(x = z + (y::real)) = (x - y = z)"]) 2);
    9.61 -by (res_inst_tac 
    9.62 -  [("x","(f(h) - sumr 0 n (%m. (diff(m)(0) / real (fact m)) * (h ^ m))) \
    9.63 -\        * real (fact n) / (h ^ n)")] exI 2);
    9.64 -by (simp_tac (HOL_ss addsimps [real_mult_assoc,real_divide_def]) 2);
    9.65 - by (rtac (CLAIM "x = (1::real) ==>  a = a * (x::real)") 2);
    9.66 -by (asm_simp_tac (HOL_ss addsimps 
    9.67 -    [CLAIM "(a::real) * (b * (c * d)) = (d * a) * (b * c)"]
    9.68 -     delsimps [realpow_Suc]) 2);
    9.69 -by (stac left_inverse 2);
    9.70 -by (stac left_inverse 3);
    9.71 -by (rtac (real_not_refl2 RS not_sym) 2);
    9.72 -by (etac zero_less_power 2);
    9.73 -by (rtac real_of_nat_fact_not_zero 2);
    9.74 -by (Simp_tac 2);
    9.75 -by (etac exE 1);
    9.76 -by (cut_inst_tac [("b","%t. f t - \
    9.77 -\      (sumr 0 n (%m. (diff m 0 / real (fact m)) * (t ^ m)) + \
    9.78 -\                       (B * ((t ^ n) / real (fact n))))")] 
    9.79 -    (CLAIM "EX g. g = b") 1);
    9.80 -by (etac exE 1);
    9.81 -by (subgoal_tac "g 0 = 0 & g h =0" 1);
    9.82 -by (asm_simp_tac (simpset() addsimps 
    9.83 -    [ARITH_PROVE "(x - y = z) = (x = z + (y::real))"]
    9.84 -    delsimps [sumr_Suc]) 2);
    9.85 -by (cut_inst_tac [("n","m"),("k","1")] sumr_offset2 2);
    9.86 -by (asm_full_simp_tac (simpset() addsimps 
    9.87 -    [ARITH_PROVE "(x = y - z) = (y = x + (z::real))"]
    9.88 -    delsimps [sumr_Suc]) 2);
    9.89 -by (cut_inst_tac [("b","%m t. diff m t - \
    9.90 -\      (sumr 0 (n - m) (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) \
    9.91 -\       + (B * ((t ^ (n - m)) / real (fact(n - m)))))")] 
    9.92 -    (CLAIM "EX difg. difg = b") 1);
    9.93 -by (etac exE 1);
    9.94 -by (subgoal_tac "difg 0 = g" 1);
    9.95 -by (asm_simp_tac (simpset() delsimps [realpow_Suc,fact_Suc]) 2);
    9.96 -by (subgoal_tac "ALL m t. m < n & 0 <= t & t <= h --> \
    9.97 -\                   DERIV (difg m) t :> difg (Suc m) t" 1);
    9.98 -by (Clarify_tac 2);
    9.99 -by (rtac DERIV_diff 2);
   9.100 -by (Asm_simp_tac 2);
   9.101 -by DERIV_tac;
   9.102 -by DERIV_tac;
   9.103 -by (rtac lemma_DERIV_subst 3);
   9.104 -by (rtac DERIV_quotient 3);
   9.105 -by (rtac DERIV_const 4);
   9.106 -by (rtac DERIV_pow 3);
   9.107 -by (asm_simp_tac (simpset() addsimps [inverse_mult_distrib,
   9.108 -    CLAIM_SIMP "(a::real) * b * c * (d * e) = a * b * (c * d) * e" 
   9.109 -    mult_ac,fact_diff_Suc]) 4);
   9.110 -by (Asm_simp_tac 3);
   9.111 -by (forw_inst_tac [("m","ma")] less_add_one 2);
   9.112 -by (Clarify_tac 2);
   9.113 -by (asm_simp_tac (simpset() addsimps 
   9.114 -    [CLAIM "Suc m = ma + d + 1 ==> m - ma = d"]
   9.115 -    delsimps [sumr_Suc]) 2);
   9.116 -by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
   9.117 -          (read_instantiate [("k","1")] sumr_offset4))] 
   9.118 -    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
   9.119 -by (rtac lemma_DERIV_subst 2);
   9.120 -by (rtac DERIV_add 2);
   9.121 -by (rtac DERIV_const 3);
   9.122 -by (rtac DERIV_sumr 2);
   9.123 -by (Clarify_tac 2);
   9.124 -by (Simp_tac 3);
   9.125 -by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc] 
   9.126 -    delsimps [fact_Suc,realpow_Suc]) 2);
   9.127 -by (rtac DERIV_cmult 2);
   9.128 -by (rtac lemma_DERIV_subst 2);
   9.129 -by DERIV_tac;
   9.130 -by (stac fact_Suc 2);
   9.131 -by (stac real_of_nat_mult 2);
   9.132 -by (simp_tac (simpset() addsimps [inverse_mult_distrib] @
   9.133 -    mult_ac) 2);
   9.134 -by (subgoal_tac "ALL ma. ma < n --> \
   9.135 -\        (EX t. 0 < t & t < h & difg (Suc ma) t = 0)" 1);
   9.136 -by (rotate_tac 11 1);
   9.137 -by (dres_inst_tac [("x","m")] spec 1);
   9.138 -by (etac impE 1);
   9.139 -by (Asm_simp_tac 1);
   9.140 -by (etac exE 1);
   9.141 -by (res_inst_tac [("x","t")] exI 1);
   9.142 -by (asm_full_simp_tac (simpset() addsimps 
   9.143 -     [ARITH_PROVE "(x - y = 0) = (y = (x::real))"] 
   9.144 -      delsimps [realpow_Suc,fact_Suc]) 1);
   9.145 -by (subgoal_tac "ALL m. m < n --> difg m 0 = 0" 1);
   9.146 -by (Clarify_tac 2);
   9.147 -by (Asm_simp_tac 2);
   9.148 -by (forw_inst_tac [("m","ma")] less_add_one 2);
   9.149 -by (Clarify_tac 2);
   9.150 -by (asm_simp_tac (simpset() delsimps [sumr_Suc]) 2);
   9.151 -by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
   9.152 -          (read_instantiate [("k","1")] sumr_offset4))] 
   9.153 -    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
   9.154 -by (subgoal_tac "ALL m. m < n --> (EX t. 0 < t & t < h & \
   9.155 -\                DERIV (difg m) t :> 0)" 1);
   9.156 -by (rtac allI 1 THEN rtac impI 1);
   9.157 -by (rotate_tac 12 1);
   9.158 -by (dres_inst_tac [("x","ma")] spec 1);
   9.159 -by (etac impE 1 THEN assume_tac 1);
   9.160 -by (etac exE 1);
   9.161 -by (res_inst_tac [("x","t")] exI 1);
   9.162 -(* do some tidying up *)
   9.163 -by (ALLGOALS(thin_tac "difg = \
   9.164 -\          (%m t. diff m t - \
   9.165 -\                 (sumr 0 (n - m) \
   9.166 -\                   (%p. diff (m + p) 0 / real (fact p) * t ^ p) + \
   9.167 -\                  B * (t ^ (n - m) / real (fact (n - m)))))"));
   9.168 -by (ALLGOALS(thin_tac "g = \
   9.169 -\          (%t. f t - \
   9.170 -\               (sumr 0 n (%m. diff m 0 / real  (fact m) * t ^ m) + \
   9.171 -\                B * (t ^ n / real (fact n))))"));
   9.172 -by (ALLGOALS(thin_tac "f h = \
   9.173 -\          sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.174 -\          B * (h ^ n / real (fact n))"));
   9.175 -(* back to business *)
   9.176 -by (Asm_simp_tac 1);
   9.177 -by (rtac DERIV_unique 1);
   9.178 -by (Blast_tac 2);
   9.179 -by (Force_tac 1);
   9.180 -by (rtac allI 1 THEN induct_tac "ma" 1);
   9.181 -by (rtac impI 1 THEN rtac Rolle 1);
   9.182 -by (assume_tac 1);
   9.183 -by (Asm_full_simp_tac 1);
   9.184 -by (Asm_full_simp_tac 1);
   9.185 -by (subgoal_tac "ALL t. 0 <= t & t <= h --> g differentiable t" 1);
   9.186 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
   9.187 -by (blast_tac (claset() addDs [DERIV_isCont]) 1);
   9.188 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
   9.189 -by (Clarify_tac 1);
   9.190 -by (res_inst_tac [("x","difg (Suc 0) t")] exI 1);
   9.191 -by (Force_tac 1);
   9.192 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
   9.193 -by (Clarify_tac 1);
   9.194 -by (res_inst_tac [("x","difg (Suc 0) x")] exI 1);
   9.195 -by (Force_tac 1);
   9.196 -by (Step_tac 1);
   9.197 -by (Force_tac 1);
   9.198 -by (subgoal_tac "EX ta. 0 < ta & ta < t & \
   9.199 -\                DERIV difg (Suc n) ta :> 0" 1);
   9.200 -by (rtac Rolle 2 THEN assume_tac 2);
   9.201 -by (Asm_full_simp_tac 2);
   9.202 -by (rotate_tac 2 2);
   9.203 -by (dres_inst_tac [("x","n")] spec 2);
   9.204 -by (ftac (ARITH_PROVE "n < m  ==> n < Suc m") 2);
   9.205 -by (rtac DERIV_unique 2);
   9.206 -by (assume_tac 3);
   9.207 -by (Force_tac 2);
   9.208 -by (subgoal_tac 
   9.209 -    "ALL ta. 0 <= ta & ta <= t --> (difg (Suc n)) differentiable ta" 2);
   9.210 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
   9.211 -by (blast_tac (claset() addSDs [DERIV_isCont]) 2);
   9.212 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
   9.213 -by (Clarify_tac 2);
   9.214 -by (res_inst_tac [("x","difg (Suc (Suc n)) ta")] exI 2);
   9.215 -by (Force_tac 2);
   9.216 -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
   9.217 -by (Clarify_tac 2);
   9.218 -by (res_inst_tac [("x","difg (Suc (Suc n)) x")] exI 2);
   9.219 -by (Force_tac 2);
   9.220 -by (Step_tac 1);
   9.221 -by (res_inst_tac [("x","ta")] exI 1);
   9.222 -by (Force_tac 1);
   9.223 -qed "Maclaurin";
   9.224 -
   9.225 -Goal "0 < h & 0 < n & diff 0 = f & \
   9.226 -\      (ALL m t. \
   9.227 -\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
   9.228 -\   --> (EX t. 0 < t & \
   9.229 -\             t < h & \
   9.230 -\             f h = \
   9.231 -\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.232 -\             diff n t / real (fact n) * h ^ n)";
   9.233 -by (blast_tac (claset() addIs [Maclaurin]) 1);
   9.234 -qed "Maclaurin_objl";
   9.235 -
   9.236 -Goal " [| 0 < h; diff 0 = f; \
   9.237 -\      ALL m t. \
   9.238 -\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
   9.239 -\   ==> EX t. 0 < t & \
   9.240 -\             t <= h & \
   9.241 -\             f h = \
   9.242 -\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.243 -\             diff n t / real (fact n) * h ^ n";
   9.244 -by (case_tac "n" 1);
   9.245 -by Auto_tac;
   9.246 -by (dtac Maclaurin 1 THEN Auto_tac);
   9.247 -qed "Maclaurin2";
   9.248 -
   9.249 -Goal "0 < h & diff 0 = f & \
   9.250 -\      (ALL m t. \
   9.251 -\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
   9.252 -\   --> (EX t. 0 < t & \
   9.253 -\             t <= h & \
   9.254 -\             f h = \
   9.255 -\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.256 -\             diff n t / real (fact n) * h ^ n)";
   9.257 -by (blast_tac (claset() addIs [Maclaurin2]) 1);
   9.258 -qed "Maclaurin2_objl";
   9.259 -
   9.260 -Goal " [| h < 0; 0 < n; diff 0 = f; \
   9.261 -\      ALL m t. \
   9.262 -\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t |] \
   9.263 -\   ==> EX t. h < t & \
   9.264 -\             t < 0 & \
   9.265 -\             f h = \
   9.266 -\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.267 -\             diff n t / real (fact n) * h ^ n";
   9.268 -by (cut_inst_tac [("f","%x. f (-x)"),
   9.269 -                 ("diff","%n x. ((- 1) ^ n) * diff n (-x)"),
   9.270 -                 ("h","-h"),("n","n")] Maclaurin_objl 1);
   9.271 -by (Asm_full_simp_tac 1);
   9.272 -by (etac impE 1 THEN Step_tac 1);
   9.273 -by (stac minus_mult_right 1);
   9.274 -by (rtac DERIV_cmult 1);
   9.275 -by (rtac lemma_DERIV_subst 1);
   9.276 -by (rtac (read_instantiate [("g","uminus")] DERIV_chain2) 1);
   9.277 -by (rtac DERIV_minus 2 THEN rtac DERIV_Id 2);
   9.278 -by (Force_tac 2);
   9.279 -by (Force_tac 1);
   9.280 -by (res_inst_tac [("x","-t")] exI 1);
   9.281 -by Auto_tac;
   9.282 -by (rtac (CLAIM "[| x = x'; y = y' |] ==> x + y = x' + (y'::real)") 1);
   9.283 -by (rtac sumr_fun_eq 1);
   9.284 -by (Asm_full_simp_tac 1);
   9.285 -by (auto_tac (claset(),simpset() addsimps [real_divide_def,
   9.286 -    CLAIM "((a * b) * c) * d = (b * c) * (a * (d::real))",
   9.287 -    power_mult_distrib RS sym]));
   9.288 -qed "Maclaurin_minus";
   9.289 -
   9.290 -Goal "(h < 0 & 0 < n & diff 0 = f & \
   9.291 -\      (ALL m t. \
   9.292 -\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t))\
   9.293 -\   --> (EX t. h < t & \
   9.294 -\             t < 0 & \
   9.295 -\             f h = \
   9.296 -\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
   9.297 -\             diff n t / real (fact n) * h ^ n)";
   9.298 -by (blast_tac (claset() addIs [Maclaurin_minus]) 1);
   9.299 -qed "Maclaurin_minus_objl";
   9.300 -
   9.301 -(* ------------------------------------------------------------------------- *)
   9.302 -(* More convenient "bidirectional" version.                                  *)
   9.303 -(* ------------------------------------------------------------------------- *)
   9.304 -
   9.305 -(* not good for PVS sin_approx, cos_approx *)
   9.306 -Goal " [| diff 0 = f; \
   9.307 -\      ALL m t. \
   9.308 -\         m < n & abs t <= abs x --> DERIV (diff m) t :> diff (Suc m) t |] \
   9.309 -\   ==> EX t. abs t <= abs x & \
   9.310 -\             f x = \
   9.311 -\             sumr 0 n (%m. diff m 0 / real (fact m) * x ^ m) + \
   9.312 -\             diff n t / real (fact n) * x ^ n";
   9.313 -by (case_tac "n = 0" 1);
   9.314 -by (Force_tac 1);
   9.315 -by (case_tac "x = 0" 1);
   9.316 -by (res_inst_tac [("x","0")] exI 1);
   9.317 -by (Asm_full_simp_tac 1);
   9.318 -by (res_inst_tac [("P","0 < n")] impE 1);
   9.319 -by (assume_tac 2 THEN assume_tac 2);
   9.320 -by (induct_tac "n" 1);
   9.321 -by (Simp_tac 1);
   9.322 -by Auto_tac;
   9.323 -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
   9.324 -by Auto_tac;
   9.325 -by (cut_inst_tac [("f","diff 0"),
   9.326 -                 ("diff","diff"),
   9.327 -                 ("h","x"),("n","n")] Maclaurin_objl 2);
   9.328 -by (Step_tac 2);
   9.329 -by (blast_tac (claset() addDs 
   9.330 -    [ARITH_PROVE "[|(0::real) <= t;t <= x |] ==> abs t <= abs x"]) 2);
   9.331 -by (res_inst_tac [("x","t")] exI 2);
   9.332 -by (force_tac (claset() addIs 
   9.333 -    [ARITH_PROVE "[| 0 < t; (t::real) < x|] ==> abs t <= abs x"],simpset()) 2);
   9.334 -by (cut_inst_tac [("f","diff 0"),
   9.335 -                 ("diff","diff"),
   9.336 -                 ("h","x"),("n","n")] Maclaurin_minus_objl 1);
   9.337 -by (Step_tac 1);
   9.338 -by (blast_tac (claset() addDs 
   9.339 -    [ARITH_PROVE "[|x <= t;t <= (0::real) |] ==> abs t <= abs x"]) 1);
   9.340 -by (res_inst_tac [("x","t")] exI 1);
   9.341 -by (force_tac (claset() addIs 
   9.342 -    [ARITH_PROVE "[| x < t; (t::real) < 0|] ==> abs t <= abs x"],simpset()) 1);
   9.343 -qed "Maclaurin_bi_le";
   9.344 -
   9.345 -Goal "[| diff 0 = f; \
   9.346 -\        ALL m x. DERIV (diff m) x :> diff(Suc m) x; \ 
   9.347 -\       x ~= 0; 0 < n \
   9.348 -\     |] ==> EX t. 0 < abs t & abs t < abs x & \
   9.349 -\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
   9.350 -\                    (diff n t / real (fact n)) * x ^ n";
   9.351 -by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1);
   9.352 -by (Blast_tac 2);
   9.353 -by (dtac Maclaurin_minus 1);
   9.354 -by (dtac Maclaurin 5);
   9.355 -by (TRYALL(assume_tac));
   9.356 -by (Blast_tac 1);
   9.357 -by (Blast_tac 2);
   9.358 -by (Step_tac 1);
   9.359 -by (ALLGOALS(res_inst_tac [("x","t")] exI));
   9.360 -by (Step_tac 1);
   9.361 -by (ALLGOALS(arith_tac));
   9.362 -qed "Maclaurin_all_lt";
   9.363 -
   9.364 -Goal "diff 0 = f & \
   9.365 -\     (ALL m x. DERIV (diff m) x :> diff(Suc m) x) & \
   9.366 -\     x ~= 0 & 0 < n \
   9.367 -\     --> (EX t. 0 < abs t & abs t < abs x & \
   9.368 -\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
   9.369 -\                    (diff n t / real (fact n)) * x ^ n)";
   9.370 -by (blast_tac (claset() addIs [Maclaurin_all_lt]) 1);
   9.371 -qed "Maclaurin_all_lt_objl";
   9.372 -
   9.373 -Goal "x = (0::real)  \
   9.374 -\     ==> 0 < n --> \
   9.375 -\         sumr 0 n (%m. (diff m (0::real) / real (fact m)) * x ^ m) = \
   9.376 -\         diff 0 0";
   9.377 -by (Asm_simp_tac 1);
   9.378 -by (induct_tac "n" 1);
   9.379 -by Auto_tac; 
   9.380 -qed_spec_mp "Maclaurin_zero";
   9.381 -
   9.382 -Goal "[| diff 0 = f; \
   9.383 -\       ALL m x. DERIV (diff m) x :> diff (Suc m) x \
   9.384 -\     |] ==> EX t. abs t <= abs x & \
   9.385 -\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
   9.386 -\                   (diff n t / real (fact n)) * x ^ n";
   9.387 -by (cut_inst_tac [("n","n"),("m","0")] 
   9.388 -       (ARITH_PROVE "n <= m | m < (n::nat)") 1);
   9.389 -by (etac disjE 1);
   9.390 -by (Force_tac 1);
   9.391 -by (case_tac "x = 0" 1);
   9.392 -by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_zero 1);
   9.393 -by (assume_tac 1);
   9.394 -by (dtac (gr_implies_not0 RS  not0_implies_Suc) 1);
   9.395 -by (res_inst_tac [("x","0")] exI 1);
   9.396 -by (Force_tac 1);
   9.397 -by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_all_lt 1);
   9.398 -by (TRYALL(assume_tac));
   9.399 -by (Step_tac 1);
   9.400 -by (res_inst_tac [("x","t")] exI 1);
   9.401 -by Auto_tac;
   9.402 -qed "Maclaurin_all_le";
   9.403 -
   9.404 -Goal "diff 0 = f & \
   9.405 -\     (ALL m x. DERIV (diff m) x :> diff (Suc m) x)  \
   9.406 -\     --> (EX t. abs t <= abs x & \
   9.407 -\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
   9.408 -\                   (diff n t / real (fact n)) * x ^ n)";
   9.409 -by (blast_tac (claset() addIs [Maclaurin_all_le]) 1);
   9.410 -qed "Maclaurin_all_le_objl";
   9.411 -
   9.412 -(* ------------------------------------------------------------------------- *)
   9.413 -(* Version for exp.                                                          *)
   9.414 -(* ------------------------------------------------------------------------- *)
   9.415 -
   9.416 -Goal "[| x ~= 0; 0 < n |] \
   9.417 -\     ==> (EX t. 0 < abs t & \
   9.418 -\               abs t < abs x & \
   9.419 -\               exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
   9.420 -\                       (exp t / real (fact n)) * x ^ n)";
   9.421 -by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
   9.422 -    Maclaurin_all_lt_objl 1);
   9.423 -by Auto_tac;
   9.424 -qed "Maclaurin_exp_lt";
   9.425 -
   9.426 -Goal "EX t. abs t <= abs x & \
   9.427 -\           exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
   9.428 -\                      (exp t / real (fact n)) * x ^ n";
   9.429 -by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
   9.430 -    Maclaurin_all_le_objl 1);
   9.431 -by Auto_tac;
   9.432 -qed "Maclaurin_exp_le";
   9.433 -
   9.434 -(* ------------------------------------------------------------------------- *)
   9.435 -(* Version for sin function                                                  *)
   9.436 -(* ------------------------------------------------------------------------- *)
   9.437 -
   9.438 -Goal "[| a < b; ALL x. a <= x & x <= b --> DERIV f x :> f'(x) |] \
   9.439 -\     ==> EX z. a < z & z < b & (f b - f a = (b - a) * f'(z))";
   9.440 -by (dtac MVT 1);
   9.441 -by (blast_tac (claset() addIs [DERIV_isCont]) 1);
   9.442 -by (force_tac (claset() addDs [order_less_imp_le],
   9.443 -    simpset() addsimps [differentiable_def]) 1);
   9.444 -by (blast_tac (claset() addDs [DERIV_unique,order_less_imp_le]) 1);
   9.445 -qed "MVT2";
   9.446 -
   9.447 -Goal "d < (4::nat) ==> d = 0 | d = 1 | d = 2 | d = 3";
   9.448 -by (case_tac "d" 1 THEN Auto_tac);
   9.449 -qed "lemma_exhaust_less_4";
   9.450 -
   9.451 -bind_thm ("real_mult_le_lemma",
   9.452 -          simplify (simpset()) (inst "b" "1" mult_right_mono));
   9.453 -
   9.454 -
   9.455 -Goal "abs(sin x - \
   9.456 -\          sumr 0 n (%m. (if even m then 0 \
   9.457 -\                         else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
   9.458 -\                         x ^ m)) \
   9.459 -\      <= inverse(real (fact n)) * abs(x) ^ n";
   9.460 -by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
   9.461 -       ("diff","%n x. if n mod 4 = 0 then sin(x) \
   9.462 -\                     else if n mod 4 = 1 then cos(x) \
   9.463 -\                     else if n mod 4 = 2 then -sin(x) \
   9.464 -\                     else -cos(x)")] Maclaurin_all_le_objl 1);
   9.465 -by (Step_tac 1);
   9.466 -by (Asm_full_simp_tac 1);
   9.467 -by (stac mod_Suc_eq_Suc_mod 1);
   9.468 -by (cut_inst_tac [("m1","m")] (CLAIM "0 < (4::nat)" RS mod_less_divisor
   9.469 -    RS lemma_exhaust_less_4) 1);
   9.470 -by (Step_tac 1);
   9.471 -by (Asm_simp_tac 1);
   9.472 -by (Asm_simp_tac 1);
   9.473 -by (Asm_simp_tac 1);
   9.474 -by (rtac DERIV_minus 1 THEN Simp_tac 1);
   9.475 -by (Asm_simp_tac 1);
   9.476 -by (rtac lemma_DERIV_subst 1 THEN rtac DERIV_minus 1 THEN rtac DERIV_cos 1);
   9.477 -by (Simp_tac 1);
   9.478 -by (dtac ssubst 1 THEN assume_tac 2);
   9.479 -by (rtac (ARITH_PROVE "[|x = y; abs u <= (v::real) |] ==> abs ((x + u) - y) <= v") 1);
   9.480 -by (rtac sumr_fun_eq 1);
   9.481 -by (Step_tac 1);
   9.482 -by (rtac (CLAIM "x = y ==> x * z = y * (z::real)") 1);
   9.483 -by (stac even_even_mod_4_iff 1);
   9.484 -by (cut_inst_tac [("m1","r")] (CLAIM "0 < (4::nat)" RS mod_less_divisor
   9.485 -    RS lemma_exhaust_less_4) 1);
   9.486 -by (Step_tac 1);
   9.487 -by (Asm_simp_tac 1);
   9.488 -by (asm_simp_tac (simpset() addsimps [even_num_iff]) 2);
   9.489 -by (asm_simp_tac (simpset() addsimps [even_num_iff]) 1);
   9.490 -by (asm_simp_tac (simpset() addsimps [even_num_iff]) 2);
   9.491 -by (dtac lemma_even_mod_4_div_2 1);
   9.492 -by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2,real_divide_def]) 1);
   9.493 -by (dtac lemma_odd_mod_4_div_2 1);
   9.494 -by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2, real_divide_def]) 1);
   9.495 -by (auto_tac (claset() addSIs [real_mult_le_lemma,mult_right_mono],
   9.496 -      simpset() addsimps [real_divide_def,abs_mult,abs_inverse,power_abs RS
   9.497 -sym]));
   9.498 -qed "Maclaurin_sin_bound";
   9.499 -
   9.500 -Goal "0 < n --> Suc (Suc (2 * n - 2)) = 2*n";
   9.501 -by (induct_tac "n" 1);
   9.502 -by Auto_tac;
   9.503 -qed_spec_mp "Suc_Suc_mult_two_diff_two";
   9.504 -Addsimps [Suc_Suc_mult_two_diff_two];
   9.505 -
   9.506 -Goal "0 < n --> Suc (Suc (4*n - 2)) = 4*n";
   9.507 -by (induct_tac "n" 1);
   9.508 -by Auto_tac;
   9.509 -qed_spec_mp "lemma_Suc_Suc_4n_diff_2";
   9.510 -Addsimps [lemma_Suc_Suc_4n_diff_2];
   9.511 -
   9.512 -Goal "0 < n --> Suc (2 * n - 1) = 2*n";
   9.513 -by (induct_tac "n" 1);
   9.514 -by Auto_tac;
   9.515 -qed_spec_mp "Suc_mult_two_diff_one";
   9.516 -Addsimps [Suc_mult_two_diff_one];
   9.517 -
   9.518 -Goal "EX t. sin x = \
   9.519 -\      (sumr 0 n (%m. (if even m then 0 \
   9.520 -\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
   9.521 -\                      x ^ m)) \
   9.522 -\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.523 -by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
   9.524 -       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
   9.525 -       Maclaurin_all_lt_objl 1);
   9.526 -by (Safe_tac);
   9.527 -by (Simp_tac 1);
   9.528 -by (Simp_tac 1);
   9.529 -by (case_tac "n" 1);
   9.530 -by (Clarify_tac 1); 
   9.531 -by (Asm_full_simp_tac 1);
   9.532 -by (dres_inst_tac [("x","0")] spec 1 THEN Asm_full_simp_tac 1);
   9.533 -by (Asm_full_simp_tac 1);
   9.534 -by (rtac ccontr 1);
   9.535 -by (Asm_full_simp_tac 1);
   9.536 -by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
   9.537 -by (dtac ssubst 1 THEN assume_tac 2);
   9.538 -by (res_inst_tac [("x","t")] exI 1);
   9.539 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.540 -by (rtac sumr_fun_eq 1);
   9.541 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.542 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
   9.543 -(*Could sin_zero_iff help?*)
   9.544 -qed "Maclaurin_sin_expansion";
   9.545 -
   9.546 -Goal "EX t. abs t <= abs x &  \
   9.547 -\      sin x = \
   9.548 -\      (sumr 0 n (%m. (if even m then 0 \
   9.549 -\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
   9.550 -\                      x ^ m)) \
   9.551 -\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.552 -
   9.553 -by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
   9.554 -       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
   9.555 -       Maclaurin_all_lt_objl 1);
   9.556 -by (Step_tac 1);
   9.557 -by (Simp_tac 1);
   9.558 -by (Simp_tac 1);
   9.559 -by (case_tac "n" 1);
   9.560 -by (Clarify_tac 1); 
   9.561 -by (Asm_full_simp_tac 1);
   9.562 -by (Asm_full_simp_tac 1);
   9.563 -by (rtac ccontr 1);
   9.564 -by (Asm_full_simp_tac 1);
   9.565 -by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
   9.566 -by (dtac ssubst 1 THEN assume_tac 2);
   9.567 -by (res_inst_tac [("x","t")] exI 1);
   9.568 -by (rtac conjI 1);
   9.569 -by (arith_tac 1);
   9.570 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.571 -by (rtac sumr_fun_eq 1);
   9.572 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.573 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
   9.574 -qed "Maclaurin_sin_expansion2";
   9.575 -
   9.576 -Goal "[| 0 < n; 0 < x |] ==> \
   9.577 -\      EX t. 0 < t & t < x & \
   9.578 -\      sin x = \
   9.579 -\      (sumr 0 n (%m. (if even m then 0 \
   9.580 -\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
   9.581 -\                      x ^ m)) \
   9.582 -\     + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)";
   9.583 -by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
   9.584 -       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
   9.585 -       Maclaurin_objl 1);
   9.586 -by (Step_tac 1);
   9.587 -by (Asm_full_simp_tac 1);
   9.588 -by (Simp_tac 1);
   9.589 -by (dtac ssubst 1 THEN assume_tac 2);
   9.590 -by (res_inst_tac [("x","t")] exI 1);
   9.591 -by (rtac conjI 1 THEN rtac conjI 2);
   9.592 -by (assume_tac 1 THEN assume_tac 1);
   9.593 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.594 -by (rtac sumr_fun_eq 1);
   9.595 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.596 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
   9.597 -qed "Maclaurin_sin_expansion3";
   9.598 -
   9.599 -Goal "0 < x ==> \
   9.600 -\      EX t. 0 < t & t <= x & \
   9.601 -\      sin x = \
   9.602 -\      (sumr 0 n (%m. (if even m then 0 \
   9.603 -\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
   9.604 -\                      x ^ m)) \
   9.605 -\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.606 -by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
   9.607 -       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
   9.608 -       Maclaurin2_objl 1);
   9.609 -by (Step_tac 1);
   9.610 -by (Asm_full_simp_tac 1);
   9.611 -by (Simp_tac 1);
   9.612 -by (dtac ssubst 1 THEN assume_tac 2);
   9.613 -by (res_inst_tac [("x","t")] exI 1);
   9.614 -by (rtac conjI 1 THEN rtac conjI 2);
   9.615 -by (assume_tac 1 THEN assume_tac 1);
   9.616 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.617 -by (rtac sumr_fun_eq 1);
   9.618 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.619 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
   9.620 -qed "Maclaurin_sin_expansion4";
   9.621 -
   9.622 -(*-----------------------------------------------------------------------------*)
   9.623 -(* Maclaurin expansion for cos                                                 *)
   9.624 -(*-----------------------------------------------------------------------------*)
   9.625 -
   9.626 -Goal "sumr 0 (Suc n) \
   9.627 -\        (%m. (if even m \
   9.628 -\              then (- 1) ^ (m div 2)/(real  (fact m)) \
   9.629 -\              else 0) * \
   9.630 -\             0 ^ m) = 1";
   9.631 -by (induct_tac "n" 1);
   9.632 -by Auto_tac;
   9.633 -qed "sumr_cos_zero_one";
   9.634 -Addsimps [sumr_cos_zero_one];
   9.635 -
   9.636 -Goal "EX t. abs t <= abs x & \
   9.637 -\      cos x = \
   9.638 -\      (sumr 0 n (%m. (if even m \
   9.639 -\                      then (- 1) ^ (m div 2)/(real (fact m)) \
   9.640 -\                      else 0) * \
   9.641 -\                      x ^ m)) \
   9.642 -\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.643 -by (cut_inst_tac [("f","cos"),("n","n"),("x","x"),
   9.644 -       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
   9.645 -       Maclaurin_all_lt_objl 1);
   9.646 -by (Step_tac 1);
   9.647 -by (Simp_tac 1);
   9.648 -by (Simp_tac 1);
   9.649 -by (case_tac "n" 1);
   9.650 -by (Asm_full_simp_tac 1);
   9.651 -by (asm_full_simp_tac (simpset() delsimps [sumr_Suc]) 1);
   9.652 -by (rtac ccontr 1);
   9.653 -by (Asm_full_simp_tac 1);
   9.654 -by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
   9.655 -by (dtac ssubst 1 THEN assume_tac 2);
   9.656 -by (res_inst_tac [("x","t")] exI 1);
   9.657 -by (rtac conjI 1);
   9.658 -by (arith_tac 1);
   9.659 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.660 -by (rtac sumr_fun_eq 1);
   9.661 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.662 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps 
   9.663 -    [fact_Suc,realpow_Suc]));
   9.664 -by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
   9.665 -qed "Maclaurin_cos_expansion";
   9.666 -
   9.667 -Goal "[| 0 < x; 0 < n |] ==> \
   9.668 -\      EX t. 0 < t & t < x & \
   9.669 -\      cos x = \
   9.670 -\      (sumr 0 n (%m. (if even m \
   9.671 -\                      then (- 1) ^ (m div 2)/(real (fact m)) \
   9.672 -\                      else 0) * \
   9.673 -\                      x ^ m)) \
   9.674 -\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.675 -by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
   9.676 -       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
   9.677 -       Maclaurin_objl 1);
   9.678 -by (Step_tac 1);
   9.679 -by (Asm_full_simp_tac 1);
   9.680 -by (Simp_tac 1);
   9.681 -by (dtac ssubst 1 THEN assume_tac 2);
   9.682 -by (res_inst_tac [("x","t")] exI 1);
   9.683 -by (rtac conjI 1 THEN rtac conjI 2);
   9.684 -by (assume_tac 1 THEN assume_tac 1);
   9.685 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.686 -by (rtac sumr_fun_eq 1);
   9.687 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.688 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
   9.689 -by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
   9.690 -qed "Maclaurin_cos_expansion2";
   9.691 -
   9.692 -Goal "[| x < 0; 0 < n |] ==> \
   9.693 -\      EX t. x < t & t < 0 & \
   9.694 -\      cos x = \
   9.695 -\      (sumr 0 n (%m. (if even m \
   9.696 -\                      then (- 1) ^ (m div 2)/(real (fact m)) \
   9.697 -\                      else 0) * \
   9.698 -\                      x ^ m)) \
   9.699 -\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
   9.700 -by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
   9.701 -       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
   9.702 -       Maclaurin_minus_objl 1);
   9.703 -by (Step_tac 1);
   9.704 -by (Asm_full_simp_tac 1);
   9.705 -by (Simp_tac 1);
   9.706 -by (dtac ssubst 1 THEN assume_tac 2);
   9.707 -by (res_inst_tac [("x","t")] exI 1);
   9.708 -by (rtac conjI 1 THEN rtac conjI 2);
   9.709 -by (assume_tac 1 THEN assume_tac 1);
   9.710 -by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
   9.711 -by (rtac sumr_fun_eq 1);
   9.712 -by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
   9.713 -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
   9.714 -by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
   9.715 -qed "Maclaurin_minus_cos_expansion";
   9.716 -
   9.717 -(* ------------------------------------------------------------------------- *)
   9.718 -(* Version for ln(1 +/- x). Where is it??                                    *)
   9.719 -(* ------------------------------------------------------------------------- *)
   9.720 -
    10.1 --- a/src/HOL/Hyperreal/MacLaurin.thy	Tue May 11 14:00:02 2004 +0200
    10.2 +++ b/src/HOL/Hyperreal/MacLaurin.thy	Tue May 11 20:11:08 2004 +0200
    10.3 @@ -4,4 +4,46 @@
    10.4      Description : MacLaurin series
    10.5  *)
    10.6  
    10.7 -MacLaurin = Log
    10.8 +theory MacLaurin = Log
    10.9 +files ("MacLaurin_lemmas.ML"):
   10.10 +
   10.11 +use "MacLaurin_lemmas.ML"
   10.12 +
   10.13 +lemma Maclaurin_sin_bound: 
   10.14 +  "abs(sin x - sumr 0 n (%m. (if even m then 0 else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * 
   10.15 +  x ^ m))  <= inverse(real (fact n)) * abs(x) ^ n"
   10.16 +proof -
   10.17 +  have "!! x (y::real). x <= 1 \<Longrightarrow> 0 <= y \<Longrightarrow> x * y \<le> 1 * y" 
   10.18 +    by (rule_tac mult_right_mono,simp_all)
   10.19 +  note est = this[simplified]
   10.20 +  show ?thesis
   10.21 +    apply (cut_tac f=sin and n=n and x=x and 
   10.22 +      diff = "%n x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)"
   10.23 +      in Maclaurin_all_le_objl)
   10.24 +    apply (tactic{* (Step_tac 1) *})
   10.25 +    apply (simp)
   10.26 +    apply (subst mod_Suc_eq_Suc_mod)
   10.27 +    apply (tactic{* cut_inst_tac [("m1","m")] (CLAIM "0 < (4::nat)" RS mod_less_divisor RS lemma_exhaust_less_4) 1*})
   10.28 +    apply (tactic{* Step_tac 1 *})
   10.29 +    apply (simp)+
   10.30 +    apply (rule DERIV_minus, simp+)
   10.31 +    apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp)
   10.32 +    apply (tactic{* dtac ssubst 1 THEN assume_tac 2 *})
   10.33 +    apply (tactic {* rtac (ARITH_PROVE "[|x = y; abs u <= (v::real) |] ==> abs ((x + u) - y) <= v") 1 *})
   10.34 +    apply (rule sumr_fun_eq)
   10.35 +    apply (tactic{* Step_tac 1 *})
   10.36 +    apply (tactic{*rtac (CLAIM "x = y ==> x * z = y * (z::real)") 1*})
   10.37 +    apply (subst even_even_mod_4_iff)
   10.38 +    apply (tactic{* cut_inst_tac [("m1","r")] (CLAIM "0 < (4::nat)" RS mod_less_divisor RS lemma_exhaust_less_4) 1 *})
   10.39 +    apply (tactic{* Step_tac 1 *})
   10.40 +    apply (simp)
   10.41 +    apply (simp_all add:even_num_iff)
   10.42 +    apply (drule lemma_even_mod_4_div_2[simplified])
   10.43 +    apply(simp add: numeral_2_eq_2 real_divide_def)
   10.44 +    apply (drule lemma_odd_mod_4_div_2 );
   10.45 +    apply (simp add: numeral_2_eq_2 real_divide_def)
   10.46 +    apply (auto intro: real_mult_le_lemma mult_right_mono simp add: est mult_pos_le mult_ac real_divide_def abs_mult abs_inverse power_abs[symmetric])
   10.47 +    done
   10.48 +qed
   10.49 +
   10.50 +end
   10.51 \ No newline at end of file
    11.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    11.2 +++ b/src/HOL/Hyperreal/MacLaurin_lemmas.ML	Tue May 11 20:11:08 2004 +0200
    11.3 @@ -0,0 +1,672 @@
    11.4 +(*  Title       : MacLaurin.thy
    11.5 +    Author      : Jacques D. Fleuriot
    11.6 +    Copyright   : 2001 University of Edinburgh
    11.7 +    Description : MacLaurin series
    11.8 +*)
    11.9 +
   11.10 +Goal "sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
   11.11 +by (induct_tac "n" 1);
   11.12 +by Auto_tac;
   11.13 +qed "sumr_offset";
   11.14 +
   11.15 +Goal "ALL f. sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
   11.16 +by (induct_tac "n" 1);
   11.17 +by Auto_tac;
   11.18 +qed "sumr_offset2";
   11.19 +
   11.20 +Goal "sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
   11.21 +by (simp_tac (simpset() addsimps [sumr_offset]) 1);
   11.22 +qed "sumr_offset3";
   11.23 +
   11.24 +Goal "ALL n f. sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
   11.25 +by (simp_tac (simpset() addsimps [sumr_offset]) 1);
   11.26 +qed "sumr_offset4";
   11.27 +
   11.28 +Goal "0 < n ==> \
   11.29 +\     sumr (Suc 0) (Suc n) (%n. (if even(n) then 0 else \
   11.30 +\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n) = \
   11.31 +\     sumr 0 (Suc n) (%n. (if even(n) then 0 else \
   11.32 +\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n)";
   11.33 +by (res_inst_tac [("n1","1")] (sumr_split_add RS subst) 1);
   11.34 +by Auto_tac;
   11.35 +qed "sumr_from_1_from_0";
   11.36 +
   11.37 +(*---------------------------------------------------------------------------*)
   11.38 +(* Maclaurin's theorem with Lagrange form of remainder                       *)
   11.39 +(*---------------------------------------------------------------------------*)
   11.40 +
   11.41 +(* Annoying: Proof is now even longer due mostly to 
   11.42 +   change in behaviour of simplifier  since Isabelle99 *)
   11.43 +Goal " [| 0 < h; 0 < n; diff 0 = f; \
   11.44 +\      ALL m t. \
   11.45 +\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
   11.46 +\   ==> EX t. 0 < t & \
   11.47 +\             t < h & \
   11.48 +\             f h = \
   11.49 +\             sumr 0 n (%m. (diff m 0 / real (fact m)) * h ^ m) + \
   11.50 +\             (diff n t / real (fact n)) * h ^ n";
   11.51 +by (case_tac "n = 0" 1);
   11.52 +by (Force_tac 1);
   11.53 +by (dtac not0_implies_Suc 1);
   11.54 +by (etac exE 1);
   11.55 +by (subgoal_tac 
   11.56 +     "EX B. f h = sumr 0 n (%m. (diff m 0 / real (fact m)) * (h ^ m)) \
   11.57 +\                  + (B * ((h ^ n) / real (fact n)))" 1);
   11.58 +
   11.59 +by (simp_tac (HOL_ss addsimps [real_add_commute, real_divide_def,
   11.60 +    ARITH_PROVE "(x = z + (y::real)) = (x - y = z)"]) 2);
   11.61 +by (res_inst_tac 
   11.62 +  [("x","(f(h) - sumr 0 n (%m. (diff(m)(0) / real (fact m)) * (h ^ m))) \
   11.63 +\        * real (fact n) / (h ^ n)")] exI 2);
   11.64 +by (simp_tac (HOL_ss addsimps [real_mult_assoc,real_divide_def]) 2);
   11.65 + by (rtac (CLAIM "x = (1::real) ==>  a = a * (x::real)") 2);
   11.66 +by (asm_simp_tac (HOL_ss addsimps 
   11.67 +    [CLAIM "(a::real) * (b * (c * d)) = (d * a) * (b * c)"]
   11.68 +     delsimps [realpow_Suc]) 2);
   11.69 +by (stac left_inverse 2);
   11.70 +by (stac left_inverse 3);
   11.71 +by (rtac (real_not_refl2 RS not_sym) 2);
   11.72 +by (etac zero_less_power 2);
   11.73 +by (rtac real_of_nat_fact_not_zero 2);
   11.74 +by (Simp_tac 2);
   11.75 +by (etac exE 1);
   11.76 +by (cut_inst_tac [("b","%t. f t - \
   11.77 +\      (sumr 0 n (%m. (diff m 0 / real (fact m)) * (t ^ m)) + \
   11.78 +\                       (B * ((t ^ n) / real (fact n))))")] 
   11.79 +    (CLAIM "EX g. g = b") 1);
   11.80 +by (etac exE 1);
   11.81 +by (subgoal_tac "g 0 = 0 & g h =0" 1);
   11.82 +by (asm_simp_tac (simpset() addsimps 
   11.83 +    [ARITH_PROVE "(x - y = z) = (x = z + (y::real))"]
   11.84 +    delsimps [sumr_Suc]) 2);
   11.85 +by (cut_inst_tac [("n","m"),("k","1")] sumr_offset2 2);
   11.86 +by (asm_full_simp_tac (simpset() addsimps 
   11.87 +    [ARITH_PROVE "(x = y - z) = (y = x + (z::real))"]
   11.88 +    delsimps [sumr_Suc]) 2);
   11.89 +by (cut_inst_tac [("b","%m t. diff m t - \
   11.90 +\      (sumr 0 (n - m) (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) \
   11.91 +\       + (B * ((t ^ (n - m)) / real (fact(n - m)))))")] 
   11.92 +    (CLAIM "EX difg. difg = b") 1);
   11.93 +by (etac exE 1);
   11.94 +by (subgoal_tac "difg 0 = g" 1);
   11.95 +by (asm_simp_tac (simpset() delsimps [realpow_Suc,fact_Suc]) 2);
   11.96 +by (subgoal_tac "ALL m t. m < n & 0 <= t & t <= h --> \
   11.97 +\                   DERIV (difg m) t :> difg (Suc m) t" 1);
   11.98 +by (Clarify_tac 2);
   11.99 +by (rtac DERIV_diff 2);
  11.100 +by (Asm_simp_tac 2);
  11.101 +by DERIV_tac;
  11.102 +by DERIV_tac;
  11.103 +by (rtac lemma_DERIV_subst 3);
  11.104 +by (rtac DERIV_quotient 3);
  11.105 +by (rtac DERIV_const 4);
  11.106 +by (rtac DERIV_pow 3);
  11.107 +by (asm_simp_tac (simpset() addsimps [inverse_mult_distrib,
  11.108 +    CLAIM_SIMP "(a::real) * b * c * (d * e) = a * b * (c * d) * e" 
  11.109 +    mult_ac,fact_diff_Suc]) 4);
  11.110 +by (Asm_simp_tac 3);
  11.111 +by (forw_inst_tac [("m","ma")] less_add_one 2);
  11.112 +by (Clarify_tac 2);
  11.113 +by (asm_simp_tac (simpset() addsimps 
  11.114 +    [CLAIM "Suc m = ma + d + 1 ==> m - ma = d"]
  11.115 +    delsimps [sumr_Suc]) 2);
  11.116 +by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
  11.117 +          (read_instantiate [("k","1")] sumr_offset4))] 
  11.118 +    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
  11.119 +by (rtac lemma_DERIV_subst 2);
  11.120 +by (rtac DERIV_add 2);
  11.121 +by (rtac DERIV_const 3);
  11.122 +by (rtac DERIV_sumr 2);
  11.123 +by (Clarify_tac 2);
  11.124 +by (Simp_tac 3);
  11.125 +by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc] 
  11.126 +    delsimps [fact_Suc,realpow_Suc]) 2);
  11.127 +by (rtac DERIV_cmult 2);
  11.128 +by (rtac lemma_DERIV_subst 2);
  11.129 +by DERIV_tac;
  11.130 +by (stac fact_Suc 2);
  11.131 +by (stac real_of_nat_mult 2);
  11.132 +by (simp_tac (simpset() addsimps [inverse_mult_distrib] @
  11.133 +    mult_ac) 2);
  11.134 +by (subgoal_tac "ALL ma. ma < n --> \
  11.135 +\        (EX t. 0 < t & t < h & difg (Suc ma) t = 0)" 1);
  11.136 +by (rotate_tac 11 1);
  11.137 +by (dres_inst_tac [("x","m")] spec 1);
  11.138 +by (etac impE 1);
  11.139 +by (Asm_simp_tac 1);
  11.140 +by (etac exE 1);
  11.141 +by (res_inst_tac [("x","t")] exI 1);
  11.142 +by (asm_full_simp_tac (simpset() addsimps 
  11.143 +     [ARITH_PROVE "(x - y = 0) = (y = (x::real))"] 
  11.144 +      delsimps [realpow_Suc,fact_Suc]) 1);
  11.145 +by (subgoal_tac "ALL m. m < n --> difg m 0 = 0" 1);
  11.146 +by (Clarify_tac 2);
  11.147 +by (Asm_simp_tac 2);
  11.148 +by (forw_inst_tac [("m","ma")] less_add_one 2);
  11.149 +by (Clarify_tac 2);
  11.150 +by (asm_simp_tac (simpset() delsimps [sumr_Suc]) 2);
  11.151 +by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
  11.152 +          (read_instantiate [("k","1")] sumr_offset4))] 
  11.153 +    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
  11.154 +by (subgoal_tac "ALL m. m < n --> (EX t. 0 < t & t < h & \
  11.155 +\                DERIV (difg m) t :> 0)" 1);
  11.156 +by (rtac allI 1 THEN rtac impI 1);
  11.157 +by (rotate_tac 12 1);
  11.158 +by (dres_inst_tac [("x","ma")] spec 1);
  11.159 +by (etac impE 1 THEN assume_tac 1);
  11.160 +by (etac exE 1);
  11.161 +by (res_inst_tac [("x","t")] exI 1);
  11.162 +(* do some tidying up *)
  11.163 +by (ALLGOALS(thin_tac "difg = \
  11.164 +\          (%m t. diff m t - \
  11.165 +\                 (sumr 0 (n - m) \
  11.166 +\                   (%p. diff (m + p) 0 / real (fact p) * t ^ p) + \
  11.167 +\                  B * (t ^ (n - m) / real (fact (n - m)))))"));
  11.168 +by (ALLGOALS(thin_tac "g = \
  11.169 +\          (%t. f t - \
  11.170 +\               (sumr 0 n (%m. diff m 0 / real  (fact m) * t ^ m) + \
  11.171 +\                B * (t ^ n / real (fact n))))"));
  11.172 +by (ALLGOALS(thin_tac "f h = \
  11.173 +\          sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.174 +\          B * (h ^ n / real (fact n))"));
  11.175 +(* back to business *)
  11.176 +by (Asm_simp_tac 1);
  11.177 +by (rtac DERIV_unique 1);
  11.178 +by (Blast_tac 2);
  11.179 +by (Force_tac 1);
  11.180 +by (rtac allI 1 THEN induct_tac "ma" 1);
  11.181 +by (rtac impI 1 THEN rtac Rolle 1);
  11.182 +by (assume_tac 1);
  11.183 +by (Asm_full_simp_tac 1);
  11.184 +by (Asm_full_simp_tac 1);
  11.185 +by (subgoal_tac "ALL t. 0 <= t & t <= h --> g differentiable t" 1);
  11.186 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
  11.187 +by (blast_tac (claset() addDs [DERIV_isCont]) 1);
  11.188 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
  11.189 +by (Clarify_tac 1);
  11.190 +by (res_inst_tac [("x","difg (Suc 0) t")] exI 1);
  11.191 +by (Force_tac 1);
  11.192 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
  11.193 +by (Clarify_tac 1);
  11.194 +by (res_inst_tac [("x","difg (Suc 0) x")] exI 1);
  11.195 +by (Force_tac 1);
  11.196 +by (Step_tac 1);
  11.197 +by (Force_tac 1);
  11.198 +by (subgoal_tac "EX ta. 0 < ta & ta < t & \
  11.199 +\                DERIV difg (Suc n) ta :> 0" 1);
  11.200 +by (rtac Rolle 2 THEN assume_tac 2);
  11.201 +by (Asm_full_simp_tac 2);
  11.202 +by (rotate_tac 2 2);
  11.203 +by (dres_inst_tac [("x","n")] spec 2);
  11.204 +by (ftac (ARITH_PROVE "n < m  ==> n < Suc m") 2);
  11.205 +by (rtac DERIV_unique 2);
  11.206 +by (assume_tac 3);
  11.207 +by (Force_tac 2);
  11.208 +by (subgoal_tac 
  11.209 +    "ALL ta. 0 <= ta & ta <= t --> (difg (Suc n)) differentiable ta" 2);
  11.210 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
  11.211 +by (blast_tac (claset() addSDs [DERIV_isCont]) 2);
  11.212 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
  11.213 +by (Clarify_tac 2);
  11.214 +by (res_inst_tac [("x","difg (Suc (Suc n)) ta")] exI 2);
  11.215 +by (Force_tac 2);
  11.216 +by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
  11.217 +by (Clarify_tac 2);
  11.218 +by (res_inst_tac [("x","difg (Suc (Suc n)) x")] exI 2);
  11.219 +by (Force_tac 2);
  11.220 +by (Step_tac 1);
  11.221 +by (res_inst_tac [("x","ta")] exI 1);
  11.222 +by (Force_tac 1);
  11.223 +qed "Maclaurin";
  11.224 +
  11.225 +Goal "0 < h & 0 < n & diff 0 = f & \
  11.226 +\      (ALL m t. \
  11.227 +\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
  11.228 +\   --> (EX t. 0 < t & \
  11.229 +\             t < h & \
  11.230 +\             f h = \
  11.231 +\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.232 +\             diff n t / real (fact n) * h ^ n)";
  11.233 +by (blast_tac (claset() addIs [Maclaurin]) 1);
  11.234 +qed "Maclaurin_objl";
  11.235 +
  11.236 +Goal " [| 0 < h; diff 0 = f; \
  11.237 +\      ALL m t. \
  11.238 +\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
  11.239 +\   ==> EX t. 0 < t & \
  11.240 +\             t <= h & \
  11.241 +\             f h = \
  11.242 +\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.243 +\             diff n t / real (fact n) * h ^ n";
  11.244 +by (case_tac "n" 1);
  11.245 +by Auto_tac;
  11.246 +by (dtac Maclaurin 1 THEN Auto_tac);
  11.247 +qed "Maclaurin2";
  11.248 +
  11.249 +Goal "0 < h & diff 0 = f & \
  11.250 +\      (ALL m t. \
  11.251 +\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
  11.252 +\   --> (EX t. 0 < t & \
  11.253 +\             t <= h & \
  11.254 +\             f h = \
  11.255 +\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.256 +\             diff n t / real (fact n) * h ^ n)";
  11.257 +by (blast_tac (claset() addIs [Maclaurin2]) 1);
  11.258 +qed "Maclaurin2_objl";
  11.259 +
  11.260 +Goal " [| h < 0; 0 < n; diff 0 = f; \
  11.261 +\      ALL m t. \
  11.262 +\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t |] \
  11.263 +\   ==> EX t. h < t & \
  11.264 +\             t < 0 & \
  11.265 +\             f h = \
  11.266 +\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.267 +\             diff n t / real (fact n) * h ^ n";
  11.268 +by (cut_inst_tac [("f","%x. f (-x)"),
  11.269 +                 ("diff","%n x. ((- 1) ^ n) * diff n (-x)"),
  11.270 +                 ("h","-h"),("n","n")] Maclaurin_objl 1);
  11.271 +by (Asm_full_simp_tac 1);
  11.272 +by (etac impE 1 THEN Step_tac 1);
  11.273 +by (stac minus_mult_right 1);
  11.274 +by (rtac DERIV_cmult 1);
  11.275 +by (rtac lemma_DERIV_subst 1);
  11.276 +by (rtac (read_instantiate [("g","uminus")] DERIV_chain2) 1);
  11.277 +by (rtac DERIV_minus 2 THEN rtac DERIV_Id 2);
  11.278 +by (Force_tac 2);
  11.279 +by (Force_tac 1);
  11.280 +by (res_inst_tac [("x","-t")] exI 1);
  11.281 +by Auto_tac;
  11.282 +by (rtac (CLAIM "[| x = x'; y = y' |] ==> x + y = x' + (y'::real)") 1);
  11.283 +by (rtac sumr_fun_eq 1);
  11.284 +by (Asm_full_simp_tac 1);
  11.285 +by (auto_tac (claset(),simpset() addsimps [real_divide_def,
  11.286 +    CLAIM "((a * b) * c) * d = (b * c) * (a * (d::real))",
  11.287 +    power_mult_distrib RS sym]));
  11.288 +qed "Maclaurin_minus";
  11.289 +
  11.290 +Goal "(h < 0 & 0 < n & diff 0 = f & \
  11.291 +\      (ALL m t. \
  11.292 +\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t))\
  11.293 +\   --> (EX t. h < t & \
  11.294 +\             t < 0 & \
  11.295 +\             f h = \
  11.296 +\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
  11.297 +\             diff n t / real (fact n) * h ^ n)";
  11.298 +by (blast_tac (claset() addIs [Maclaurin_minus]) 1);
  11.299 +qed "Maclaurin_minus_objl";
  11.300 +
  11.301 +(* ------------------------------------------------------------------------- *)
  11.302 +(* More convenient "bidirectional" version.                                  *)
  11.303 +(* ------------------------------------------------------------------------- *)
  11.304 +
  11.305 +(* not good for PVS sin_approx, cos_approx *)
  11.306 +Goal " [| diff 0 = f; \
  11.307 +\      ALL m t. \
  11.308 +\         m < n & abs t <= abs x --> DERIV (diff m) t :> diff (Suc m) t |] \
  11.309 +\   ==> EX t. abs t <= abs x & \
  11.310 +\             f x = \
  11.311 +\             sumr 0 n (%m. diff m 0 / real (fact m) * x ^ m) + \
  11.312 +\             diff n t / real (fact n) * x ^ n";
  11.313 +by (case_tac "n = 0" 1);
  11.314 +by (Force_tac 1);
  11.315 +by (case_tac "x = 0" 1);
  11.316 +by (res_inst_tac [("x","0")] exI 1);
  11.317 +by (Asm_full_simp_tac 1);
  11.318 +by (res_inst_tac [("P","0 < n")] impE 1);
  11.319 +by (assume_tac 2 THEN assume_tac 2);
  11.320 +by (induct_tac "n" 1);
  11.321 +by (Simp_tac 1);
  11.322 +by Auto_tac;
  11.323 +by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
  11.324 +by Auto_tac;
  11.325 +by (cut_inst_tac [("f","diff 0"),
  11.326 +                 ("diff","diff"),
  11.327 +                 ("h","x"),("n","n")] Maclaurin_objl 2);
  11.328 +by (Step_tac 2);
  11.329 +by (blast_tac (claset() addDs 
  11.330 +    [ARITH_PROVE "[|(0::real) <= t;t <= x |] ==> abs t <= abs x"]) 2);
  11.331 +by (res_inst_tac [("x","t")] exI 2);
  11.332 +by (force_tac (claset() addIs 
  11.333 +    [ARITH_PROVE "[| 0 < t; (t::real) < x|] ==> abs t <= abs x"],simpset()) 2);
  11.334 +by (cut_inst_tac [("f","diff 0"),
  11.335 +                 ("diff","diff"),
  11.336 +                 ("h","x"),("n","n")] Maclaurin_minus_objl 1);
  11.337 +by (Step_tac 1);
  11.338 +by (blast_tac (claset() addDs 
  11.339 +    [ARITH_PROVE "[|x <= t;t <= (0::real) |] ==> abs t <= abs x"]) 1);
  11.340 +by (res_inst_tac [("x","t")] exI 1);
  11.341 +by (force_tac (claset() addIs 
  11.342 +    [ARITH_PROVE "[| x < t; (t::real) < 0|] ==> abs t <= abs x"],simpset()) 1);
  11.343 +qed "Maclaurin_bi_le";
  11.344 +
  11.345 +Goal "[| diff 0 = f; \
  11.346 +\        ALL m x. DERIV (diff m) x :> diff(Suc m) x; \ 
  11.347 +\       x ~= 0; 0 < n \
  11.348 +\     |] ==> EX t. 0 < abs t & abs t < abs x & \
  11.349 +\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
  11.350 +\                    (diff n t / real (fact n)) * x ^ n";
  11.351 +by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1);
  11.352 +by (Blast_tac 2);
  11.353 +by (dtac Maclaurin_minus 1);
  11.354 +by (dtac Maclaurin 5);
  11.355 +by (TRYALL(assume_tac));
  11.356 +by (Blast_tac 1);
  11.357 +by (Blast_tac 2);
  11.358 +by (Step_tac 1);
  11.359 +by (ALLGOALS(res_inst_tac [("x","t")] exI));
  11.360 +by (Step_tac 1);
  11.361 +by (ALLGOALS(arith_tac));
  11.362 +qed "Maclaurin_all_lt";
  11.363 +
  11.364 +Goal "diff 0 = f & \
  11.365 +\     (ALL m x. DERIV (diff m) x :> diff(Suc m) x) & \
  11.366 +\     x ~= 0 & 0 < n \
  11.367 +\     --> (EX t. 0 < abs t & abs t < abs x & \
  11.368 +\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
  11.369 +\                    (diff n t / real (fact n)) * x ^ n)";
  11.370 +by (blast_tac (claset() addIs [Maclaurin_all_lt]) 1);
  11.371 +qed "Maclaurin_all_lt_objl";
  11.372 +
  11.373 +Goal "x = (0::real)  \
  11.374 +\     ==> 0 < n --> \
  11.375 +\         sumr 0 n (%m. (diff m (0::real) / real (fact m)) * x ^ m) = \
  11.376 +\         diff 0 0";
  11.377 +by (Asm_simp_tac 1);
  11.378 +by (induct_tac "n" 1);
  11.379 +by Auto_tac; 
  11.380 +qed_spec_mp "Maclaurin_zero";
  11.381 +
  11.382 +Goal "[| diff 0 = f; \
  11.383 +\       ALL m x. DERIV (diff m) x :> diff (Suc m) x \
  11.384 +\     |] ==> EX t. abs t <= abs x & \
  11.385 +\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
  11.386 +\                   (diff n t / real (fact n)) * x ^ n";
  11.387 +by (cut_inst_tac [("n","n"),("m","0")] 
  11.388 +       (ARITH_PROVE "n <= m | m < (n::nat)") 1);
  11.389 +by (etac disjE 1);
  11.390 +by (Force_tac 1);
  11.391 +by (case_tac "x = 0" 1);
  11.392 +by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_zero 1);
  11.393 +by (assume_tac 1);
  11.394 +by (dtac (gr_implies_not0 RS  not0_implies_Suc) 1);
  11.395 +by (res_inst_tac [("x","0")] exI 1);
  11.396 +by (Force_tac 1);
  11.397 +by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_all_lt 1);
  11.398 +by (TRYALL(assume_tac));
  11.399 +by (Step_tac 1);
  11.400 +by (res_inst_tac [("x","t")] exI 1);
  11.401 +by Auto_tac;
  11.402 +qed "Maclaurin_all_le";
  11.403 +
  11.404 +Goal "diff 0 = f & \
  11.405 +\     (ALL m x. DERIV (diff m) x :> diff (Suc m) x)  \
  11.406 +\     --> (EX t. abs t <= abs x & \
  11.407 +\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
  11.408 +\                   (diff n t / real (fact n)) * x ^ n)";
  11.409 +by (blast_tac (claset() addIs [Maclaurin_all_le]) 1);
  11.410 +qed "Maclaurin_all_le_objl";
  11.411 +
  11.412 +(* ------------------------------------------------------------------------- *)
  11.413 +(* Version for exp.                                                          *)
  11.414 +(* ------------------------------------------------------------------------- *)
  11.415 +
  11.416 +Goal "[| x ~= 0; 0 < n |] \
  11.417 +\     ==> (EX t. 0 < abs t & \
  11.418 +\               abs t < abs x & \
  11.419 +\               exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
  11.420 +\                       (exp t / real (fact n)) * x ^ n)";
  11.421 +by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
  11.422 +    Maclaurin_all_lt_objl 1);
  11.423 +by Auto_tac;
  11.424 +qed "Maclaurin_exp_lt";
  11.425 +
  11.426 +Goal "EX t. abs t <= abs x & \
  11.427 +\           exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
  11.428 +\                      (exp t / real (fact n)) * x ^ n";
  11.429 +by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
  11.430 +    Maclaurin_all_le_objl 1);
  11.431 +by Auto_tac;
  11.432 +qed "Maclaurin_exp_le";
  11.433 +
  11.434 +(* ------------------------------------------------------------------------- *)
  11.435 +(* Version for sin function                                                  *)
  11.436 +(* ------------------------------------------------------------------------- *)
  11.437 +
  11.438 +Goal "[| a < b; ALL x. a <= x & x <= b --> DERIV f x :> f'(x) |] \
  11.439 +\     ==> EX z. a < z & z < b & (f b - f a = (b - a) * f'(z))";
  11.440 +by (dtac MVT 1);
  11.441 +by (blast_tac (claset() addIs [DERIV_isCont]) 1);
  11.442 +by (force_tac (claset() addDs [order_less_imp_le],
  11.443 +    simpset() addsimps [differentiable_def]) 1);
  11.444 +by (blast_tac (claset() addDs [DERIV_unique,order_less_imp_le]) 1);
  11.445 +qed "MVT2";
  11.446 +
  11.447 +Goal "d < (4::nat) ==> d = 0 | d = 1 | d = 2 | d = 3";
  11.448 +by (case_tac "d" 1 THEN Auto_tac);
  11.449 +qed "lemma_exhaust_less_4";
  11.450 +
  11.451 +bind_thm ("real_mult_le_lemma",
  11.452 +          simplify (simpset()) (inst "b" "1" mult_right_mono));
  11.453 +
  11.454 +
  11.455 +Goal "0 < n --> Suc (Suc (2 * n - 2)) = 2*n";
  11.456 +by (induct_tac "n" 1);
  11.457 +by Auto_tac;
  11.458 +qed_spec_mp "Suc_Suc_mult_two_diff_two";
  11.459 +Addsimps [Suc_Suc_mult_two_diff_two];
  11.460 +
  11.461 +Goal "0 < n --> Suc (Suc (4*n - 2)) = 4*n";
  11.462 +by (induct_tac "n" 1);
  11.463 +by Auto_tac;
  11.464 +qed_spec_mp "lemma_Suc_Suc_4n_diff_2";
  11.465 +Addsimps [lemma_Suc_Suc_4n_diff_2];
  11.466 +
  11.467 +Goal "0 < n --> Suc (2 * n - 1) = 2*n";
  11.468 +by (induct_tac "n" 1);
  11.469 +by Auto_tac;
  11.470 +qed_spec_mp "Suc_mult_two_diff_one";
  11.471 +Addsimps [Suc_mult_two_diff_one];
  11.472 +
  11.473 +Goal "EX t. sin x = \
  11.474 +\      (sumr 0 n (%m. (if even m then 0 \
  11.475 +\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
  11.476 +\                      x ^ m)) \
  11.477 +\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.478 +by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
  11.479 +       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
  11.480 +       Maclaurin_all_lt_objl 1);
  11.481 +by (Safe_tac);
  11.482 +by (Simp_tac 1);
  11.483 +by (Simp_tac 1);
  11.484 +by (case_tac "n" 1);
  11.485 +by (Clarify_tac 1); 
  11.486 +by (Asm_full_simp_tac 1);
  11.487 +by (dres_inst_tac [("x","0")] spec 1 THEN Asm_full_simp_tac 1);
  11.488 +by (Asm_full_simp_tac 1);
  11.489 +by (rtac ccontr 1);
  11.490 +by (Asm_full_simp_tac 1);
  11.491 +by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
  11.492 +by (dtac ssubst 1 THEN assume_tac 2);
  11.493 +by (res_inst_tac [("x","t")] exI 1);
  11.494 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.495 +by (rtac sumr_fun_eq 1);
  11.496 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.497 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
  11.498 +(*Could sin_zero_iff help?*)
  11.499 +qed "Maclaurin_sin_expansion";
  11.500 +
  11.501 +Goal "EX t. abs t <= abs x &  \
  11.502 +\      sin x = \
  11.503 +\      (sumr 0 n (%m. (if even m then 0 \
  11.504 +\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
  11.505 +\                      x ^ m)) \
  11.506 +\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.507 +
  11.508 +by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
  11.509 +       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
  11.510 +       Maclaurin_all_lt_objl 1);
  11.511 +by (Step_tac 1);
  11.512 +by (Simp_tac 1);
  11.513 +by (Simp_tac 1);
  11.514 +by (case_tac "n" 1);
  11.515 +by (Clarify_tac 1); 
  11.516 +by (Asm_full_simp_tac 1);
  11.517 +by (Asm_full_simp_tac 1);
  11.518 +by (rtac ccontr 1);
  11.519 +by (Asm_full_simp_tac 1);
  11.520 +by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
  11.521 +by (dtac ssubst 1 THEN assume_tac 2);
  11.522 +by (res_inst_tac [("x","t")] exI 1);
  11.523 +by (rtac conjI 1);
  11.524 +by (arith_tac 1);
  11.525 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.526 +by (rtac sumr_fun_eq 1);
  11.527 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.528 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
  11.529 +qed "Maclaurin_sin_expansion2";
  11.530 +
  11.531 +Goal "[| 0 < n; 0 < x |] ==> \
  11.532 +\      EX t. 0 < t & t < x & \
  11.533 +\      sin x = \
  11.534 +\      (sumr 0 n (%m. (if even m then 0 \
  11.535 +\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
  11.536 +\                      x ^ m)) \
  11.537 +\     + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)";
  11.538 +by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
  11.539 +       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
  11.540 +       Maclaurin_objl 1);
  11.541 +by (Step_tac 1);
  11.542 +by (Asm_full_simp_tac 1);
  11.543 +by (Simp_tac 1);
  11.544 +by (dtac ssubst 1 THEN assume_tac 2);
  11.545 +by (res_inst_tac [("x","t")] exI 1);
  11.546 +by (rtac conjI 1 THEN rtac conjI 2);
  11.547 +by (assume_tac 1 THEN assume_tac 1);
  11.548 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.549 +by (rtac sumr_fun_eq 1);
  11.550 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.551 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
  11.552 +qed "Maclaurin_sin_expansion3";
  11.553 +
  11.554 +Goal "0 < x ==> \
  11.555 +\      EX t. 0 < t & t <= x & \
  11.556 +\      sin x = \
  11.557 +\      (sumr 0 n (%m. (if even m then 0 \
  11.558 +\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
  11.559 +\                      x ^ m)) \
  11.560 +\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.561 +by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
  11.562 +       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
  11.563 +       Maclaurin2_objl 1);
  11.564 +by (Step_tac 1);
  11.565 +by (Asm_full_simp_tac 1);
  11.566 +by (Simp_tac 1);
  11.567 +by (dtac ssubst 1 THEN assume_tac 2);
  11.568 +by (res_inst_tac [("x","t")] exI 1);
  11.569 +by (rtac conjI 1 THEN rtac conjI 2);
  11.570 +by (assume_tac 1 THEN assume_tac 1);
  11.571 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.572 +by (rtac sumr_fun_eq 1);
  11.573 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.574 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
  11.575 +qed "Maclaurin_sin_expansion4";
  11.576 +
  11.577 +(*-----------------------------------------------------------------------------*)
  11.578 +(* Maclaurin expansion for cos                                                 *)
  11.579 +(*-----------------------------------------------------------------------------*)
  11.580 +
  11.581 +Goal "sumr 0 (Suc n) \
  11.582 +\        (%m. (if even m \
  11.583 +\              then (- 1) ^ (m div 2)/(real  (fact m)) \
  11.584 +\              else 0) * \
  11.585 +\             0 ^ m) = 1";
  11.586 +by (induct_tac "n" 1);
  11.587 +by Auto_tac;
  11.588 +qed "sumr_cos_zero_one";
  11.589 +Addsimps [sumr_cos_zero_one];
  11.590 +
  11.591 +Goal "EX t. abs t <= abs x & \
  11.592 +\      cos x = \
  11.593 +\      (sumr 0 n (%m. (if even m \
  11.594 +\                      then (- 1) ^ (m div 2)/(real (fact m)) \
  11.595 +\                      else 0) * \
  11.596 +\                      x ^ m)) \
  11.597 +\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.598 +by (cut_inst_tac [("f","cos"),("n","n"),("x","x"),
  11.599 +       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
  11.600 +       Maclaurin_all_lt_objl 1);
  11.601 +by (Step_tac 1);
  11.602 +by (Simp_tac 1);
  11.603 +by (Simp_tac 1);
  11.604 +by (case_tac "n" 1);
  11.605 +by (Asm_full_simp_tac 1);
  11.606 +by (asm_full_simp_tac (simpset() delsimps [sumr_Suc]) 1);
  11.607 +by (rtac ccontr 1);
  11.608 +by (Asm_full_simp_tac 1);
  11.609 +by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
  11.610 +by (dtac ssubst 1 THEN assume_tac 2);
  11.611 +by (res_inst_tac [("x","t")] exI 1);
  11.612 +by (rtac conjI 1);
  11.613 +by (arith_tac 1);
  11.614 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.615 +by (rtac sumr_fun_eq 1);
  11.616 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.617 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps 
  11.618 +    [fact_Suc,realpow_Suc]));
  11.619 +by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
  11.620 +qed "Maclaurin_cos_expansion";
  11.621 +
  11.622 +Goal "[| 0 < x; 0 < n |] ==> \
  11.623 +\      EX t. 0 < t & t < x & \
  11.624 +\      cos x = \
  11.625 +\      (sumr 0 n (%m. (if even m \
  11.626 +\                      then (- 1) ^ (m div 2)/(real (fact m)) \
  11.627 +\                      else 0) * \
  11.628 +\                      x ^ m)) \
  11.629 +\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.630 +by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
  11.631 +       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
  11.632 +       Maclaurin_objl 1);
  11.633 +by (Step_tac 1);
  11.634 +by (Asm_full_simp_tac 1);
  11.635 +by (Simp_tac 1);
  11.636 +by (dtac ssubst 1 THEN assume_tac 2);
  11.637 +by (res_inst_tac [("x","t")] exI 1);
  11.638 +by (rtac conjI 1 THEN rtac conjI 2);
  11.639 +by (assume_tac 1 THEN assume_tac 1);
  11.640 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.641 +by (rtac sumr_fun_eq 1);
  11.642 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.643 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
  11.644 +by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
  11.645 +qed "Maclaurin_cos_expansion2";
  11.646 +
  11.647 +Goal "[| x < 0; 0 < n |] ==> \
  11.648 +\      EX t. x < t & t < 0 & \
  11.649 +\      cos x = \
  11.650 +\      (sumr 0 n (%m. (if even m \
  11.651 +\                      then (- 1) ^ (m div 2)/(real (fact m)) \
  11.652 +\                      else 0) * \
  11.653 +\                      x ^ m)) \
  11.654 +\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
  11.655 +by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
  11.656 +       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
  11.657 +       Maclaurin_minus_objl 1);
  11.658 +by (Step_tac 1);
  11.659 +by (Asm_full_simp_tac 1);
  11.660 +by (Simp_tac 1);
  11.661 +by (dtac ssubst 1 THEN assume_tac 2);
  11.662 +by (res_inst_tac [("x","t")] exI 1);
  11.663 +by (rtac conjI 1 THEN rtac conjI 2);
  11.664 +by (assume_tac 1 THEN assume_tac 1);
  11.665 +by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
  11.666 +by (rtac sumr_fun_eq 1);
  11.667 +by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
  11.668 +by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
  11.669 +by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
  11.670 +qed "Maclaurin_minus_cos_expansion";
  11.671 +
  11.672 +(* ------------------------------------------------------------------------- *)
  11.673 +(* Version for ln(1 +/- x). Where is it??                                    *)
  11.674 +(* ------------------------------------------------------------------------- *)
  11.675 +
    12.1 --- a/src/HOL/IMP/Compiler.thy	Tue May 11 14:00:02 2004 +0200
    12.2 +++ b/src/HOL/IMP/Compiler.thy	Tue May 11 20:11:08 2004 +0200
    12.3 @@ -86,7 +86,7 @@
    12.4  lemma [simp]:
    12.5   "\<And>m n. closed m n (C1@C2) =
    12.6           (closed m (n+size C2) C1 \<and> closed (m+size C1) n C2)"
    12.7 -by(induct C1, simp, simp add:plus_ac0)
    12.8 +by(induct C1, simp, simp add:add_ac)
    12.9  
   12.10  theorem [simp]: "\<And>m n. closed m n (compile c)"
   12.11  by(induct c, simp_all add:backws_def forws_def)
    13.1 --- a/src/HOL/Integ/Bin.thy	Tue May 11 14:00:02 2004 +0200
    13.2 +++ b/src/HOL/Integ/Bin.thy	Tue May 11 20:11:08 2004 +0200
    13.3 @@ -10,7 +10,7 @@
    13.4  
    13.5  theory Bin = IntDef + Numeral:
    13.6  
    13.7 -axclass number_ring \<subseteq> number, ring
    13.8 +axclass number_ring \<subseteq> number, comm_ring_1
    13.9    number_of_Pls: "number_of bin.Pls = 0"
   13.10    number_of_Min: "number_of bin.Min = - 1"
   13.11    number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
   13.12 @@ -118,7 +118,7 @@
   13.13  
   13.14  subsection{*Comparisons, for Ordered Rings*}
   13.15  
   13.16 -lemma double_eq_0_iff: "(a + a = 0) = (a = (0::'a::ordered_ring))"
   13.17 +lemma double_eq_0_iff: "(a + a = 0) = (a = (0::'a::ordered_idom))"
   13.18  proof -
   13.19    have "a + a = (1+1)*a" by (simp add: left_distrib)
   13.20    with zero_less_two [where 'a = 'a]
   13.21 @@ -157,7 +157,7 @@
   13.22  
   13.23  
   13.24  text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
   13.25 -lemma Ints_odd_nonzero: "a \<in> Ints ==> 1 + a + a \<noteq> (0::'a::ordered_ring)"
   13.26 +lemma Ints_odd_nonzero: "a \<in> Ints ==> 1 + a + a \<noteq> (0::'a::ordered_idom)"
   13.27  proof (unfold Ints_def) 
   13.28    assume "a \<in> range of_int"
   13.29    then obtain z where a: "a = of_int z" ..
   13.30 @@ -175,17 +175,17 @@
   13.31  
   13.32  lemma iszero_number_of_BIT:
   13.33       "iszero (number_of (w BIT x)::'a) = 
   13.34 -      (~x & iszero (number_of w::'a::{ordered_ring,number_ring}))"
   13.35 +      (~x & iszero (number_of w::'a::{ordered_idom,number_ring}))"
   13.36  by (simp add: iszero_def compare_rls zero_reorient double_eq_0_iff 
   13.37                number_of Ints_odd_nonzero [OF Ints_number_of])
   13.38  
   13.39  lemma iszero_number_of_0:
   13.40 -     "iszero (number_of (w BIT False) :: 'a::{ordered_ring,number_ring}) = 
   13.41 +     "iszero (number_of (w BIT False) :: 'a::{ordered_idom,number_ring}) = 
   13.42        iszero (number_of w :: 'a)"
   13.43  by (simp only: iszero_number_of_BIT simp_thms)
   13.44  
   13.45  lemma iszero_number_of_1:
   13.46 -     "~ iszero (number_of (w BIT True)::'a::{ordered_ring,number_ring})"
   13.47 +     "~ iszero (number_of (w BIT True)::'a::{ordered_idom,number_ring})"
   13.48  by (simp only: iszero_number_of_BIT simp_thms)
   13.49  
   13.50  
   13.51 @@ -193,7 +193,7 @@
   13.52  subsection{*The Less-Than Relation*}
   13.53  
   13.54  lemma less_number_of_eq_neg:
   13.55 -    "((number_of x::'a::{ordered_ring,number_ring}) < number_of y)
   13.56 +    "((number_of x::'a::{ordered_idom,number_ring}) < number_of y)
   13.57       = neg (number_of (bin_add x (bin_minus y)) :: 'a)"
   13.58  apply (subst less_iff_diff_less_0) 
   13.59  apply (simp add: neg_def diff_minus number_of_add number_of_minus)
   13.60 @@ -202,14 +202,14 @@
   13.61  text{*If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
   13.62    @{term Numeral0} IS @{term "number_of Pls"} *}
   13.63  lemma not_neg_number_of_Pls:
   13.64 -     "~ neg (number_of bin.Pls ::'a::{ordered_ring,number_ring})"
   13.65 +     "~ neg (number_of bin.Pls ::'a::{ordered_idom,number_ring})"
   13.66  by (simp add: neg_def numeral_0_eq_0)
   13.67  
   13.68  lemma neg_number_of_Min:
   13.69 -     "neg (number_of bin.Min ::'a::{ordered_ring,number_ring})"
   13.70 +     "neg (number_of bin.Min ::'a::{ordered_idom,number_ring})"
   13.71  by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1)
   13.72  
   13.73 -lemma double_less_0_iff: "(a + a < 0) = (a < (0::'a::ordered_ring))"
   13.74 +lemma double_less_0_iff: "(a + a < 0) = (a < (0::'a::ordered_idom))"
   13.75  proof -
   13.76    have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib)
   13.77    also have "... = (a < 0)"
   13.78 @@ -231,7 +231,7 @@
   13.79  
   13.80  text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
   13.81  lemma Ints_odd_less_0: 
   13.82 -     "a \<in> Ints ==> (1 + a + a < 0) = (a < (0::'a::ordered_ring))";
   13.83 +     "a \<in> Ints ==> (1 + a + a < 0) = (a < (0::'a::ordered_idom))";
   13.84  proof (unfold Ints_def) 
   13.85    assume "a \<in> range of_int"
   13.86    then obtain z where a: "a = of_int z" ..
   13.87 @@ -244,7 +244,7 @@
   13.88  
   13.89  lemma neg_number_of_BIT:
   13.90       "neg (number_of (w BIT x)::'a) = 
   13.91 -      neg (number_of w :: 'a::{ordered_ring,number_ring})"
   13.92 +      neg (number_of w :: 'a::{ordered_idom,number_ring})"
   13.93  by (simp add: number_of neg_def double_less_0_iff
   13.94                Ints_odd_less_0 [OF Ints_number_of])
   13.95  
   13.96 @@ -257,7 +257,7 @@
   13.97                            standard]
   13.98  
   13.99  lemma le_number_of_eq:
  13.100 -    "((number_of x::'a::{ordered_ring,number_ring}) \<le> number_of y)
  13.101 +    "((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y)
  13.102       = (~ (neg (number_of (bin_add y (bin_minus x)) :: 'a)))"
  13.103  by (simp add: le_number_of_eq_not_less less_number_of_eq_neg)
  13.104  
  13.105 @@ -265,7 +265,7 @@
  13.106  text{*Absolute value (@{term abs})*}
  13.107  
  13.108  lemma abs_number_of:
  13.109 -     "abs(number_of x::'a::{ordered_ring,number_ring}) =
  13.110 +     "abs(number_of x::'a::{ordered_idom,number_ring}) =
  13.111        (if number_of x < (0::'a) then -number_of x else number_of x)"
  13.112  by (simp add: abs_if)
  13.113  
    14.1 --- a/src/HOL/Integ/IntArith.thy	Tue May 11 14:00:02 2004 +0200
    14.2 +++ b/src/HOL/Integ/IntArith.thy	Tue May 11 20:11:08 2004 +0200
    14.3 @@ -126,11 +126,11 @@
    14.4    finally show ?thesis .
    14.5  qed
    14.6  
    14.7 -lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_ring,number_ring})"
    14.8 +lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})"
    14.9  by (simp add: abs_if)
   14.10  
   14.11  lemma abs_power_minus_one [simp]:
   14.12 -     "abs(-1 ^ n) = (1::'a::{ordered_ring,number_ring,ringpower})"
   14.13 +     "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,ringpower})"
   14.14  by (simp add: power_abs)
   14.15  
   14.16  lemma of_int_number_of_eq:
    15.1 --- a/src/HOL/Integ/IntDef.thy	Tue May 11 14:00:02 2004 +0200
    15.2 +++ b/src/HOL/Integ/IntDef.thy	Tue May 11 20:11:08 2004 +0200
    15.3 @@ -153,7 +153,7 @@
    15.4  
    15.5  lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
    15.6  
    15.7 -lemmas zmult_ac = Ring_and_Field.mult_ac
    15.8 +lemmas zmult_ac = OrderedGroup.mult_ac
    15.9  
   15.10  lemma zadd_int: "(int m) + (int n) = int (m + n)"
   15.11  by (simp add: int_def add)
   15.12 @@ -164,7 +164,7 @@
   15.13  lemma int_Suc: "int (Suc m) = 1 + (int m)"
   15.14  by (simp add: One_int_def zadd_int)
   15.15  
   15.16 -(*also for the instance declaration int :: plus_ac0*)
   15.17 +(*also for the instance declaration int :: comm_monoid_add*)
   15.18  lemma zadd_0: "(0::int) + z = z"
   15.19  apply (simp add: Zero_int_def int_def)
   15.20  apply (cases z, simp add: add)
   15.21 @@ -235,8 +235,8 @@
   15.22  by (rule trans [OF zmult_commute zmult_1])
   15.23  
   15.24  
   15.25 -text{*The Integers Form A Ring*}
   15.26 -instance int :: ring
   15.27 +text{*The Integers Form A comm_ring_1*}
   15.28 +instance int :: comm_ring_1
   15.29  proof
   15.30    fix i j k :: int
   15.31    show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
   15.32 @@ -368,8 +368,8 @@
   15.33      zabs_def:  "abs(i::int) == if i < 0 then -i else i"
   15.34  
   15.35  
   15.36 -text{*The Integers Form an Ordered Ring*}
   15.37 -instance int :: ordered_ring
   15.38 +text{*The Integers Form an Ordered comm_ring_1*}
   15.39 +instance int :: ordered_idom
   15.40  proof
   15.41    fix i j k :: int
   15.42    show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono)
   15.43 @@ -510,7 +510,7 @@
   15.44        in theory @{text Ring_and_Field}.
   15.45        But is it really better than just rewriting with @{text abs_if}?*}
   15.46  lemma abs_split [arith_split]:
   15.47 -     "P(abs(a::'a::ordered_ring)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
   15.48 +     "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
   15.49  by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
   15.50  
   15.51  
   15.52 @@ -519,11 +519,11 @@
   15.53  
   15.54  constdefs
   15.55  
   15.56 -  neg   :: "'a::ordered_ring => bool"
   15.57 +  neg   :: "'a::ordered_idom => bool"
   15.58    "neg(Z) == Z < 0"
   15.59  
   15.60    (*For simplifying equalities*)
   15.61 -  iszero :: "'a::semiring => bool"
   15.62 +  iszero :: "'a::comm_semiring_1_cancel => bool"
   15.63    "iszero z == z = (0)"
   15.64  
   15.65  
   15.66 @@ -560,9 +560,9 @@
   15.67  by (simp add: linorder_not_less neg_def)
   15.68  
   15.69  
   15.70 -subsection{*Embedding of the Naturals into any Semiring: @{term of_nat}*}
   15.71 +subsection{*Embedding of the Naturals into any comm_semiring_1_cancel: @{term of_nat}*}
   15.72  
   15.73 -consts of_nat :: "nat => 'a::semiring"
   15.74 +consts of_nat :: "nat => 'a::comm_semiring_1_cancel"
   15.75  
   15.76  primrec
   15.77    of_nat_0:   "of_nat 0 = 0"
   15.78 @@ -581,27 +581,27 @@
   15.79  apply (simp_all add: mult_ac add_ac right_distrib)
   15.80  done
   15.81  
   15.82 -lemma zero_le_imp_of_nat: "0 \<le> (of_nat m::'a::ordered_semiring)"
   15.83 +lemma zero_le_imp_of_nat: "0 \<le> (of_nat m::'a::ordered_semidom)"
   15.84  apply (induct m, simp_all)
   15.85  apply (erule order_trans)
   15.86  apply (rule less_add_one [THEN order_less_imp_le])
   15.87  done
   15.88  
   15.89  lemma less_imp_of_nat_less:
   15.90 -     "m < n ==> of_nat m < (of_nat n::'a::ordered_semiring)"
   15.91 +     "m < n ==> of_nat m < (of_nat n::'a::ordered_semidom)"
   15.92  apply (induct m n rule: diff_induct, simp_all)
   15.93  apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force)
   15.94  done
   15.95  
   15.96  lemma of_nat_less_imp_less:
   15.97 -     "of_nat m < (of_nat n::'a::ordered_semiring) ==> m < n"
   15.98 +     "of_nat m < (of_nat n::'a::ordered_semidom) ==> m < n"
   15.99  apply (induct m n rule: diff_induct, simp_all)
  15.100  apply (insert zero_le_imp_of_nat)
  15.101  apply (force simp add: linorder_not_less [symmetric])
  15.102  done
  15.103  
  15.104  lemma of_nat_less_iff [simp]:
  15.105 -     "(of_nat m < (of_nat n::'a::ordered_semiring)) = (m<n)"
  15.106 +     "(of_nat m < (of_nat n::'a::ordered_semidom)) = (m<n)"
  15.107  by (blast intro: of_nat_less_imp_less less_imp_of_nat_less)
  15.108  
  15.109  text{*Special cases where either operand is zero*}
  15.110 @@ -609,17 +609,17 @@
  15.111  declare of_nat_less_iff [of _ 0, simplified, simp]
  15.112  
  15.113  lemma of_nat_le_iff [simp]:
  15.114 -     "(of_nat m \<le> (of_nat n::'a::ordered_semiring)) = (m \<le> n)"
  15.115 +     "(of_nat m \<le> (of_nat n::'a::ordered_semidom)) = (m \<le> n)"
  15.116  by (simp add: linorder_not_less [symmetric])
  15.117  
  15.118  text{*Special cases where either operand is zero*}
  15.119  declare of_nat_le_iff [of 0, simplified, simp]
  15.120  declare of_nat_le_iff [of _ 0, simplified, simp]
  15.121  
  15.122 -text{*The ordering on the semiring is necessary to exclude the possibility of
  15.123 +text{*The ordering on the comm_semiring_1_cancel is necessary to exclude the possibility of
  15.124  a finite field, which indeed wraps back to zero.*}
  15.125  lemma of_nat_eq_iff [simp]:
  15.126 -     "(of_nat m = (of_nat n::'a::ordered_semiring)) = (m = n)"
  15.127 +     "(of_nat m = (of_nat n::'a::ordered_semidom)) = (m = n)"
  15.128  by (simp add: order_eq_iff)
  15.129  
  15.130  text{*Special cases where either operand is zero*}
  15.131 @@ -627,7 +627,7 @@
  15.132  declare of_nat_eq_iff [of _ 0, simplified, simp]
  15.133  
  15.134  lemma of_nat_diff [simp]:
  15.135 -     "n \<le> m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::ring)"
  15.136 +     "n \<le> m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::comm_ring_1)"
  15.137  by (simp del: of_nat_add
  15.138  	 add: compare_rls of_nat_add [symmetric] split add: nat_diff_split)
  15.139  
  15.140 @@ -635,7 +635,7 @@
  15.141  subsection{*The Set of Natural Numbers*}
  15.142  
  15.143  constdefs
  15.144 -   Nats  :: "'a::semiring set"
  15.145 +   Nats  :: "'a::comm_semiring_1_cancel set"
  15.146      "Nats == range of_nat"
  15.147  
  15.148  syntax (xsymbols)    Nats :: "'a set"   ("\<nat>")
  15.149 @@ -681,10 +681,10 @@
  15.150  qed
  15.151  
  15.152  
  15.153 -subsection{*Embedding of the Integers into any Ring: @{term of_int}*}
  15.154 +subsection{*Embedding of the Integers into any comm_ring_1: @{term of_int}*}
  15.155  
  15.156  constdefs
  15.157 -   of_int :: "int => 'a::ring"
  15.158 +   of_int :: "int => 'a::comm_ring_1"
  15.159     "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
  15.160  
  15.161  
  15.162 @@ -719,7 +719,7 @@
  15.163  done
  15.164  
  15.165  lemma of_int_le_iff [simp]:
  15.166 -     "(of_int w \<le> (of_int z::'a::ordered_ring)) = (w \<le> z)"
  15.167 +     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
  15.168  apply (cases w)
  15.169  apply (cases z)
  15.170  apply (simp add: compare_rls of_int le diff_int_def add minus
  15.171 @@ -731,16 +731,16 @@
  15.172  declare of_int_le_iff [of _ 0, simplified, simp]
  15.173  
  15.174  lemma of_int_less_iff [simp]:
  15.175 -     "(of_int w < (of_int z::'a::ordered_ring)) = (w < z)"
  15.176 +     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
  15.177  by (simp add: linorder_not_le [symmetric])
  15.178  
  15.179  text{*Special cases where either operand is zero*}
  15.180  declare of_int_less_iff [of 0, simplified, simp]
  15.181  declare of_int_less_iff [of _ 0, simplified, simp]
  15.182  
  15.183 -text{*The ordering on the ring is necessary. See @{text of_nat_eq_iff} above.*}
  15.184 +text{*The ordering on the comm_ring_1 is necessary. See @{text of_nat_eq_iff} above.*}
  15.185  lemma of_int_eq_iff [simp]:
  15.186 -     "(of_int w = (of_int z::'a::ordered_ring)) = (w = z)"
  15.187 +     "(of_int w = (of_int z::'a::ordered_idom)) = (w = z)"
  15.188  by (simp add: order_eq_iff)
  15.189  
  15.190  text{*Special cases where either operand is zero*}
  15.191 @@ -759,7 +759,7 @@
  15.192  subsection{*The Set of Integers*}
  15.193  
  15.194  constdefs
  15.195 -   Ints  :: "'a::ring set"
  15.196 +   Ints  :: "'a::comm_ring_1 set"
  15.197      "Ints == range of_int"
  15.198  
  15.199  
    16.1 --- a/src/HOL/Integ/NatBin.thy	Tue May 11 14:00:02 2004 +0200
    16.2 +++ b/src/HOL/Integ/NatBin.thy	Tue May 11 20:11:08 2004 +0200
    16.3 @@ -256,42 +256,42 @@
    16.4  We cannot prove general results about the numeral @{term "-1"}, so we have to
    16.5  use @{term "- 1"} instead.*}
    16.6  
    16.7 -lemma power2_eq_square: "(a::'a::{semiring,ringpower})\<twosuperior> = a * a"
    16.8 +lemma power2_eq_square: "(a::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = a * a"
    16.9    by (simp add: numeral_2_eq_2 Power.power_Suc)
   16.10  
   16.11 -lemma [simp]: "(0::'a::{semiring,ringpower})\<twosuperior> = 0"
   16.12 +lemma [simp]: "(0::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = 0"
   16.13    by (simp add: power2_eq_square)
   16.14  
   16.15 -lemma [simp]: "(1::'a::{semiring,ringpower})\<twosuperior> = 1"
   16.16 +lemma [simp]: "(1::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = 1"
   16.17    by (simp add: power2_eq_square)
   16.18  
   16.19  text{*Squares of literal numerals will be evaluated.*}
   16.20  declare power2_eq_square [of "number_of w", standard, simp]
   16.21  
   16.22 -lemma zero_le_power2 [simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_ring,ringpower})"
   16.23 +lemma zero_le_power2 [simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_idom,ringpower})"
   16.24    by (simp add: power2_eq_square zero_le_square)
   16.25  
   16.26  lemma zero_less_power2 [simp]:
   16.27 -     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_ring,ringpower}))"
   16.28 +     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_idom,ringpower}))"
   16.29    by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
   16.30  
   16.31  lemma zero_eq_power2 [simp]:
   16.32 -     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_ring,ringpower}))"
   16.33 +     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_idom,ringpower}))"
   16.34    by (force simp add: power2_eq_square mult_eq_0_iff)
   16.35  
   16.36  lemma abs_power2 [simp]:
   16.37 -     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_ring,ringpower})"
   16.38 +     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_idom,ringpower})"
   16.39    by (simp add: power2_eq_square abs_mult abs_mult_self)
   16.40  
   16.41  lemma power2_abs [simp]:
   16.42 -     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_ring,ringpower})"
   16.43 +     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_idom,ringpower})"
   16.44    by (simp add: power2_eq_square abs_mult_self)
   16.45  
   16.46  lemma power2_minus [simp]:
   16.47 -     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{ring,ringpower})"
   16.48 +     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{comm_ring_1,ringpower})"
   16.49    by (simp add: power2_eq_square)
   16.50  
   16.51 -lemma power_minus1_even: "(- 1) ^ (2*n) = (1::'a::{ring,ringpower})"
   16.52 +lemma power_minus1_even: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,ringpower})"
   16.53  apply (induct_tac "n")
   16.54  apply (auto simp add: power_Suc power_add)
   16.55  done
   16.56 @@ -303,11 +303,11 @@
   16.57  by (simp add: power_even_eq) 
   16.58  
   16.59  lemma power_minus_even [simp]:
   16.60 -     "(-a) ^ (2*n) = (a::'a::{ring,ringpower}) ^ (2*n)"
   16.61 +     "(-a) ^ (2*n) = (a::'a::{comm_ring_1,ringpower}) ^ (2*n)"
   16.62  by (simp add: power_minus1_even power_minus [of a]) 
   16.63  
   16.64  lemma zero_le_even_power:
   16.65 -     "0 \<le> (a::'a::{ordered_ring,ringpower}) ^ (2*n)"
   16.66 +     "0 \<le> (a::'a::{ordered_idom,ringpower}) ^ (2*n)"
   16.67  proof (induct "n")
   16.68    case 0
   16.69      show ?case by (simp add: zero_le_one)
   16.70 @@ -320,7 +320,7 @@
   16.71  qed
   16.72  
   16.73  lemma odd_power_less_zero:
   16.74 -     "(a::'a::{ordered_ring,ringpower}) < 0 ==> a ^ Suc(2*n) < 0"
   16.75 +     "(a::'a::{ordered_idom,ringpower}) < 0 ==> a ^ Suc(2*n) < 0"
   16.76  proof (induct "n")
   16.77    case 0
   16.78      show ?case by (simp add: Power.power_Suc)
   16.79 @@ -333,7 +333,7 @@
   16.80  qed
   16.81  
   16.82  lemma odd_0_le_power_imp_0_le:
   16.83 -     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_ring,ringpower})"
   16.84 +     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_idom,ringpower})"
   16.85  apply (insert odd_power_less_zero [of a n]) 
   16.86  apply (force simp add: linorder_not_less [symmetric]) 
   16.87  done
   16.88 @@ -473,7 +473,7 @@
   16.89       "((number_of (v BIT x) ::int) = number_of (w BIT y)) =  
   16.90        (x=y & (((number_of v) ::int) = number_of w))"
   16.91  by (simp only: simp_thms number_of_BIT lemma1 lemma2 eq_commute
   16.92 -               Ring_and_Field.add_left_cancel add_assoc Ring_and_Field.add_0
   16.93 +               OrderedGroup.add_left_cancel add_assoc OrderedGroup.add_0
   16.94              split add: split_if cong: imp_cong) 
   16.95  
   16.96  
    17.1 --- a/src/HOL/Integ/Parity.thy	Tue May 11 14:00:02 2004 +0200
    17.2 +++ b/src/HOL/Integ/Parity.thy	Tue May 11 20:11:08 2004 +0200
    17.3 @@ -249,7 +249,7 @@
    17.4    by (rule neg_one_even_odd_power [THEN conjunct2, THEN mp], assumption)
    17.5  
    17.6  lemma neg_power_if:
    17.7 -     "(-x::'a::{ring,ringpower}) ^ n = 
    17.8 +     "(-x::'a::{comm_ring_1,ringpower}) ^ n = 
    17.9        (if even n then (x ^ n) else -(x ^ n))"
   17.10    by (induct n, simp_all split: split_if_asm add: power_Suc) 
   17.11  
   17.12 @@ -257,13 +257,13 @@
   17.13  subsection {* An Equivalence for @{term "0 \<le> a^n"} *}
   17.14  
   17.15  lemma even_power_le_0_imp_0:
   17.16 -     "a ^ (2*k) \<le> (0::'a::{ordered_ring,ringpower}) ==> a=0"
   17.17 +     "a ^ (2*k) \<le> (0::'a::{ordered_idom,ringpower}) ==> a=0"
   17.18  apply (induct k) 
   17.19  apply (auto simp add: zero_le_mult_iff mult_le_0_iff power_Suc)  
   17.20  done
   17.21  
   17.22  lemma zero_le_power_iff:
   17.23 -     "(0 \<le> a^n) = (0 \<le> (a::'a::{ordered_ring,ringpower}) | even n)"
   17.24 +     "(0 \<le> a^n) = (0 \<le> (a::'a::{ordered_idom,ringpower}) | even n)"
   17.25        (is "?P n")
   17.26  proof cases
   17.27    assume even: "even n"
    18.1 --- a/src/HOL/Integ/Presburger.thy	Tue May 11 14:00:02 2004 +0200
    18.2 +++ b/src/HOL/Integ/Presburger.thy	Tue May 11 20:11:08 2004 +0200
    18.3 @@ -704,7 +704,7 @@
    18.4      have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast
    18.5      also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)"
    18.6        using minus[THEN spec, of "x - i * d"]
    18.7 -      by (simp add:int_distrib Ring_and_Field.diff_diff_eq[symmetric])
    18.8 +      by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric])
    18.9      ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast
   18.10    qed
   18.11  qed
    19.1 --- a/src/HOL/Integ/int_arith1.ML	Tue May 11 14:00:02 2004 +0200
    19.2 +++ b/src/HOL/Integ/int_arith1.ML	Tue May 11 20:11:08 2004 +0200
    19.3 @@ -383,20 +383,20 @@
    19.4        "(l::'a::number_ring) = m * n"],
    19.5       EqCancelNumerals.proc),
    19.6      ("intless_cancel_numerals",
    19.7 -     ["(l::'a::{ordered_ring,number_ring}) + m < n",
    19.8 -      "(l::'a::{ordered_ring,number_ring}) < m + n",
    19.9 -      "(l::'a::{ordered_ring,number_ring}) - m < n",
   19.10 -      "(l::'a::{ordered_ring,number_ring}) < m - n",
   19.11 -      "(l::'a::{ordered_ring,number_ring}) * m < n",
   19.12 -      "(l::'a::{ordered_ring,number_ring}) < m * n"],
   19.13 +     ["(l::'a::{ordered_idom,number_ring}) + m < n",
   19.14 +      "(l::'a::{ordered_idom,number_ring}) < m + n",
   19.15 +      "(l::'a::{ordered_idom,number_ring}) - m < n",
   19.16 +      "(l::'a::{ordered_idom,number_ring}) < m - n",
   19.17 +      "(l::'a::{ordered_idom,number_ring}) * m < n",
   19.18 +      "(l::'a::{ordered_idom,number_ring}) < m * n"],
   19.19       LessCancelNumerals.proc),
   19.20      ("intle_cancel_numerals",
   19.21 -     ["(l::'a::{ordered_ring,number_ring}) + m <= n",
   19.22 -      "(l::'a::{ordered_ring,number_ring}) <= m + n",
   19.23 -      "(l::'a::{ordered_ring,number_ring}) - m <= n",
   19.24 -      "(l::'a::{ordered_ring,number_ring}) <= m - n",
   19.25 -      "(l::'a::{ordered_ring,number_ring}) * m <= n",
   19.26 -      "(l::'a::{ordered_ring,number_ring}) <= m * n"],
   19.27 +     ["(l::'a::{ordered_idom,number_ring}) + m <= n",
   19.28 +      "(l::'a::{ordered_idom,number_ring}) <= m + n",
   19.29 +      "(l::'a::{ordered_idom,number_ring}) - m <= n",
   19.30 +      "(l::'a::{ordered_idom,number_ring}) <= m - n",
   19.31 +      "(l::'a::{ordered_idom,number_ring}) * m <= n",
   19.32 +      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
   19.33       LeCancelNumerals.proc)];
   19.34  
   19.35  
   19.36 @@ -488,7 +488,7 @@
   19.37  
   19.38  val assoc_fold_simproc =
   19.39    Bin_Simprocs.prep_simproc
   19.40 -   ("semiring_assoc_fold", ["(a::'a::semiring) * b"],
   19.41 +   ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
   19.42      Semiring_Times_Assoc.proc);
   19.43  
   19.44  Addsimprocs [assoc_fold_simproc];
   19.45 @@ -545,9 +545,9 @@
   19.46  val fast_int_arith_simproc =
   19.47    Simplifier.simproc (Theory.sign_of (the_context()))
   19.48    "fast_int_arith" 
   19.49 -     ["(m::'a::{ordered_ring,number_ring}) < n",
   19.50 -      "(m::'a::{ordered_ring,number_ring}) <= n",
   19.51 -      "(m::'a::{ordered_ring,number_ring}) = n"] Fast_Arith.lin_arith_prover;
   19.52 +     ["(m::'a::{ordered_idom,number_ring}) < n",
   19.53 +      "(m::'a::{ordered_idom,number_ring}) <= n",
   19.54 +      "(m::'a::{ordered_idom,number_ring}) = n"] Fast_Arith.lin_arith_prover;
   19.55  
   19.56  Addsimprocs [fast_int_arith_simproc]
   19.57  
    20.1 --- a/src/HOL/Integ/int_factor_simprocs.ML	Tue May 11 14:00:02 2004 +0200
    20.2 +++ b/src/HOL/Integ/int_factor_simprocs.ML	Tue May 11 20:11:08 2004 +0200
    20.3 @@ -116,16 +116,16 @@
    20.4  val ring_cancel_numeral_factors =
    20.5    map Bin_Simprocs.prep_simproc
    20.6     [("ring_eq_cancel_numeral_factor",
    20.7 -     ["(l::'a::{ordered_ring,number_ring}) * m = n",
    20.8 -      "(l::'a::{ordered_ring,number_ring}) = m * n"],
    20.9 +     ["(l::'a::{ordered_idom,number_ring}) * m = n",
   20.10 +      "(l::'a::{ordered_idom,number_ring}) = m * n"],
   20.11       EqCancelNumeralFactor.proc),
   20.12      ("ring_less_cancel_numeral_factor",
   20.13 -     ["(l::'a::{ordered_ring,number_ring}) * m < n",
   20.14 -      "(l::'a::{ordered_ring,number_ring}) < m * n"],
   20.15 +     ["(l::'a::{ordered_idom,number_ring}) * m < n",
   20.16 +      "(l::'a::{ordered_idom,number_ring}) < m * n"],
   20.17       LessCancelNumeralFactor.proc),
   20.18      ("ring_le_cancel_numeral_factor",
   20.19 -     ["(l::'a::{ordered_ring,number_ring}) * m <= n",
   20.20 -      "(l::'a::{ordered_ring,number_ring}) <= m * n"],
   20.21 +     ["(l::'a::{ordered_idom,number_ring}) * m <= n",
   20.22 +      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
   20.23       LeCancelNumeralFactor.proc),
   20.24      ("int_div_cancel_numeral_factors",
   20.25       ["((l::int) * m) div n", "(l::int) div (m * n)"],
   20.26 @@ -249,7 +249,7 @@
   20.27    val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
   20.28    end;
   20.29  
   20.30 -(*mult_cancel_left requires an ordered ring, such as int. The version in
   20.31 +(*mult_cancel_left requires an ordered comm_ring_1, such as int. The version in
   20.32    rat_arith.ML works for all fields.*)
   20.33  structure EqCancelFactor = ExtractCommonTermFun
   20.34   (open CancelFactorCommon
    21.1 --- a/src/HOL/IsaMakefile	Tue May 11 14:00:02 2004 +0200
    21.2 +++ b/src/HOL/IsaMakefile	Tue May 11 20:11:08 2004 +0200
    21.3 @@ -151,7 +151,7 @@
    21.4    Hyperreal/HyperDef.thy Hyperreal/HyperNat.thy\
    21.5    Hyperreal/HyperPow.thy Hyperreal/Hyperreal.thy\
    21.6    Hyperreal/Lim.thy Hyperreal/Log.thy\
    21.7 -  Hyperreal/MacLaurin.ML Hyperreal/MacLaurin.thy Hyperreal/NatStar.thy\
    21.8 +  Hyperreal/MacLaurin_lemmas.ML Hyperreal/MacLaurin.thy Hyperreal/NatStar.thy\
    21.9    Hyperreal/NSA.thy Hyperreal/NthRoot.thy Hyperreal/Poly.thy\
   21.10    Hyperreal/SEQ.ML Hyperreal/SEQ.thy Hyperreal/Series.thy\
   21.11    Hyperreal/Star.thy Hyperreal/Transcendental.ML\
    22.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    22.2 +++ b/src/HOL/LOrder.thy	Tue May 11 20:11:08 2004 +0200
    22.3 @@ -0,0 +1,331 @@
    22.4 +(*  Title:   HOL/LOrder.thy
    22.5 +    ID:      $Id$
    22.6 +    Author:  Steven Obua, TU Muenchen
    22.7 +    License: GPL (GNU GENERAL PUBLIC LICENSE)
    22.8 +*)
    22.9 +
   22.10 +header {* Lattice Orders *}
   22.11 +
   22.12 +theory LOrder = HOL:
   22.13 +
   22.14 +text {*
   22.15 +  The theory of lattices developed here is taken from the book:
   22.16 +  \begin{itemize}
   22.17 +  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979. 
   22.18 +  \end{itemize}
   22.19 +*}
   22.20 +
   22.21 +constdefs
   22.22 +  is_meet :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
   22.23 +  "is_meet m == ! a b x. m a b \<le> a \<and> m a b \<le> b \<and> (x \<le> a \<and> x \<le> b \<longrightarrow> x \<le> m a b)"
   22.24 +  is_join :: "(('a::order) \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
   22.25 +  "is_join j == ! a b x. a \<le> j a b \<and> b \<le> j a b \<and> (a \<le> x \<and> b \<le> x \<longrightarrow> j a b \<le> x)"  
   22.26 +
   22.27 +lemma is_meet_unique: 
   22.28 +  assumes "is_meet u" "is_meet v" shows "u = v"
   22.29 +proof -
   22.30 +  {
   22.31 +    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
   22.32 +    assume a: "is_meet a"
   22.33 +    assume b: "is_meet b"
   22.34 +    {
   22.35 +      fix x y 
   22.36 +      let ?za = "a x y"
   22.37 +      let ?zb = "b x y"
   22.38 +      from a have za_le: "?za <= x & ?za <= y" by (auto simp add: is_meet_def)
   22.39 +      with b have "?za <= ?zb" by (auto simp add: is_meet_def)
   22.40 +    }
   22.41 +  }
   22.42 +  note f_le = this
   22.43 +  show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
   22.44 +qed
   22.45 +
   22.46 +lemma is_join_unique: 
   22.47 +  assumes "is_join u" "is_join v" shows "u = v"
   22.48 +proof -
   22.49 +  {
   22.50 +    fix a b :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
   22.51 +    assume a: "is_join a"
   22.52 +    assume b: "is_join b"
   22.53 +    {
   22.54 +      fix x y 
   22.55 +      let ?za = "a x y"
   22.56 +      let ?zb = "b x y"
   22.57 +      from a have za_le: "x <= ?za & y <= ?za" by (auto simp add: is_join_def)
   22.58 +      with b have "?zb <= ?za" by (auto simp add: is_join_def)
   22.59 +    }
   22.60 +  }
   22.61 +  note f_le = this
   22.62 +  show "u = v" by ((rule ext)+, simp_all add: order_antisym prems f_le) 
   22.63 +qed
   22.64 +
   22.65 +axclass join_semilorder < order
   22.66 +  join_exists: "? j. is_join j"
   22.67 +
   22.68 +axclass meet_semilorder < order
   22.69 +  meet_exists: "? m. is_meet m"
   22.70 +
   22.71 +axclass lorder < join_semilorder, meet_semilorder
   22.72 +
   22.73 +constdefs
   22.74 +  meet :: "('a::meet_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
   22.75 +  "meet == THE m. is_meet m"
   22.76 +  join :: "('a::join_semilorder) \<Rightarrow> 'a \<Rightarrow> 'a"
   22.77 +  "join ==  THE j. is_join j"
   22.78 +
   22.79 +lemma is_meet_meet: "is_meet (meet::'a \<Rightarrow> 'a \<Rightarrow> ('a::meet_semilorder))"
   22.80 +proof -
   22.81 +  from meet_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_meet k" ..
   22.82 +  with is_meet_unique[of _ k] show ?thesis
   22.83 +    by (simp add: meet_def theI[of is_meet])    
   22.84 +qed
   22.85 +
   22.86 +lemma meet_unique: "(is_meet m) = (m = meet)" 
   22.87 +by (insert is_meet_meet, auto simp add: is_meet_unique)
   22.88 +
   22.89 +lemma is_join_join: "is_join (join::'a \<Rightarrow> 'a \<Rightarrow> ('a::join_semilorder))"
   22.90 +proof -
   22.91 +  from join_exists obtain k::"'a \<Rightarrow> 'a \<Rightarrow> 'a" where "is_join k" ..
   22.92 +  with is_join_unique[of _ k] show ?thesis
   22.93 +    by (simp add: join_def theI[of is_join])    
   22.94 +qed
   22.95 +
   22.96 +lemma join_unique: "(is_join j) = (j = join)"
   22.97 +by (insert is_join_join, auto simp add: is_join_unique)
   22.98 +
   22.99 +lemma meet_left_le: "meet a b \<le> (a::'a::meet_semilorder)"
  22.100 +by (insert is_meet_meet, auto simp add: is_meet_def)
  22.101 +
  22.102 +lemma meet_right_le: "meet a b \<le> (b::'a::meet_semilorder)"
  22.103 +by (insert is_meet_meet, auto simp add: is_meet_def)
  22.104 +
  22.105 +lemma meet_imp_le: "x \<le> a \<Longrightarrow> x \<le> b \<Longrightarrow> x \<le> meet a (b::'a::meet_semilorder)"
  22.106 +by (insert is_meet_meet, auto simp add: is_meet_def)
  22.107 +
  22.108 +lemma join_left_le: "a \<le> join a (b::'a::join_semilorder)"
  22.109 +by (insert is_join_join, auto simp add: is_join_def)
  22.110 +
  22.111 +lemma join_right_le: "b \<le> join a (b::'a::join_semilorder)"
  22.112 +by (insert is_join_join, auto simp add: is_join_def)
  22.113 +
  22.114 +lemma join_imp_le: "a \<le> x \<Longrightarrow> b \<le> x \<Longrightarrow> join a b \<le> (x::'a::join_semilorder)"
  22.115 +by (insert is_join_join, auto simp add: is_join_def)
  22.116 +
  22.117 +lemmas meet_join_le = meet_left_le meet_right_le join_left_le join_right_le
  22.118 +
  22.119 +lemma is_meet_min: "is_meet (min::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
  22.120 +by (auto simp add: is_meet_def min_def)
  22.121 +
  22.122 +lemma is_join_max: "is_join (max::'a \<Rightarrow> 'a \<Rightarrow> ('a::linorder))"
  22.123 +by (auto simp add: is_join_def max_def)
  22.124 +
  22.125 +instance linorder \<subseteq> meet_semilorder
  22.126 +proof
  22.127 +  from is_meet_min show "? (m::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_meet m" by auto
  22.128 +qed
  22.129 +
  22.130 +instance linorder \<subseteq> join_semilorder
  22.131 +proof
  22.132 +  from is_join_max show "? (j::'a\<Rightarrow>'a\<Rightarrow>('a::linorder)). is_join j" by auto 
  22.133 +qed
  22.134 +    
  22.135 +instance linorder \<subseteq> lorder ..
  22.136 +
  22.137 +lemma meet_min: "meet = (min :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))" 
  22.138 +by (simp add: is_meet_meet is_meet_min is_meet_unique)
  22.139 +
  22.140 +lemma join_max: "join = (max :: 'a\<Rightarrow>'a\<Rightarrow>('a::linorder))"
  22.141 +by (simp add: is_join_join is_join_max is_join_unique)
  22.142 +
  22.143 +lemma meet_idempotent[simp]: "meet x x = x"
  22.144 +by (rule order_antisym, simp_all add: meet_left_le meet_imp_le)
  22.145 +
  22.146 +lemma join_idempotent[simp]: "join x x = x"
  22.147 +by (rule order_antisym, simp_all add: join_left_le join_imp_le)
  22.148 +
  22.149 +lemma meet_comm: "meet x y = meet y x" 
  22.150 +by (rule order_antisym, (simp add: meet_left_le meet_right_le meet_imp_le)+)
  22.151 +
  22.152 +lemma join_comm: "join x y = join y x"
  22.153 +by (rule order_antisym, (simp add: join_right_le join_left_le join_imp_le)+)
  22.154 +
  22.155 +lemma meet_assoc: "meet (meet x y) z = meet x (meet y z)" (is "?l=?r")
  22.156 +proof - 
  22.157 +  have "?l <= meet x y & meet x y <= x & ?l <= z & meet x y <= y" by (simp add: meet_left_le meet_right_le)
  22.158 +  hence "?l <= x & ?l <= y & ?l <= z" by auto
  22.159 +  hence "?l <= ?r" by (simp add: meet_imp_le)
  22.160 +  hence a:"?l <= meet x (meet y z)" by (simp add: meet_imp_le)
  22.161 +  have "?r <= meet y z & meet y z <= y & meet y z <= z & ?r <= x" by (simp add: meet_left_le meet_right_le)  
  22.162 +  hence "?r <= x & ?r <= y & ?r <= z" by (auto) 
  22.163 +  hence "?r <= meet x y & ?r <= z" by (simp add: meet_imp_le)
  22.164 +  hence b:"?r <= ?l" by (simp add: meet_imp_le)
  22.165 +  from a b show "?l = ?r" by auto
  22.166 +qed
  22.167 +
  22.168 +lemma join_assoc: "join (join x y) z = join x (join y z)" (is "?l=?r")
  22.169 +proof -
  22.170 +  have "join x y <= ?l & x <= join x y & z <= ?l & y <= join x y" by (simp add: join_left_le join_right_le)
  22.171 +  hence "x <= ?l & y <= ?l & z <= ?l" by auto
  22.172 +  hence "join y z <= ?l & x <= ?l" by (simp add: join_imp_le)
  22.173 +  hence a:"?r <= ?l" by (simp add: join_imp_le)
  22.174 +  have "join y z <= ?r & y <= join y z & z <= join y z & x <= ?r" by (simp add: join_left_le join_right_le)
  22.175 +  hence "y <= ?r & z <= ?r & x <= ?r" by auto
  22.176 +  hence "join x y <= ?r & z <= ?r" by (simp add: join_imp_le)
  22.177 +  hence b:"?l <= ?r" by (simp add: join_imp_le)
  22.178 +  from a b show "?l = ?r" by auto
  22.179 +qed
  22.180 +
  22.181 +lemma meet_left_comm: "meet a (meet b c) = meet b (meet a c)"
  22.182 +by (simp add: meet_assoc[symmetric, of a b c], simp add: meet_comm[of a b], simp add: meet_assoc)
  22.183 +
  22.184 +lemma meet_left_idempotent: "meet y (meet y x) = meet y x"
  22.185 +by (simp add: meet_assoc meet_comm meet_left_comm)
  22.186 +
  22.187 +lemma join_left_comm: "join a (join b c) = join b (join a c)"
  22.188 +by (simp add: join_assoc[symmetric, of a b c], simp add: join_comm[of a b], simp add: join_assoc)
  22.189 +
  22.190 +lemma join_left_idempotent: "join y (join y x) = join y x"
  22.191 +by (simp add: join_assoc join_comm join_left_comm)
  22.192 +    
  22.193 +lemmas meet_aci = meet_assoc meet_comm meet_left_comm meet_left_idempotent
  22.194 +
  22.195 +lemmas join_aci = join_assoc join_comm join_left_comm join_left_idempotent
  22.196 +
  22.197 +lemma le_def_meet: "(x <= y) = (meet x y = x)" 
  22.198 +proof -
  22.199 +  have u: "x <= y \<longrightarrow> meet x y = x"
  22.200 +  proof 
  22.201 +    assume "x <= y"
  22.202 +    hence "x <= meet x y & meet x y <= x" by (simp add: meet_imp_le meet_left_le)
  22.203 +    thus "meet x y = x" by auto
  22.204 +  qed
  22.205 +  have v:"meet x y = x \<longrightarrow> x <= y" 
  22.206 +  proof 
  22.207 +    have a:"meet x y <= y" by (simp add: meet_right_le)
  22.208 +    assume "meet x y = x"
  22.209 +    hence "x = meet x y" by auto
  22.210 +    with a show "x <= y" by (auto)
  22.211 +  qed
  22.212 +  from u v show ?thesis by blast
  22.213 +qed
  22.214 +
  22.215 +lemma le_def_join: "(x <= y) = (join x y = y)" 
  22.216 +proof -
  22.217 +  have u: "x <= y \<longrightarrow> join x y = y"
  22.218 +  proof 
  22.219 +    assume "x <= y"
  22.220 +    hence "join x y <= y & y <= join x y" by (simp add: join_imp_le join_right_le)
  22.221 +    thus "join x y = y" by auto
  22.222 +  qed
  22.223 +  have v:"join x y = y \<longrightarrow> x <= y" 
  22.224 +  proof 
  22.225 +    have a:"x <= join x y" by (simp add: join_left_le)
  22.226 +    assume "join x y = y"
  22.227 +    hence "y = join x y" by auto
  22.228 +    with a show "x <= y" by (auto)
  22.229 +  qed
  22.230 +  from u v show ?thesis by blast
  22.231 +qed
  22.232 +
  22.233 +lemma meet_join_absorp: "meet x (join x y) = x"
  22.234 +proof -
  22.235 +  have a:"meet x (join x y) <= x" by (simp add: meet_left_le)
  22.236 +  have b:"x <= meet x (join x y)" by (rule meet_imp_le, simp_all add: join_left_le)
  22.237 +  from a b show ?thesis by auto
  22.238 +qed
  22.239 +
  22.240 +lemma join_meet_absorp: "join x (meet x y) = x"
  22.241 +proof - 
  22.242 +  have a:"x <= join x (meet x y)" by (simp add: join_left_le)
  22.243 +  have b:"join x (meet x y) <= x" by (rule join_imp_le, simp_all add: meet_left_le)
  22.244 +  from a b show ?thesis by auto
  22.245 +qed
  22.246 +
  22.247 +lemma meet_mono: "y \<le> z \<Longrightarrow> meet x y \<le> meet x z"
  22.248 +proof -
  22.249 +  assume a: "y <= z"
  22.250 +  have "meet x y <= x & meet x y <= y" by (simp add: meet_left_le meet_right_le)
  22.251 +  with a have "meet x y <= x & meet x y <= z" by auto 
  22.252 +  thus "meet x y <= meet x z" by (simp add: meet_imp_le)
  22.253 +qed
  22.254 +
  22.255 +lemma join_mono: "y \<le> z \<Longrightarrow> join x y \<le> join x z"
  22.256 +proof -
  22.257 +  assume a: "y \<le> z"
  22.258 +  have "x <= join x z & z <= join x z" by (simp add: join_left_le join_right_le)
  22.259 +  with a have "x <= join x z & y <= join x z" by auto
  22.260 +  thus "join x y <= join x z" by (simp add: join_imp_le)
  22.261 +qed
  22.262 +
  22.263 +lemma distrib_join_le: "join x (meet y z) \<le> meet (join x y) (join x z)" (is "_ <= ?r")
  22.264 +proof -
  22.265 +  have a: "x <= ?r" by (rule meet_imp_le, simp_all add: join_left_le)
  22.266 +  from meet_join_le have b: "meet y z <= ?r" 
  22.267 +    by (rule_tac meet_imp_le, (blast intro: order_trans)+)
  22.268 +  from a b show ?thesis by (simp add: join_imp_le)
  22.269 +qed
  22.270 +  
  22.271 +lemma distrib_meet_le: "join (meet x y) (meet x z) \<le> meet x (join y z)" (is "?l <= _") 
  22.272 +proof -
  22.273 +  have a: "?l <= x" by (rule join_imp_le, simp_all add: meet_left_le)
  22.274 +  from meet_join_le have b: "?l <= join y z" 
  22.275 +    by (rule_tac join_imp_le, (blast intro: order_trans)+)
  22.276 +  from a b show ?thesis by (simp add: meet_imp_le)
  22.277 +qed
  22.278 +
  22.279 +lemma meet_join_eq_imp_le: "a = c \<or> a = d \<or> b = c \<or> b = d \<Longrightarrow> meet a b \<le> join c d"
  22.280 +by (insert meet_join_le, blast intro: order_trans)
  22.281 +
  22.282 +lemma modular_le: "x \<le> z \<Longrightarrow> join x (meet y z) \<le> meet (join x y) z" (is "_ \<Longrightarrow> ?t <= _")
  22.283 +proof -
  22.284 +  assume a: "x <= z"
  22.285 +  have b: "?t <= join x y" by (rule join_imp_le, simp_all add: meet_join_le meet_join_eq_imp_le)
  22.286 +  have c: "?t <= z" by (rule join_imp_le, simp_all add: meet_join_le a)
  22.287 +  from b c show ?thesis by (simp add: meet_imp_le)
  22.288 +qed
  22.289 +
  22.290 +ML {*
  22.291 +val is_meet_unique = thm "is_meet_unique";
  22.292 +val is_join_unique = thm "is_join_unique";
  22.293 +val join_exists = thm "join_exists";
  22.294 +val meet_exists = thm "meet_exists";
  22.295 +val is_meet_meet = thm "is_meet_meet";
  22.296 +val meet_unique = thm "meet_unique";
  22.297 +val is_join_join = thm "is_join_join";
  22.298 +val join_unique = thm "join_unique";
  22.299 +val meet_left_le = thm "meet_left_le";
  22.300 +val meet_right_le = thm "meet_right_le";
  22.301 +val meet_imp_le = thm "meet_imp_le";
  22.302 +val join_left_le = thm "join_left_le";
  22.303 +val join_right_le = thm "join_right_le";
  22.304 +val join_imp_le = thm "join_imp_le";
  22.305 +val meet_join_le = thms "meet_join_le";
  22.306 +val is_meet_min = thm "is_meet_min";
  22.307 +val is_join_max = thm "is_join_max";
  22.308 +val meet_min = thm "meet_min";
  22.309 +val join_max = thm "join_max";
  22.310 +val meet_idempotent = thm "meet_idempotent";
  22.311 +val join_idempotent = thm "join_idempotent";
  22.312 +val meet_comm = thm "meet_comm";
  22.313 +val join_comm = thm "join_comm";
  22.314 +val meet_assoc = thm "meet_assoc";
  22.315 +val join_assoc = thm "join_assoc";
  22.316 +val meet_left_comm = thm "meet_left_comm";
  22.317 +val meet_left_idempotent = thm "meet_left_idempotent";
  22.318 +val join_left_comm = thm "join_left_comm";
  22.319 +val join_left_idempotent = thm "join_left_idempotent";
  22.320 +val meet_aci = thms "meet_aci";
  22.321 +val join_aci = thms "join_aci";
  22.322 +val le_def_meet = thm "le_def_meet";
  22.323 +val le_def_join = thm "le_def_join";
  22.324 +val meet_join_absorp = thm "meet_join_absorp";
  22.325 +val join_meet_absorp = thm "join_meet_absorp";
  22.326 +val meet_mono = thm "meet_mono";
  22.327 +val join_mono = thm "join_mono";
  22.328 +val distrib_join_le = thm "distrib_join_le";
  22.329 +val distrib_meet_le = thm "distrib_meet_le";
  22.330 +val meet_join_eq_imp_le = thm "meet_join_eq_imp_le";
  22.331 +val modular_le = thm "modular_le";
  22.332 +*}
  22.333 +
  22.334 +end
  22.335 \ No newline at end of file
    23.1 --- a/src/HOL/Library/Multiset.thy	Tue May 11 14:00:02 2004 +0200
    23.2 +++ b/src/HOL/Library/Multiset.thy	Tue May 11 20:11:08 2004 +0200
    23.3 @@ -108,7 +108,7 @@
    23.4  
    23.5  theorems union_ac = union_assoc union_commute union_lcomm
    23.6  
    23.7 -instance multiset :: (type) plus_ac0
    23.8 +instance multiset :: (type) comm_monoid_add
    23.9  proof 
   23.10    fix a b c :: "'a multiset"
   23.11    show "(a + b) + c = a + (b + c)" by (rule union_assoc)
    24.1 --- a/src/HOL/Matrix/Matrix.thy	Tue May 11 14:00:02 2004 +0200
    24.2 +++ b/src/HOL/Matrix/Matrix.thy	Tue May 11 20:11:08 2004 +0200
    24.3 @@ -195,7 +195,7 @@
    24.4  
    24.5  (*
    24.6  constdefs
    24.7 -  one_matrix :: "nat \<Rightarrow> ('a::semiring) matrix"
    24.8 +  one_matrix :: "nat \<Rightarrow> ('a::comm_semiring_1_cancel) matrix"
    24.9    "one_matrix n == Abs_matrix (% j i. if j = i & j < n then 1 else 0)"
   24.10  
   24.11  lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)"
   24.12 @@ -235,9 +235,9 @@
   24.13  by (simp add: max_def nrows)
   24.14  
   24.15  constdefs
   24.16 -  right_inverse_matrix :: "('a::semiring) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
   24.17 +  right_inverse_matrix :: "('a::comm_semiring_1_cancel) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
   24.18    "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X)))"
   24.19 -  inverse_matrix :: "('a::semiring) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
   24.20 +  inverse_matrix :: "('a::comm_semiring_1_cancel) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
   24.21    "inverse_matrix A X == (right_inverse_matrix A X) \<and> (right_inverse_matrix X A)"
   24.22  
   24.23  lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"
    25.1 --- a/src/HOL/Nat.thy	Tue May 11 14:00:02 2004 +0200
    25.2 +++ b/src/HOL/Nat.thy	Tue May 11 20:11:08 2004 +0200
    25.3 @@ -712,8 +712,8 @@
    25.4    by (induct m) (simp_all add: add_mult_distrib)
    25.5  
    25.6  
    25.7 -text{*The Naturals Form A Semiring*}
    25.8 -instance nat :: semiring
    25.9 +text{*The Naturals Form A comm_semiring_1_cancel*}
   25.10 +instance nat :: comm_semiring_1_cancel
   25.11  proof
   25.12    fix i j k :: nat
   25.13    show "(i + j) + k = i + (j + k)" by (rule nat_add_assoc)
   25.14 @@ -760,8 +760,8 @@
   25.15    done
   25.16  
   25.17  
   25.18 -text{*The Naturals Form an Ordered Semiring*}
   25.19 -instance nat :: ordered_semiring
   25.20 +text{*The Naturals Form an Ordered comm_semiring_1_cancel*}
   25.21 +instance nat :: ordered_semidom
   25.22  proof
   25.23    fix i j k :: nat
   25.24    show "0 < (1::nat)" by simp
    26.1 --- a/src/HOL/NumberTheory/WilsonBij.thy	Tue May 11 14:00:02 2004 +0200
    26.2 +++ b/src/HOL/NumberTheory/WilsonBij.thy	Tue May 11 20:11:08 2004 +0200
    26.3 @@ -75,7 +75,7 @@
    26.4  lemma aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
    26.5    -- {* same as @{text WilsonRuss} *}
    26.6    apply (unfold zcong_def)
    26.7 -  apply (simp add: Ring_and_Field.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
    26.8 +  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
    26.9    apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
   26.10     apply (simp add: mult_commute)
   26.11    apply (subst zdvd_zminus_iff)
    27.1 --- a/src/HOL/NumberTheory/WilsonRuss.thy	Tue May 11 14:00:02 2004 +0200
    27.2 +++ b/src/HOL/NumberTheory/WilsonRuss.thy	Tue May 11 20:11:08 2004 +0200
    27.3 @@ -86,7 +86,7 @@
    27.4  
    27.5  lemma inv_not_p_minus_1_aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
    27.6    apply (unfold zcong_def)
    27.7 -  apply (simp add: Ring_and_Field.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
    27.8 +  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
    27.9    apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
   27.10     apply (simp add: mult_commute)
   27.11    apply (subst zdvd_zminus_iff)
    28.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    28.2 +++ b/src/HOL/OrderedGroup.thy	Tue May 11 20:11:08 2004 +0200
    28.3 @@ -0,0 +1,950 @@
    28.4 +(*  Title:   HOL/Group.thy
    28.5 +    ID:      $Id$
    28.6 +    Author:  Gertrud Bauer and Markus Wenzel, TU Muenchen
    28.7 +             Lawrence C Paulson, University of Cambridge
    28.8 +             Revised and decoupled from Ring_and_Field.thy 
    28.9 +             by Steven Obua, TU Muenchen, in May 2004
   28.10 +    License: GPL (GNU GENERAL PUBLIC LICENSE)
   28.11 +*)
   28.12 +
   28.13 +header {* Ordered Groups *}
   28.14 +
   28.15 +theory OrderedGroup = Inductive + LOrder:
   28.16 +
   28.17 +text {*
   28.18 +  The theory of partially ordered groups is taken from the books:
   28.19 +  \begin{itemize}
   28.20 +  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
   28.21 +  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
   28.22 +  \end{itemize}
   28.23 +  Most of the used notions can also be looked up in 
   28.24 +  \begin{itemize}
   28.25 +  \item \emph{www.mathworld.com} by Eric Weisstein et. al.
   28.26 +  \item \emph{Algebra I} by van der Waerden, Springer.
   28.27 +  \end{itemize}
   28.28 +*}
   28.29 +
   28.30 +subsection {* Semigroups, Groups *}
   28.31 + 
   28.32 +axclass semigroup_add \<subseteq> plus
   28.33 +  add_assoc: "(a + b) + c = a + (b + c)"
   28.34 +
   28.35 +axclass ab_semigroup_add \<subseteq> semigroup_add
   28.36 +  add_commute: "a + b = b + a"
   28.37 +
   28.38 +lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::ab_semigroup_add))"
   28.39 +  by (rule mk_left_commute [of "op +", OF add_assoc add_commute])
   28.40 +
   28.41 +theorems add_ac = add_assoc add_commute add_left_commute
   28.42 +
   28.43 +axclass semigroup_mult \<subseteq> times
   28.44 +  mult_assoc: "(a * b) * c = a * (b * c)"
   28.45 +
   28.46 +axclass ab_semigroup_mult \<subseteq> semigroup_mult
   28.47 +  mult_commute: "a * b = b * a"
   28.48 +
   28.49 +lemma mult_left_commute: "a * (b * c) = b * (a * (c::'a::ab_semigroup_mult))"
   28.50 +  by (rule mk_left_commute [of "op *", OF mult_assoc mult_commute])
   28.51 +
   28.52 +theorems mult_ac = mult_assoc mult_commute mult_left_commute
   28.53 +
   28.54 +axclass comm_monoid_add \<subseteq> zero, ab_semigroup_add
   28.55 +  add_0[simp]: "0 + a = a"
   28.56 +
   28.57 +axclass monoid_mult \<subseteq> one, semigroup_mult
   28.58 +  mult_1_left[simp]: "1 * a  = a"
   28.59 +  mult_1_right[simp]: "a * 1 = a"
   28.60 +
   28.61 +axclass comm_monoid_mult \<subseteq> one, ab_semigroup_mult
   28.62 +  mult_1: "1 * a = a"
   28.63 +
   28.64 +instance comm_monoid_mult \<subseteq> monoid_mult
   28.65 +by (intro_classes, insert mult_1, simp_all add: mult_commute, auto)
   28.66 +
   28.67 +axclass cancel_semigroup_add \<subseteq> semigroup_add
   28.68 +  add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
   28.69 +  add_right_imp_eq: "b + a = c + a \<Longrightarrow> b = c"
   28.70 +
   28.71 +axclass cancel_ab_semigroup_add \<subseteq> ab_semigroup_add
   28.72 +  add_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
   28.73 +
   28.74 +instance cancel_ab_semigroup_add \<subseteq> cancel_semigroup_add
   28.75 +proof
   28.76 +  {
   28.77 +    fix a b c :: 'a
   28.78 +    assume "a + b = a + c"
   28.79 +    thus "b = c" by (rule add_imp_eq)
   28.80 +  }
   28.81 +  note f = this
   28.82 +  fix a b c :: 'a
   28.83 +  assume "b + a = c + a"
   28.84 +  hence "a + b = a + c" by (simp only: add_commute)
   28.85 +  thus "b = c" by (rule f)
   28.86 +qed
   28.87 +
   28.88 +axclass ab_group_add \<subseteq> minus, comm_monoid_add
   28.89 +  left_minus[simp]: " - a + a = 0"
   28.90 +  diff_minus: "a - b = a + (-b)"
   28.91 +
   28.92 +instance ab_group_add \<subseteq> cancel_ab_semigroup_add
   28.93 +proof 
   28.94 +  fix a b c :: 'a
   28.95 +  assume "a + b = a + c"
   28.96 +  hence "-a + a + b = -a + a + c" by (simp only: add_assoc)
   28.97 +  thus "b = c" by simp 
   28.98 +qed
   28.99 +
  28.100 +lemma add_0_right [simp]: "a + 0 = (a::'a::comm_monoid_add)"
  28.101 +proof -
  28.102 +  have "a + 0 = 0 + a" by (simp only: add_commute)
  28.103 +  also have "... = a" by simp
  28.104 +  finally show ?thesis .
  28.105 +qed
  28.106 +
  28.107 +lemma add_left_cancel [simp]:
  28.108 +     "(a + b = a + c) = (b = (c::'a::cancel_semigroup_add))"
  28.109 +by (blast dest: add_left_imp_eq) 
  28.110 +
  28.111 +lemma add_right_cancel [simp]:
  28.112 +     "(b + a = c + a) = (b = (c::'a::cancel_semigroup_add))"
  28.113 +  by (blast dest: add_right_imp_eq)
  28.114 +
  28.115 +lemma right_minus [simp]: "a + -(a::'a::ab_group_add) = 0"
  28.116 +proof -
  28.117 +  have "a + -a = -a + a" by (simp add: add_ac)
  28.118 +  also have "... = 0" by simp
  28.119 +  finally show ?thesis .
  28.120 +qed
  28.121 +
  28.122 +lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::ab_group_add))"
  28.123 +proof
  28.124 +  have "a = a - b + b" by (simp add: diff_minus add_ac)
  28.125 +  also assume "a - b = 0"
  28.126 +  finally show "a = b" by simp
  28.127 +next
  28.128 +  assume "a = b"
  28.129 +  thus "a - b = 0" by (simp add: diff_minus)
  28.130 +qed
  28.131 +
  28.132 +lemma minus_minus [simp]: "- (- (a::'a::ab_group_add)) = a"
  28.133 +proof (rule add_left_cancel [of "-a", THEN iffD1])
  28.134 +  show "(-a + -(-a) = -a + a)"
  28.135 +  by simp
  28.136 +qed
  28.137 +
  28.138 +lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::ab_group_add)"
  28.139 +apply (rule right_minus_eq [THEN iffD1, symmetric])
  28.140 +apply (simp add: diff_minus add_commute) 
  28.141 +done
  28.142 +
  28.143 +lemma minus_zero [simp]: "- 0 = (0::'a::ab_group_add)"
  28.144 +by (simp add: equals_zero_I)
  28.145 +
  28.146 +lemma diff_self [simp]: "a - (a::'a::ab_group_add) = 0"
  28.147 +  by (simp add: diff_minus)
  28.148 +
  28.149 +lemma diff_0 [simp]: "(0::'a::ab_group_add) - a = -a"
  28.150 +by (simp add: diff_minus)
  28.151 +
  28.152 +lemma diff_0_right [simp]: "a - (0::'a::ab_group_add) = a" 
  28.153 +by (simp add: diff_minus)
  28.154 +
  28.155 +lemma diff_minus_eq_add [simp]: "a - - b = a + (b::'a::ab_group_add)"
  28.156 +by (simp add: diff_minus)
  28.157 +
  28.158 +lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::ab_group_add))" 
  28.159 +proof 
  28.160 +  assume "- a = - b"
  28.161 +  hence "- (- a) = - (- b)"
  28.162 +    by simp
  28.163 +  thus "a=b" by simp
  28.164 +next
  28.165 +  assume "a=b"
  28.166 +  thus "-a = -b" by simp
  28.167 +qed
  28.168 +
  28.169 +lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::ab_group_add))"
  28.170 +by (subst neg_equal_iff_equal [symmetric], simp)
  28.171 +
  28.172 +lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::ab_group_add))"
  28.173 +by (subst neg_equal_iff_equal [symmetric], simp)
  28.174 +
  28.175 +text{*The next two equations can make the simplifier loop!*}
  28.176 +
  28.177 +lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::ab_group_add))"
  28.178 +proof -
  28.179 +  have "(- (-a) = - b) = (- a = b)" by (rule neg_equal_iff_equal)
  28.180 +  thus ?thesis by (simp add: eq_commute)
  28.181 +qed
  28.182 +
  28.183 +lemma minus_equation_iff: "(- a = b) = (- (b::'a::ab_group_add) = a)"
  28.184 +proof -
  28.185 +  have "(- a = - (-b)) = (a = -b)" by (rule neg_equal_iff_equal)
  28.186 +  thus ?thesis by (simp add: eq_commute)
  28.187 +qed
  28.188 +
  28.189 +lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::ab_group_add)"
  28.190 +apply (rule equals_zero_I)
  28.191 +apply (simp add: add_ac) 
  28.192 +done
  28.193 +
  28.194 +lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ab_group_add)"
  28.195 +by (simp add: diff_minus add_commute)
  28.196 +
  28.197 +subsection {* (Partially) Ordered Groups *} 
  28.198 +
  28.199 +axclass pordered_ab_semigroup_add \<subseteq> order, ab_semigroup_add
  28.200 +  add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
  28.201 +
  28.202 +axclass pordered_cancel_ab_semigroup_add \<subseteq> pordered_ab_semigroup_add, cancel_ab_semigroup_add
  28.203 +
  28.204 +instance pordered_cancel_ab_semigroup_add \<subseteq> pordered_ab_semigroup_add ..
  28.205 +
  28.206 +axclass pordered_ab_semigroup_add_imp_le \<subseteq> pordered_cancel_ab_semigroup_add
  28.207 +  add_le_imp_le_left: "c + a \<le> c + b \<Longrightarrow> a \<le> b"
  28.208 +
  28.209 +axclass pordered_ab_group_add \<subseteq> ab_group_add, pordered_ab_semigroup_add
  28.210 +
  28.211 +instance pordered_ab_group_add \<subseteq> pordered_ab_semigroup_add_imp_le
  28.212 +proof
  28.213 +  fix a b c :: 'a
  28.214 +  assume "c + a \<le> c + b"
  28.215 +  hence "(-c) + (c + a) \<le> (-c) + (c + b)" by (rule add_left_mono)
  28.216 +  hence "((-c) + c) + a \<le> ((-c) + c) + b" by (simp only: add_assoc)
  28.217 +  thus "a \<le> b" by simp
  28.218 +qed
  28.219 +
  28.220 +axclass ordered_cancel_ab_semigroup_add \<subseteq> pordered_cancel_ab_semigroup_add, linorder
  28.221 +
  28.222 +instance ordered_cancel_ab_semigroup_add \<subseteq> pordered_ab_semigroup_add_imp_le
  28.223 +proof
  28.224 +  fix a b c :: 'a
  28.225 +  assume le: "c + a <= c + b"  
  28.226 +  show "a <= b"
  28.227 +  proof (rule ccontr)
  28.228 +    assume w: "~ a \<le> b"
  28.229 +    hence "b <= a" by (simp add: linorder_not_le)
  28.230 +    hence le2: "c+b <= c+a" by (rule add_left_mono)
  28.231 +    have "a = b" 
  28.232 +      apply (insert le)
  28.233 +      apply (insert le2)
  28.234 +      apply (drule order_antisym, simp_all)
  28.235 +      done
  28.236 +    with w  show False 
  28.237 +      by (simp add: linorder_not_le [symmetric])
  28.238 +  qed
  28.239 +qed
  28.240 +
  28.241 +lemma add_right_mono: "a \<le> (b::'a::pordered_ab_semigroup_add) ==> a + c \<le> b + c"
  28.242 +by (simp add: add_commute[of _ c] add_left_mono)
  28.243 +
  28.244 +text {* non-strict, in both arguments *}
  28.245 +lemma add_mono:
  28.246 +     "[|a \<le> b;  c \<le> d|] ==> a + c \<le> b + (d::'a::pordered_ab_semigroup_add)"
  28.247 +  apply (erule add_right_mono [THEN order_trans])
  28.248 +  apply (simp add: add_commute add_left_mono)
  28.249 +  done
  28.250 +
  28.251 +lemma add_strict_left_mono:
  28.252 +     "a < b ==> c + a < c + (b::'a::pordered_cancel_ab_semigroup_add)"
  28.253 + by (simp add: order_less_le add_left_mono) 
  28.254 +
  28.255 +lemma add_strict_right_mono:
  28.256 +     "a < b ==> a + c < b + (c::'a::pordered_cancel_ab_semigroup_add)"
  28.257 + by (simp add: add_commute [of _ c] add_strict_left_mono)
  28.258 +
  28.259 +text{*Strict monotonicity in both arguments*}
  28.260 +lemma add_strict_mono: "[|a<b; c<d|] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"
  28.261 +apply (erule add_strict_right_mono [THEN order_less_trans])
  28.262 +apply (erule add_strict_left_mono)
  28.263 +done
  28.264 +
  28.265 +lemma add_less_le_mono:
  28.266 +     "[| a<b; c\<le>d |] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"
  28.267 +apply (erule add_strict_right_mono [THEN order_less_le_trans])
  28.268 +apply (erule add_left_mono) 
  28.269 +done
  28.270 +
  28.271 +lemma add_le_less_mono:
  28.272 +     "[| a\<le>b; c<d |] ==> a + c < b + (d::'a::pordered_cancel_ab_semigroup_add)"
  28.273 +apply (erule add_right_mono [THEN order_le_less_trans])
  28.274 +apply (erule add_strict_left_mono) 
  28.275 +done
  28.276 +
  28.277 +lemma add_less_imp_less_left:
  28.278 +      assumes less: "c + a < c + b"  shows "a < (b::'a::pordered_ab_semigroup_add_imp_le)"
  28.279 +proof -
  28.280 +  from less have le: "c + a <= c + b" by (simp add: order_le_less)
  28.281 +  have "a <= b" 
  28.282 +    apply (insert le)
  28.283 +    apply (drule add_le_imp_le_left)
  28.284 +    by (insert le, drule add_le_imp_le_left, assumption)
  28.285 +  moreover have "a \<noteq> b"
  28.286 +  proof (rule ccontr)
  28.287 +    assume "~(a \<noteq> b)"
  28.288 +    then have "a = b" by simp
  28.289 +    then have "c + a = c + b" by simp
  28.290 +    with less show "False"by simp
  28.291 +  qed
  28.292 +  ultimately show "a < b" by (simp add: order_le_less)
  28.293 +qed
  28.294 +
  28.295 +lemma add_less_imp_less_right:
  28.296 +      "a + c < b + c ==> a < (b::'a::pordered_ab_semigroup_add_imp_le)"
  28.297 +apply (rule add_less_imp_less_left [of c])
  28.298 +apply (simp add: add_commute)  
  28.299 +done
  28.300 +
  28.301 +lemma add_less_cancel_left [simp]:
  28.302 +    "(c+a < c+b) = (a < (b::'a::pordered_ab_semigroup_add_imp_le))"
  28.303 +by (blast intro: add_less_imp_less_left add_strict_left_mono) 
  28.304 +
  28.305 +lemma add_less_cancel_right [simp]:
  28.306 +    "(a+c < b+c) = (a < (b::'a::pordered_ab_semigroup_add_imp_le))"
  28.307 +by (blast intro: add_less_imp_less_right add_strict_right_mono)
  28.308 +
  28.309 +lemma add_le_cancel_left [simp]:
  28.310 +    "(c+a \<le> c+b) = (a \<le> (b::'a::pordered_ab_semigroup_add_imp_le))"
  28.311 +by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
  28.312 +
  28.313 +lemma add_le_cancel_right [simp]:
  28.314 +    "(a+c \<le> b+c) = (a \<le> (b::'a::pordered_ab_semigroup_add_imp_le))"
  28.315 +by (simp add: add_commute[of a c] add_commute[of b c])
  28.316 +
  28.317 +lemma add_le_imp_le_right:
  28.318 +      "a + c \<le> b + c ==> a \<le> (b::'a::pordered_ab_semigroup_add_imp_le)"
  28.319 +by simp
  28.320 +
  28.321 +lemma add_increasing: "[|0\<le>a; b\<le>c|] ==> b \<le> a + (c::'a::{pordered_ab_semigroup_add_imp_le, comm_monoid_add})"
  28.322 +by (insert add_mono [of 0 a b c], simp)
  28.323 +
  28.324 +subsection {* Ordering Rules for Unary Minus *}
  28.325 +
  28.326 +lemma le_imp_neg_le:
  28.327 +      assumes "a \<le> (b::'a::{pordered_ab_semigroup_add_imp_le, ab_group_add})" shows "-b \<le> -a"
  28.328 +proof -
  28.329 +  have "-a+a \<le> -a+b"
  28.330 +    by (rule add_left_mono) 
  28.331 +  hence "0 \<le> -a+b"
  28.332 +    by simp
  28.333 +  hence "0 + (-b) \<le> (-a + b) + (-b)"
  28.334 +    by (rule add_right_mono) 
  28.335 +  thus ?thesis
  28.336 +    by (simp add: add_assoc)
  28.337 +qed
  28.338 +
  28.339 +lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::pordered_ab_group_add))"
  28.340 +proof 
  28.341 +  assume "- b \<le> - a"
  28.342 +  hence "- (- a) \<le> - (- b)"
  28.343 +    by (rule le_imp_neg_le)
  28.344 +  thus "a\<le>b" by simp
  28.345 +next
  28.346 +  assume "a\<le>b"
  28.347 +  thus "-b \<le> -a" by (rule le_imp_neg_le)
  28.348 +qed
  28.349 +
  28.350 +lemma neg_le_0_iff_le [simp]: "(-a \<le> 0) = (0 \<le> (a::'a::pordered_ab_group_add))"
  28.351 +by (subst neg_le_iff_le [symmetric], simp)
  28.352 +
  28.353 +lemma neg_0_le_iff_le [simp]: "(0 \<le> -a) = (a \<le> (0::'a::pordered_ab_group_add))"
  28.354 +by (subst neg_le_iff_le [symmetric], simp)
  28.355 +
  28.356 +lemma neg_less_iff_less [simp]: "(-b < -a) = (a < (b::'a::pordered_ab_group_add))"
  28.357 +by (force simp add: order_less_le) 
  28.358 +
  28.359 +lemma neg_less_0_iff_less [simp]: "(-a < 0) = (0 < (a::'a::pordered_ab_group_add))"
  28.360 +by (subst neg_less_iff_less [symmetric], simp)
  28.361 +
  28.362 +lemma neg_0_less_iff_less [simp]: "(0 < -a) = (a < (0::'a::pordered_ab_group_add))"
  28.363 +by (subst neg_less_iff_less [symmetric], simp)
  28.364 +
  28.365 +text{*The next several equations can make the simplifier loop!*}
  28.366 +
  28.367 +lemma less_minus_iff: "(a < - b) = (b < - (a::'a::pordered_ab_group_add))"
  28.368 +proof -
  28.369 +  have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)
  28.370 +  thus ?thesis by simp
  28.371 +qed
  28.372 +
  28.373 +lemma minus_less_iff: "(- a < b) = (- b < (a::'a::pordered_ab_group_add))"
  28.374 +proof -
  28.375 +  have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)
  28.376 +  thus ?thesis by simp
  28.377 +qed
  28.378 +
  28.379 +lemma le_minus_iff: "(a \<le> - b) = (b \<le> - (a::'a::pordered_ab_group_add))"
  28.380 +proof -
  28.381 +  have mm: "!! a (b::'a). (-(-a)) < -b \<Longrightarrow> -(-b) < -a" by (simp only: minus_less_iff)
  28.382 +  have "(- (- a) <= -b) = (b <= - a)" 
  28.383 +    apply (auto simp only: order_le_less)
  28.384 +    apply (drule mm)
  28.385 +    apply (simp_all)
  28.386 +    apply (drule mm[simplified], assumption)
  28.387 +    done
  28.388 +  then show ?thesis by simp
  28.389 +qed
  28.390 +
  28.391 +lemma minus_le_iff: "(- a \<le> b) = (- b \<le> (a::'a::pordered_ab_group_add))"
  28.392 +by (auto simp add: order_le_less minus_less_iff)
  28.393 +
  28.394 +lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::ab_group_add)"
  28.395 +by (simp add: diff_minus add_ac)
  28.396 +
  28.397 +lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::ab_group_add)"
  28.398 +by (simp add: diff_minus add_ac)
  28.399 +
  28.400 +lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::ab_group_add))"
  28.401 +by (auto simp add: diff_minus add_assoc)
  28.402 +
  28.403 +lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::ab_group_add) = c)"
  28.404 +by (auto simp add: diff_minus add_assoc)
  28.405 +
  28.406 +lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::ab_group_add))"
  28.407 +by (simp add: diff_minus add_ac)
  28.408 +
  28.409 +lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::ab_group_add)"
  28.410 +by (simp add: diff_minus add_ac)
  28.411 +
  28.412 +lemma diff_add_cancel: "a - b + b = (a::'a::ab_group_add)"
  28.413 +by (simp add: diff_minus add_ac)
  28.414 +
  28.415 +lemma add_diff_cancel: "a + b - b = (a::'a::ab_group_add)"
  28.416 +by (simp add: diff_minus add_ac)
  28.417 +
  28.418 +text{*Further subtraction laws for ordered rings*}
  28.419 +
  28.420 +lemma less_iff_diff_less_0: "(a < b) = (a - b < (0::'a::pordered_ab_group_add))"
  28.421 +proof -
  28.422 +  have  "(a < b) = (a + (- b) < b + (-b))"  
  28.423 +    by (simp only: add_less_cancel_right)
  28.424 +  also have "... =  (a - b < 0)" by (simp add: diff_minus)
  28.425 +  finally show ?thesis .
  28.426 +qed
  28.427 +
  28.428 +lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::pordered_ab_group_add))"
  28.429 +apply (subst less_iff_diff_less_0)
  28.430 +apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
  28.431 +apply (simp add: diff_minus add_ac)
  28.432 +done
  28.433 +
  28.434 +lemma less_diff_eq: "(a < c-b) = (a + (b::'a::pordered_ab_group_add) < c)"
  28.435 +apply (subst less_iff_diff_less_0)
  28.436 +apply (rule less_iff_diff_less_0 [of _ "c-b", THEN ssubst])
  28.437 +apply (simp add: diff_minus add_ac)
  28.438 +done
  28.439 +
  28.440 +lemma diff_le_eq: "(a-b \<le> c) = (a \<le> c + (b::'a::pordered_ab_group_add))"
  28.441 +by (auto simp add: order_le_less diff_less_eq diff_add_cancel add_diff_cancel)
  28.442 +
  28.443 +lemma le_diff_eq: "(a \<le> c-b) = (a + (b::'a::pordered_ab_group_add) \<le> c)"
  28.444 +by (auto simp add: order_le_less less_diff_eq diff_add_cancel add_diff_cancel)
  28.445 +
  28.446 +text{*This list of rewrites simplifies (in)equalities by bringing subtractions
  28.447 +  to the top and then moving negative terms to the other side.
  28.448 +  Use with @{text add_ac}*}
  28.449 +lemmas compare_rls =
  28.450 +       diff_minus [symmetric]
  28.451 +       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
  28.452 +       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
  28.453 +       diff_eq_eq eq_diff_eq
  28.454 +
  28.455 +
  28.456 +subsection{*Lemmas for the @{text cancel_numerals} simproc*}
  28.457 +
  28.458 +lemma eq_iff_diff_eq_0: "(a = b) = (a-b = (0::'a::ab_group_add))"
  28.459 +by (simp add: compare_rls)
  28.460 +
  28.461 +lemma le_iff_diff_le_0: "(a \<le> b) = (a-b \<le> (0::'a::pordered_ab_group_add))"
  28.462 +by (simp add: compare_rls)
  28.463 +
  28.464 +subsection {* Lattice Ordered (Abelian) Groups *}
  28.465 +
  28.466 +axclass lordered_ab_group_meet < pordered_ab_group_add, meet_semilorder
  28.467 +
  28.468 +axclass lordered_ab_group_join < pordered_ab_group_add, join_semilorder
  28.469 +
  28.470 +lemma add_meet_distrib_left: "a + (meet b c) = meet (a + b) (a + (c::'a::{pordered_ab_group_add, meet_semilorder}))"
  28.471 +apply (rule order_antisym)
  28.472 +apply (rule meet_imp_le, simp_all add: meet_join_le)
  28.473 +apply (rule add_le_imp_le_left [of "-a"])
  28.474 +apply (simp only: add_assoc[symmetric], simp)
  28.475 +apply (rule meet_imp_le)
  28.476 +apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp add: meet_join_le)+
  28.477 +done
  28.478 +
  28.479 +lemma add_join_distrib_left: "a + (join b c) = join (a + b) (a+ (c::'a::{pordered_ab_group_add, join_semilorder}))" 
  28.480 +apply (rule order_antisym)
  28.481 +apply (rule add_le_imp_le_left [of "-a"])
  28.482 +apply (simp only: add_assoc[symmetric], simp)
  28.483 +apply (rule join_imp_le)
  28.484 +apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp add: meet_join_le)+
  28.485 +apply (rule join_imp_le)
  28.486 +apply (simp_all add: meet_join_le)
  28.487 +done
  28.488 +
  28.489 +lemma is_join_neg_meet: "is_join (% (a::'a::{pordered_ab_group_add, meet_semilorder}) b. - (meet (-a) (-b)))"
  28.490 +apply (auto simp add: is_join_def)
  28.491 +apply (rule_tac c="meet (-a) (-b)" in add_le_imp_le_right, simp, simp add: add_meet_distrib_left meet_join_le)
  28.492 +apply (rule_tac c="meet (-a) (-b)" in add_le_imp_le_right, simp, simp add: add_meet_distrib_left meet_join_le)
  28.493 +apply (subst neg_le_iff_le[symmetric]) 
  28.494 +apply (simp add: meet_imp_le)
  28.495 +done
  28.496 +
  28.497 +lemma is_meet_neg_join: "is_meet (% (a::'a::{pordered_ab_group_add, join_semilorder}) b. - (join (-a) (-b)))"
  28.498 +apply (auto simp add: is_meet_def)
  28.499 +apply (rule_tac c="join (-a) (-b)" in add_le_imp_le_right, simp, simp add: add_join_distrib_left meet_join_le)
  28.500 +apply (rule_tac c="join (-a) (-b)" in add_le_imp_le_right, simp, simp add: add_join_distrib_left meet_join_le)
  28.501 +apply (subst neg_le_iff_le[symmetric]) 
  28.502 +apply (simp add: join_imp_le)
  28.503 +done
  28.504 +
  28.505 +axclass lordered_ab_group \<subseteq> pordered_ab_group_add, lorder
  28.506 +
  28.507 +instance lordered_ab_group_meet \<subseteq> lordered_ab_group
  28.508 +proof 
  28.509 +  show "? j. is_join (j::'a\<Rightarrow>'a\<Rightarrow>('a::lordered_ab_group_meet))" by (blast intro: is_join_neg_meet)
  28.510 +qed
  28.511 +
  28.512 +instance lordered_ab_group_join \<subseteq> lordered_ab_group
  28.513 +proof
  28.514 +  show "? m. is_meet (m::'a\<Rightarrow>'a\<Rightarrow>('a::lordered_ab_group_join))" by (blast intro: is_meet_neg_join)
  28.515 +qed
  28.516 +
  28.517 +lemma add_join_distrib_right: "(join a b) + (c::'a::lordered_ab_group) = join (a+c) (b+c)"
  28.518 +proof -
  28.519 +  have "c + (join a b) = join (c+a) (c+b)" by (simp add: add_join_distrib_left)
  28.520 +  thus ?thesis by (simp add: add_commute)
  28.521 +qed
  28.522 +
  28.523 +lemma add_meet_distrib_right: "(meet a b) + (c::'a::lordered_ab_group) = meet (a+c) (b+c)"
  28.524 +proof -
  28.525 +  have "c + (meet a b) = meet (c+a) (c+b)" by (simp add: add_meet_distrib_left)
  28.526 +  thus ?thesis by (simp add: add_commute)
  28.527 +qed
  28.528 +
  28.529 +lemmas add_meet_join_distribs = add_meet_distrib_right add_meet_distrib_left add_join_distrib_right add_join_distrib_left
  28.530 +
  28.531 +lemma join_eq_neg_meet: "join a (b::'a::lordered_ab_group) = - meet (-a) (-b)"
  28.532 +by (simp add: is_join_unique[OF is_join_join is_join_neg_meet])
  28.533 +
  28.534 +lemma meet_eq_neg_join: "meet a (b::'a::lordered_ab_group) = - join (-a) (-b)"
  28.535 +by (simp add: is_meet_unique[OF is_meet_meet is_meet_neg_join])
  28.536 +
  28.537 +lemma add_eq_meet_join: "a + b = (join a b) + (meet a (b::'a::lordered_ab_group))"
  28.538 +proof -
  28.539 +  have "0 = - meet 0 (a-b) + meet (a-b) 0" by (simp add: meet_comm)
  28.540 +  hence "0 = join 0 (b-a) + meet (a-b) 0" by (simp add: meet_eq_neg_join)
  28.541 +  hence "0 = (-a + join a b) + (meet a b + (-b))"
  28.542 +    apply (simp add: add_join_distrib_left add_meet_distrib_right)
  28.543 +    by (simp add: diff_minus add_commute)
  28.544 +  thus ?thesis
  28.545 +    apply (simp add: compare_rls)
  28.546 +    apply (subst add_left_cancel[symmetric, of "a+b" "join a b + meet a b" "-a"])
  28.547 +    apply (simp only: add_assoc, simp add: add_assoc[symmetric])
  28.548 +    done
  28.549 +qed
  28.550 +
  28.551 +subsection {* Positive Part, Negative Part, Absolute Value *}
  28.552 +
  28.553 +constdefs
  28.554 +  pprt :: "'a \<Rightarrow> ('a::lordered_ab_group)"
  28.555 +  "pprt x == join x 0"
  28.556 +  nprt :: "'a \<Rightarrow> ('a::lordered_ab_group)"
  28.557 +  "nprt x == meet x 0"
  28.558 +
  28.559 +lemma prts: "a = pprt a + nprt a"
  28.560 +by (simp add: pprt_def nprt_def add_eq_meet_join[symmetric])
  28.561 +
  28.562 +lemma zero_le_pprt[simp]: "0 \<le> pprt a"
  28.563 +by (simp add: pprt_def meet_join_le)
  28.564 +
  28.565 +lemma nprt_le_zero[simp]: "nprt a \<le> 0"
  28.566 +by (simp add: nprt_def meet_join_le)
  28.567 +
  28.568 +lemma le_eq_neg: "(a \<le> -b) = (a + b \<le> (0::_::lordered_ab_group))" (is "?l = ?r")
  28.569 +proof -
  28.570 +  have a: "?l \<longrightarrow> ?r"
  28.571 +    apply (auto)
  28.572 +    apply (rule add_le_imp_le_right[of _ "-b" _])
  28.573 +    apply (simp add: add_assoc)
  28.574 +    done
  28.575 +  have b: "?r \<longrightarrow> ?l"
  28.576 +    apply (auto)
  28.577 +    apply (rule add_le_imp_le_right[of _ "b" _])
  28.578 +    apply (simp)
  28.579 +    done
  28.580 +  from a b show ?thesis by blast
  28.581 +qed
  28.582 +
  28.583 +lemma join_0_imp_0: "join a (-a) = 0 \<Longrightarrow> a = (0::'a::lordered_ab_group)"
  28.584 +proof -
  28.585 +  {
  28.586 +    fix a::'a
  28.587 +    assume hyp: "join a (-a) = 0"
  28.588 +    hence "join a (-a) + a = a" by (simp)
  28.589 +    hence "join (a+a) 0 = a" by (simp add: add_join_distrib_right) 
  28.590 +    hence "join (a+a) 0 <= a" by (simp)
  28.591 +    hence "0 <= a" by (blast intro: order_trans meet_join_le)
  28.592 +  }
  28.593 +  note p = this
  28.594 +  thm p
  28.595 +  assume hyp:"join a (-a) = 0"
  28.596 +  hence hyp2:"join (-a) (-(-a)) = 0" by (simp add: join_comm)
  28.597 +  from p[OF hyp] p[OF hyp2] show "a = 0" by simp
  28.598 +qed
  28.599 +
  28.600 +lemma meet_0_imp_0: "meet a (-a) = 0 \<Longrightarrow> a = (0::'a::lordered_ab_group)"
  28.601 +apply (simp add: meet_eq_neg_join)
  28.602 +apply (simp add: join_comm)
  28.603 +apply (subst join_0_imp_0)
  28.604 +by auto
  28.605 +
  28.606 +lemma join_0_eq_0[simp]: "(join a (-a) = 0) = (a = (0::'a::lordered_ab_group))"
  28.607 +by (auto, erule join_0_imp_0)
  28.608 +
  28.609 +lemma meet_0_eq_0[simp]: "(meet a (-a) = 0) = (a = (0::'a::lordered_ab_group))"
  28.610 +by (auto, erule meet_0_imp_0)
  28.611 +
  28.612 +lemma zero_le_double_add_iff_zero_le_single_add[simp]: "(0 \<le> a + a) = (0 \<le> (a::'a::lordered_ab_group))"
  28.613 +proof
  28.614 +  assume "0 <= a + a"
  28.615 +  hence a:"meet (a+a) 0 = 0" by (simp add: le_def_meet meet_comm)
  28.616 +  have "(meet a 0)+(meet a 0) = meet (meet (a+a) 0) a" (is "?l=_") by (simp add: add_meet_join_distribs meet_aci)
  28.617 +  hence "?l = 0 + meet a 0" by (simp add: a, simp add: meet_comm)
  28.618 +  hence "meet a 0 = 0" by (simp only: add_right_cancel)
  28.619 +  then show "0 <= a" by (simp add: le_def_meet meet_comm)    
  28.620 +next  
  28.621 +  assume a: "0 <= a"
  28.622 +  show "0 <= a + a" by (simp add: add_mono[OF a a, simplified])
  28.623 +qed
  28.624 +
  28.625 +lemma double_add_le_zero_iff_single_add_le_zero[simp]: "(a + a <= 0) = ((a::'a::lordered_ab_group) <= 0)" 
  28.626 +proof -
  28.627 +  have "(a + a <= 0) = (0 <= -(a+a))" by (subst le_minus_iff, simp)
  28.628 +  moreover have "\<dots> = (a <= 0)" by (simp add: zero_le_double_add_iff_zero_le_single_add)
  28.629 +  ultimately show ?thesis by blast
  28.630 +qed
  28.631 +
  28.632 +lemma double_add_less_zero_iff_single_less_zero[simp]: "(a+a<0) = ((a::'a::{pordered_ab_group_add,linorder}) < 0)" (is ?s)
  28.633 +proof cases
  28.634 +  assume a: "a < 0"
  28.635 +  thus ?s by (simp add:  add_strict_mono[OF a a, simplified])
  28.636 +next
  28.637 +  assume "~(a < 0)" 
  28.638 +  hence a:"0 <= a" by (simp)
  28.639 +  hence "0 <= a+a" by (simp add: add_mono[OF a a, simplified])
  28.640 +  hence "~(a+a < 0)" by simp
  28.641 +  with a show ?thesis by simp 
  28.642 +qed
  28.643 +
  28.644 +axclass lordered_ab_group_abs \<subseteq> lordered_ab_group
  28.645 +  abs_lattice: "abs x = join x (-x)"
  28.646 +
  28.647 +lemma abs_zero[simp]: "abs 0 = (0::'a::lordered_ab_group_abs)"
  28.648 +by (simp add: abs_lattice)
  28.649 +
  28.650 +lemma abs_eq_0[simp]: "(abs a = 0) = (a = (0::'a::lordered_ab_group_abs))"
  28.651 +by (simp add: abs_lattice)
  28.652 +
  28.653 +lemma abs_0_eq[simp]: "(0 = abs a) = (a = (0::'a::lordered_ab_group_abs))"
  28.654 +proof -
  28.655 +  have "(0 = abs a) = (abs a = 0)" by (simp only: eq_ac)
  28.656 +  thus ?thesis by simp
  28.657 +qed
  28.658 +
  28.659 +lemma neg_meet_eq_join[simp]: "- meet a (b::_::lordered_ab_group) = join (-a) (-b)"
  28.660 +by (simp add: meet_eq_neg_join)
  28.661 +
  28.662 +lemma neg_join_eq_meet[simp]: "- join a (b::_::lordered_ab_group) = meet (-a) (-b)"
  28.663 +by (simp del: neg_meet_eq_join add: join_eq_neg_meet)
  28.664 +
  28.665 +lemma join_eq_if: "join a (-a) = (if a < 0 then -a else (a::'a::{lordered_ab_group, linorder}))"
  28.666 +proof -
  28.667 +  note b = add_le_cancel_right[of a a "-a",symmetric,simplified]
  28.668 +  have c: "a + a = 0 \<Longrightarrow> -a = a" by (rule add_right_imp_eq[of _ a], simp)
  28.669 +  show ?thesis
  28.670 +    apply (auto simp add: join_max max_def b linorder_not_less)
  28.671 +    apply (drule order_antisym, auto)
  28.672 +    done
  28.673 +qed
  28.674 +
  28.675 +lemma abs_if_lattice: "\<bar>a\<bar> = (if a < 0 then -a else (a::'a::{lordered_ab_group_abs, linorder}))"
  28.676 +proof -
  28.677 +  show ?thesis by (simp add: abs_lattice join_eq_if)
  28.678 +qed
  28.679 +
  28.680 +lemma abs_ge_zero[simp]: "0 \<le> abs (a::'a::lordered_ab_group_abs)"
  28.681 +proof -
  28.682 +  have a:"a <= abs a" and b:"-a <= abs a" by (auto simp add: abs_lattice meet_join_le)
  28.683 +  show ?thesis by (rule add_mono[OF a b, simplified])
  28.684 +qed
  28.685 +  
  28.686 +lemma abs_le_zero_iff [simp]: "(abs a \<le> (0::'a::lordered_ab_group_abs)) = (a = 0)" 
  28.687 +proof
  28.688 +  assume "abs a <= 0"
  28.689 +  hence "abs a = 0" by (auto dest: order_antisym)
  28.690 +  thus "a = 0" by simp
  28.691 +next
  28.692 +  assume "a = 0"
  28.693 +  thus "abs a <= 0" by simp
  28.694 +qed
  28.695 +
  28.696 +lemma zero_less_abs_iff [simp]: "(0 < abs a) = (a \<noteq> (0::'a::lordered_ab_group_abs))"
  28.697 +by (simp add: order_less_le)
  28.698 +
  28.699 +lemma abs_not_less_zero [simp]: "~ abs a < (0::'a::lordered_ab_group_abs)"
  28.700 +proof -
  28.701 +  have a:"!! x (y::_::order). x <= y \<Longrightarrow> ~(y < x)" by auto
  28.702 +  show ?thesis by (simp add: a)
  28.703 +qed
  28.704 +
  28.705 +lemma abs_ge_self: "a \<le> abs (a::'a::lordered_ab_group_abs)"
  28.706 +by (simp add: abs_lattice meet_join_le)
  28.707 +
  28.708 +lemma abs_ge_minus_self: "-a \<le> abs (a::'a::lordered_ab_group_abs)"
  28.709 +by (simp add: abs_lattice meet_join_le)
  28.710 +
  28.711 +lemma le_imp_join_eq: "a \<le> b \<Longrightarrow> join a b = b" 
  28.712 +by (simp add: le_def_join)
  28.713 +
  28.714 +lemma ge_imp_join_eq: "b \<le> a \<Longrightarrow> join a b = a"
  28.715 +by (simp add: le_def_join join_aci)
  28.716 +
  28.717 +lemma le_imp_meet_eq: "a \<le> b \<Longrightarrow> meet a b = a"
  28.718 +by (simp add: le_def_meet)
  28.719 +
  28.720 +lemma ge_imp_meet_eq: "b \<le> a \<Longrightarrow> meet a b = b"
  28.721 +by (simp add: le_def_meet meet_aci)
  28.722 +
  28.723 +lemma abs_prts: "abs (a::_::lordered_ab_group_abs) = pprt a - nprt a"
  28.724 +apply (simp add: pprt_def nprt_def diff_minus)
  28.725 +apply (simp add: add_meet_join_distribs join_aci abs_lattice[symmetric])
  28.726 +apply (subst le_imp_join_eq, auto)
  28.727 +done
  28.728 +
  28.729 +lemma abs_minus_cancel [simp]: "abs (-a) = abs(a::'a::lordered_ab_group_abs)"
  28.730 +by (simp add: abs_lattice join_comm)
  28.731 +
  28.732 +lemma abs_idempotent [simp]: "abs (abs a) = abs (a::'a::lordered_ab_group_abs)"
  28.733 +apply (simp add: abs_lattice[of "abs a"])
  28.734 +apply (subst ge_imp_join_eq)
  28.735 +apply (rule order_trans[of _ 0])
  28.736 +by auto
  28.737 +
  28.738 +lemma zero_le_iff_zero_nprt: "(0 \<le> a) = (nprt a = 0)"
  28.739 +by (simp add: le_def_meet nprt_def meet_comm)
  28.740 +
  28.741 +lemma le_zero_iff_zero_pprt: "(a \<le> 0) = (pprt a = 0)"
  28.742 +by (simp add: le_def_join pprt_def join_comm)
  28.743 +
  28.744 +lemma le_zero_iff_pprt_id: "(0 \<le> a) = (pprt a = a)"
  28.745 +by (simp add: le_def_join pprt_def join_comm)
  28.746 +
  28.747 +lemma zero_le_iff_nprt_id: "(a \<le> 0) = (nprt a = a)"
  28.748 +by (simp add: le_def_meet nprt_def meet_comm)
  28.749 +
  28.750 +lemma iff2imp: "(A=B) \<Longrightarrow> (A \<Longrightarrow> B)"
  28.751 +by (simp)
  28.752 +
  28.753 +lemma imp_abs_id: "0 \<le> a \<Longrightarrow> abs a = (a::'a::lordered_ab_group_abs)"
  28.754 +by (simp add: iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_pprt_id] abs_prts)
  28.755 +
  28.756 +lemma imp_abs_neg_id: "a \<le> 0 \<Longrightarrow> abs a = -(a::'a::lordered_ab_group_abs)"
  28.757 +by (simp add: iff2imp[OF le_zero_iff_zero_pprt] iff2imp[OF zero_le_iff_nprt_id] abs_prts)
  28.758 +
  28.759 +lemma abs_leI: "[|a \<le> b; -a \<le> b|] ==> abs a \<le> (b::'a::lordered_ab_group_abs)"
  28.760 +by (simp add: abs_lattice join_imp_le)
  28.761 +
  28.762 +lemma le_minus_self_iff: "(a \<le> -a) = (a \<le> (0::'a::lordered_ab_group))"
  28.763 +proof -
  28.764 +  from add_le_cancel_left[of "-a" "a+a" "0"] have "(a <= -a) = (a+a <= 0)" 
  28.765 +    by (simp add: add_assoc[symmetric])
  28.766 +  thus ?thesis by simp
  28.767 +qed
  28.768 +
  28.769 +lemma minus_le_self_iff: "(-a \<le> a) = (0 \<le> (a::'a::lordered_ab_group))"
  28.770 +proof -
  28.771 +  from add_le_cancel_left[of "-a" "0" "a+a"] have "(-a <= a) = (0 <= a+a)" 
  28.772 +    by (simp add: add_assoc[symmetric])
  28.773 +  thus ?thesis by simp
  28.774 +qed
  28.775 +
  28.776 +lemma abs_le_D1: "abs a \<le> b ==> a \<le> (b::'a::lordered_ab_group_abs)"
  28.777 +by (insert abs_ge_self, blast intro: order_trans)
  28.778 +
  28.779 +lemma abs_le_D2: "abs a \<le> b ==> -a \<le> (b::'a::lordered_ab_group_abs)"
  28.780 +by (insert abs_le_D1 [of "-a"], simp)
  28.781 +
  28.782 +lemma abs_le_iff: "(abs a \<le> b) = (a \<le> b & -a \<le> (b::'a::lordered_ab_group_abs))"
  28.783 +by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
  28.784 +
  28.785 +lemma abs_triangle_ineq: "abs (a+b) \<le> abs a + abs (b::'a::lordered_ab_group_abs)"
  28.786 +proof -
  28.787 +  have g:"abs a + abs b = join (a+b) (join (-a-b) (join (-a+b) (a + (-b))))" (is "_=join ?m ?n")
  28.788 +    apply (simp add: abs_lattice add_meet_join_distribs join_aci)
  28.789 +    by (simp only: diff_minus)
  28.790 +  have a:"a+b <= join ?m ?n" by (simp add: meet_join_le)
  28.791 +  have b:"-a-b <= ?n" by (simp add: meet_join_le) 
  28.792 +  have c:"?n <= join ?m ?n" by (simp add: meet_join_le)
  28.793 +  from b c have d: "-a-b <= join ?m ?n" by simp
  28.794 +  have e:"-a-b = -(a+b)" by (simp add: diff_minus)
  28.795 +  from a d e have "abs(a+b) <= join ?m ?n" 
  28.796 +    by (drule_tac abs_leI, auto)
  28.797 +  with g[symmetric] show ?thesis by simp
  28.798 +qed
  28.799 +
  28.800 +lemma abs_diff_triangle_ineq:
  28.801 +     "\<bar>(a::'a::lordered_ab_group_abs) + b - (c+d)\<bar> \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>"
  28.802 +proof -
  28.803 +  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
  28.804 +  also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
  28.805 +  finally show ?thesis .
  28.806 +qed
  28.807 +
  28.808 +ML {*
  28.809 +val add_zero_left = thm"add_0";
  28.810 +val add_zero_right = thm"add_0_right";
  28.811 +*}
  28.812 +
  28.813 +ML {*
  28.814 +val add_assoc = thm "add_assoc";
  28.815 +val add_commute = thm "add_commute";
  28.816 +val add_left_commute = thm "add_left_commute";
  28.817 +val add_ac = thms "add_ac";
  28.818 +val mult_assoc = thm "mult_assoc";
  28.819 +val mult_commute = thm "mult_commute";
  28.820 +val mult_left_commute = thm "mult_left_commute";
  28.821 +val mult_ac = thms "mult_ac";
  28.822 +val add_0 = thm "add_0";
  28.823 +val mult_1_left = thm "mult_1_left";
  28.824 +val mult_1_right = thm "mult_1_right";
  28.825 +val mult_1 = thm "mult_1";
  28.826 +val add_left_imp_eq = thm "add_left_imp_eq";
  28.827 +val add_right_imp_eq = thm "add_right_imp_eq";
  28.828 +val add_imp_eq = thm "add_imp_eq";
  28.829 +val left_minus = thm "left_minus";
  28.830 +val diff_minus = thm "diff_minus";
  28.831 +val add_0_right = thm "add_0_right";
  28.832 +val add_left_cancel = thm "add_left_cancel";
  28.833 +val add_right_cancel = thm "add_right_cancel";
  28.834 +val right_minus = thm "right_minus";
  28.835 +val right_minus_eq = thm "right_minus_eq";
  28.836 +val minus_minus = thm "minus_minus";
  28.837 +val equals_zero_I = thm "equals_zero_I";
  28.838 +val minus_zero = thm "minus_zero";
  28.839 +val diff_self = thm "diff_self";
  28.840 +val diff_0 = thm "diff_0";
  28.841 +val diff_0_right = thm "diff_0_right";
  28.842 +val diff_minus_eq_add = thm "diff_minus_eq_add";
  28.843 +val neg_equal_iff_equal = thm "neg_equal_iff_equal";
  28.844 +val neg_equal_0_iff_equal = thm "neg_equal_0_iff_equal";
  28.845 +val neg_0_equal_iff_equal = thm "neg_0_equal_iff_equal";
  28.846 +val equation_minus_iff = thm "equation_minus_iff";
  28.847 +val minus_equation_iff = thm "minus_equation_iff";
  28.848 +val minus_add_distrib = thm "minus_add_distrib";
  28.849 +val minus_diff_eq = thm "minus_diff_eq";
  28.850 +val add_left_mono = thm "add_left_mono";
  28.851 +val add_le_imp_le_left = thm "add_le_imp_le_left";
  28.852 +val add_right_mono = thm "add_right_mono";
  28.853 +val add_mono = thm "add_mono";
  28.854 +val add_strict_left_mono = thm "add_strict_left_mono";
  28.855 +val add_strict_right_mono = thm "add_strict_right_mono";
  28.856 +val add_strict_mono = thm "add_strict_mono";
  28.857 +val add_less_le_mono = thm "add_less_le_mono";
  28.858 +val add_le_less_mono = thm "add_le_less_mono";
  28.859 +val add_less_imp_less_left = thm "add_less_imp_less_left";
  28.860 +val add_less_imp_less_right = thm "add_less_imp_less_right";
  28.861 +val add_less_cancel_left = thm "add_less_cancel_left";
  28.862 +val add_less_cancel_right = thm "add_less_cancel_right";
  28.863 +val add_le_cancel_left = thm "add_le_cancel_left";
  28.864 +val add_le_cancel_right = thm "add_le_cancel_right";
  28.865 +val add_le_imp_le_right = thm "add_le_imp_le_right";
  28.866 +val add_increasing = thm "add_increasing";
  28.867 +val le_imp_neg_le = thm "le_imp_neg_le";
  28.868 +val neg_le_iff_le = thm "neg_le_iff_le";
  28.869 +val neg_le_0_iff_le = thm "neg_le_0_iff_le";
  28.870 +val neg_0_le_iff_le = thm "neg_0_le_iff_le";
  28.871 +val neg_less_iff_less = thm "neg_less_iff_less";
  28.872 +val neg_less_0_iff_less = thm "neg_less_0_iff_less";
  28.873 +val neg_0_less_iff_less = thm "neg_0_less_iff_less";
  28.874 +val less_minus_iff = thm "less_minus_iff";
  28.875 +val minus_less_iff = thm "minus_less_iff";
  28.876 +val le_minus_iff = thm "le_minus_iff";
  28.877 +val minus_le_iff = thm "minus_le_iff";
  28.878 +val add_diff_eq = thm "add_diff_eq";
  28.879 +val diff_add_eq = thm "diff_add_eq";
  28.880 +val diff_eq_eq = thm "diff_eq_eq";
  28.881 +val eq_diff_eq = thm "eq_diff_eq";
  28.882 +val diff_diff_eq = thm "diff_diff_eq";
  28.883 +val diff_diff_eq2 = thm "diff_diff_eq2";
  28.884 +val diff_add_cancel = thm "diff_add_cancel";
  28.885 +val add_diff_cancel = thm "add_diff_cancel";
  28.886 +val less_iff_diff_less_0 = thm "less_iff_diff_less_0";
  28.887 +val diff_less_eq = thm "diff_less_eq";
  28.888 +val less_diff_eq = thm "less_diff_eq";
  28.889 +val diff_le_eq = thm "diff_le_eq";
  28.890 +val le_diff_eq = thm "le_diff_eq";
  28.891 +val compare_rls = thms "compare_rls";
  28.892 +val eq_iff_diff_eq_0 = thm "eq_iff_diff_eq_0";
  28.893 +val le_iff_diff_le_0 = thm "le_iff_diff_le_0";
  28.894 +val add_meet_distrib_left = thm "add_meet_distrib_left";
  28.895 +val add_join_distrib_left = thm "add_join_distrib_left";
  28.896 +val is_join_neg_meet = thm "is_join_neg_meet";
  28.897 +val is_meet_neg_join = thm "is_meet_neg_join";
  28.898 +val add_join_distrib_right = thm "add_join_distrib_right";
  28.899 +val add_meet_distrib_right = thm "add_meet_distrib_right";
  28.900 +val add_meet_join_distribs = thms "add_meet_join_distribs";
  28.901 +val join_eq_neg_meet = thm "join_eq_neg_meet";
  28.902 +val meet_eq_neg_join = thm "meet_eq_neg_join";
  28.903 +val add_eq_meet_join = thm "add_eq_meet_join";
  28.904 +val prts = thm "prts";
  28.905 +val zero_le_pprt = thm "zero_le_pprt";
  28.906 +val nprt_le_zero = thm "nprt_le_zero";
  28.907 +val le_eq_neg = thm "le_eq_neg";
  28.908 +val join_0_imp_0 = thm "join_0_imp_0";
  28.909 +val meet_0_imp_0 = thm "meet_0_imp_0";
  28.910 +val join_0_eq_0 = thm "join_0_eq_0";
  28.911 +val meet_0_eq_0 = thm "meet_0_eq_0";
  28.912 +val zero_le_double_add_iff_zero_le_single_add = thm "zero_le_double_add_iff_zero_le_single_add";
  28.913 +val double_add_le_zero_iff_single_add_le_zero = thm "double_add_le_zero_iff_single_add_le_zero";
  28.914 +val double_add_less_zero_iff_single_less_zero = thm "double_add_less_zero_iff_single_less_zero";
  28.915 +val abs_lattice = thm "abs_lattice";
  28.916 +val abs_zero = thm "abs_zero";
  28.917 +val abs_eq_0 = thm "abs_eq_0";
  28.918 +val abs_0_eq = thm "abs_0_eq";
  28.919 +val neg_meet_eq_join = thm "neg_meet_eq_join";
  28.920 +val neg_join_eq_meet = thm "neg_join_eq_meet";
  28.921 +val join_eq_if = thm "join_eq_if";
  28.922 +val abs_if_lattice = thm "abs_if_lattice";
  28.923 +val abs_ge_zero = thm "abs_ge_zero";
  28.924 +val abs_le_zero_iff = thm "abs_le_zero_iff";
  28.925 +val zero_less_abs_iff = thm "zero_less_abs_iff";
  28.926 +val abs_not_less_zero = thm "abs_not_less_zero";
  28.927 +val abs_ge_self = thm "abs_ge_self";
  28.928 +val abs_ge_minus_self = thm "abs_ge_minus_self";
  28.929 +val le_imp_join_eq = thm "le_imp_join_eq";
  28.930 +val ge_imp_join_eq = thm "ge_imp_join_eq";
  28.931 +val le_imp_meet_eq = thm "le_imp_meet_eq";
  28.932 +val ge_imp_meet_eq = thm "ge_imp_meet_eq";
  28.933 +val abs_prts = thm "abs_prts";
  28.934 +val abs_minus_cancel = thm "abs_minus_cancel";
  28.935 +val abs_idempotent = thm "abs_idempotent";
  28.936 +val zero_le_iff_zero_nprt = thm "zero_le_iff_zero_nprt";
  28.937 +val le_zero_iff_zero_pprt = thm "le_zero_iff_zero_pprt";
  28.938 +val le_zero_iff_pprt_id = thm "le_zero_iff_pprt_id";
  28.939 +val zero_le_iff_nprt_id = thm "zero_le_iff_nprt_id";
  28.940 +val iff2imp = thm "iff2imp";
  28.941 +val imp_abs_id = thm "imp_abs_id";
  28.942 +val imp_abs_neg_id = thm "imp_abs_neg_id";
  28.943 +val abs_leI = thm "abs_leI";
  28.944 +val le_minus_self_iff = thm "le_minus_self_iff";
  28.945 +val minus_le_self_iff = thm "minus_le_self_iff";
  28.946 +val abs_le_D1 = thm "abs_le_D1";
  28.947 +val abs_le_D2 = thm "abs_le_D2";
  28.948 +val abs_le_iff = thm "abs_le_iff";
  28.949 +val abs_triangle_ineq = thm "abs_triangle_ineq";
  28.950 +val abs_diff_triangle_ineq = thm "abs_diff_triangle_ineq";
  28.951 +*}
  28.952 +
  28.953 +end
    29.1 --- a/src/HOL/Power.thy	Tue May 11 14:00:02 2004 +0200
    29.2 +++ b/src/HOL/Power.thy	Tue May 11 20:11:08 2004 +0200
    29.3 @@ -11,7 +11,7 @@
    29.4  
    29.5  subsection{*Powers for Arbitrary (Semi)Rings*}
    29.6  
    29.7 -axclass ringpower \<subseteq> semiring, power
    29.8 +axclass ringpower \<subseteq> comm_semiring_1_cancel, power
    29.9    power_0 [simp]:   "a ^ 0       = 1"
   29.10    power_Suc: "a ^ (Suc n) = a * (a ^ n)"
   29.11  
   29.12 @@ -46,31 +46,31 @@
   29.13  done
   29.14  
   29.15  lemma zero_less_power:
   29.16 -     "0 < (a::'a::{ordered_semiring,ringpower}) ==> 0 < a^n"
   29.17 +     "0 < (a::'a::{ordered_semidom,ringpower}) ==> 0 < a^n"
   29.18  apply (induct_tac "n")
   29.19  apply (simp_all add: power_Suc zero_less_one mult_pos)
   29.20  done
   29.21  
   29.22  lemma zero_le_power:
   29.23 -     "0 \<le> (a::'a::{ordered_semiring,ringpower}) ==> 0 \<le> a^n"
   29.24 +     "0 \<le> (a::'a::{ordered_semidom,ringpower}) ==> 0 \<le> a^n"
   29.25  apply (simp add: order_le_less)
   29.26  apply (erule disjE)
   29.27  apply (simp_all add: zero_less_power zero_less_one power_0_left)
   29.28  done
   29.29  
   29.30  lemma one_le_power:
   29.31 -     "1 \<le> (a::'a::{ordered_semiring,ringpower}) ==> 1 \<le> a^n"
   29.32 +     "1 \<le> (a::'a::{ordered_semidom,ringpower}) ==> 1 \<le> a^n"
   29.33  apply (induct_tac "n")
   29.34  apply (simp_all add: power_Suc)
   29.35  apply (rule order_trans [OF _ mult_mono [of 1 _ 1]])
   29.36  apply (simp_all add: zero_le_one order_trans [OF zero_le_one])
   29.37  done
   29.38  
   29.39 -lemma gt1_imp_ge0: "1 < a ==> 0 \<le> (a::'a::ordered_semiring)"
   29.40 +lemma gt1_imp_ge0: "1 < a ==> 0 \<le> (a::'a::ordered_semidom)"
   29.41    by (simp add: order_trans [OF zero_le_one order_less_imp_le])
   29.42  
   29.43  lemma power_gt1_lemma:
   29.44 -  assumes gt1: "1 < (a::'a::{ordered_semiring,ringpower})"
   29.45 +  assumes gt1: "1 < (a::'a::{ordered_semidom,ringpower})"
   29.46    shows "1 < a * a^n"
   29.47  proof -
   29.48    have "1*1 < a*1" using gt1 by simp
   29.49 @@ -81,11 +81,11 @@
   29.50  qed
   29.51  
   29.52  lemma power_gt1:
   29.53 -     "1 < (a::'a::{ordered_semiring,ringpower}) ==> 1 < a ^ (Suc n)"
   29.54 +     "1 < (a::'a::{ordered_semidom,ringpower}) ==> 1 < a ^ (Suc n)"
   29.55  by (simp add: power_gt1_lemma power_Suc)
   29.56  
   29.57  lemma power_le_imp_le_exp:
   29.58 -  assumes gt1: "(1::'a::{ringpower,ordered_semiring}) < a"
   29.59 +  assumes gt1: "(1::'a::{ringpower,ordered_semidom}) < a"
   29.60    shows "!!n. a^m \<le> a^n ==> m \<le> n"
   29.61  proof (induct m)
   29.62    case 0
   29.63 @@ -109,26 +109,26 @@
   29.64  
   29.65  text{*Surely we can strengthen this? It holds for @{text "0<a<1"} too.*}
   29.66  lemma power_inject_exp [simp]:
   29.67 -     "1 < (a::'a::{ordered_semiring,ringpower}) ==> (a^m = a^n) = (m=n)"
   29.68 +     "1 < (a::'a::{ordered_semidom,ringpower}) ==> (a^m = a^n) = (m=n)"
   29.69    by (force simp add: order_antisym power_le_imp_le_exp)
   29.70  
   29.71  text{*Can relax the first premise to @{term "0<a"} in the case of the
   29.72  natural numbers.*}
   29.73  lemma power_less_imp_less_exp:
   29.74 -     "[| (1::'a::{ringpower,ordered_semiring}) < a; a^m < a^n |] ==> m < n"
   29.75 +     "[| (1::'a::{ringpower,ordered_semidom}) < a; a^m < a^n |] ==> m < n"
   29.76  by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"]
   29.77                power_le_imp_le_exp)
   29.78  
   29.79  
   29.80  lemma power_mono:
   29.81 -     "[|a \<le> b; (0::'a::{ringpower,ordered_semiring}) \<le> a|] ==> a^n \<le> b^n"
   29.82 +     "[|a \<le> b; (0::'a::{ringpower,ordered_semidom}) \<le> a|] ==> a^n \<le> b^n"
   29.83  apply (induct_tac "n")
   29.84  apply (simp_all add: power_Suc)
   29.85  apply (auto intro: mult_mono zero_le_power order_trans [of 0 a b])
   29.86  done
   29.87  
   29.88  lemma power_strict_mono [rule_format]:
   29.89 -     "[|a < b; (0::'a::{ringpower,ordered_semiring}) \<le> a|]
   29.90 +     "[|a < b; (0::'a::{ringpower,ordered_semidom}) \<le> a|]
   29.91        ==> 0 < n --> a^n < b^n"
   29.92  apply (induct_tac "n")
   29.93  apply (auto simp add: mult_strict_mono zero_le_power power_Suc
   29.94 @@ -136,7 +136,7 @@
   29.95  done
   29.96  
   29.97  lemma power_eq_0_iff [simp]:
   29.98 -     "(a^n = 0) = (a = (0::'a::{ordered_ring,ringpower}) & 0<n)"
   29.99 +     "(a^n = 0) = (a = (0::'a::{ordered_idom,ringpower}) & 0<n)"
  29.100  apply (induct_tac "n")
  29.101  apply (auto simp add: power_Suc zero_neq_one [THEN not_sym])
  29.102  done
  29.103 @@ -174,13 +174,13 @@
  29.104  apply assumption
  29.105  done
  29.106  
  29.107 -lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_ring,ringpower}) ^ n"
  29.108 +lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,ringpower}) ^ n"
  29.109  apply (induct_tac "n")
  29.110  apply (auto simp add: power_Suc abs_mult)
  29.111  done
  29.112  
  29.113  lemma zero_less_power_abs_iff [simp]:
  29.114 -     "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_ring,ringpower}) | n=0)"
  29.115 +     "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_idom,ringpower}) | n=0)"
  29.116  proof (induct "n")
  29.117    case 0
  29.118      show ?case by (simp add: zero_less_one)
  29.119 @@ -190,12 +190,12 @@
  29.120  qed
  29.121  
  29.122  lemma zero_le_power_abs [simp]:
  29.123 -     "(0::'a::{ordered_ring,ringpower}) \<le> (abs a)^n"
  29.124 +     "(0::'a::{ordered_idom,ringpower}) \<le> (abs a)^n"
  29.125  apply (induct_tac "n")
  29.126  apply (auto simp add: zero_le_one zero_le_power)
  29.127  done
  29.128  
  29.129 -lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{ring,ringpower}) ^ n"
  29.130 +lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{comm_ring_1,ringpower}) ^ n"
  29.131  proof -
  29.132    have "-a = (- 1) * a"  by (simp add: minus_mult_left [symmetric])
  29.133    thus ?thesis by (simp only: power_mult_distrib)
  29.134 @@ -203,14 +203,14 @@
  29.135  
  29.136  text{*Lemma for @{text power_strict_decreasing}*}
  29.137  lemma power_Suc_less:
  29.138 -     "[|(0::'a::{ordered_semiring,ringpower}) < a; a < 1|]
  29.139 +     "[|(0::'a::{ordered_semidom,ringpower}) < a; a < 1|]
  29.140        ==> a * a^n < a^n"
  29.141  apply (induct_tac n)
  29.142  apply (auto simp add: power_Suc mult_strict_left_mono)
  29.143  done
  29.144  
  29.145  lemma power_strict_decreasing:
  29.146 -     "[|n < N; 0 < a; a < (1::'a::{ordered_semiring,ringpower})|]
  29.147 +     "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,ringpower})|]
  29.148        ==> a^N < a^n"
  29.149  apply (erule rev_mp)
  29.150  apply (induct_tac "N")
  29.151 @@ -223,7 +223,7 @@
  29.152  
  29.153  text{*Proof resembles that of @{text power_strict_decreasing}*}
  29.154  lemma power_decreasing:
  29.155 -     "[|n \<le> N; 0 \<le> a; a \<le> (1::'a::{ordered_semiring,ringpower})|]
  29.156 +     "[|n \<le> N; 0 \<le> a; a \<le> (1::'a::{ordered_semidom,ringpower})|]
  29.157        ==> a^N \<le> a^n"
  29.158  apply (erule rev_mp)
  29.159  apply (induct_tac "N")
  29.160 @@ -235,13 +235,13 @@
  29.161  done
  29.162  
  29.163  lemma power_Suc_less_one:
  29.164 -     "[| 0 < a; a < (1::'a::{ordered_semiring,ringpower}) |] ==> a ^ Suc n < 1"
  29.165 +     "[| 0 < a; a < (1::'a::{ordered_semidom,ringpower}) |] ==> a ^ Suc n < 1"
  29.166  apply (insert power_strict_decreasing [of 0 "Suc n" a], simp)
  29.167  done
  29.168  
  29.169  text{*Proof again resembles that of @{text power_strict_decreasing}*}
  29.170  lemma power_increasing:
  29.171 -     "[|n \<le> N; (1::'a::{ordered_semiring,ringpower}) \<le> a|] ==> a^n \<le> a^N"
  29.172 +     "[|n \<le> N; (1::'a::{ordered_semidom,ringpower}) \<le> a|] ==> a^n \<le> a^N"
  29.173  apply (erule rev_mp)
  29.174  apply (induct_tac "N")
  29.175  apply (auto simp add: power_Suc le_Suc_eq)
  29.176 @@ -253,13 +253,13 @@
  29.177  
  29.178  text{*Lemma for @{text power_strict_increasing}*}
  29.179  lemma power_less_power_Suc:
  29.180 -     "(1::'a::{ordered_semiring,ringpower}) < a ==> a^n < a * a^n"
  29.181 +     "(1::'a::{ordered_semidom,ringpower}) < a ==> a^n < a * a^n"
  29.182  apply (induct_tac n)
  29.183  apply (auto simp add: power_Suc mult_strict_left_mono order_less_trans [OF zero_less_one])
  29.184  done
  29.185  
  29.186  lemma power_strict_increasing:
  29.187 -     "[|n < N; (1::'a::{ordered_semiring,ringpower}) < a|] ==> a^n < a^N"
  29.188 +     "[|n < N; (1::'a::{ordered_semidom,ringpower}) < a|] ==> a^n < a^N"
  29.189  apply (erule rev_mp)
  29.190  apply (induct_tac "N")
  29.191  apply (auto simp add: power_less_power_Suc power_Suc less_Suc_eq)
  29.192 @@ -272,7 +272,7 @@
  29.193  
  29.194  lemma power_le_imp_le_base:
  29.195    assumes le: "a ^ Suc n \<le> b ^ Suc n"
  29.196 -      and xnonneg: "(0::'a::{ordered_semiring,ringpower}) \<le> a"
  29.197 +      and xnonneg: "(0::'a::{ordered_semidom,ringpower}) \<le> a"
  29.198        and ynonneg: "0 \<le> b"
  29.199    shows "a \<le> b"
  29.200   proof (rule ccontr)
  29.201 @@ -286,7 +286,7 @@
  29.202  
  29.203  lemma power_inject_base:
  29.204       "[| a ^ Suc n = b ^ Suc n; 0 \<le> a; 0 \<le> b |]
  29.205 -      ==> a = (b::'a::{ordered_semiring,ringpower})"
  29.206 +      ==> a = (b::'a::{ordered_semidom,ringpower})"
  29.207  by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym)
  29.208  
  29.209  
    30.1 --- a/src/HOL/Presburger.thy	Tue May 11 14:00:02 2004 +0200
    30.2 +++ b/src/HOL/Presburger.thy	Tue May 11 20:11:08 2004 +0200
    30.3 @@ -704,7 +704,7 @@
    30.4      have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast
    30.5      also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)"
    30.6        using minus[THEN spec, of "x - i * d"]
    30.7 -      by (simp add:int_distrib Ring_and_Field.diff_diff_eq[symmetric])
    30.8 +      by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric])
    30.9      ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast
   30.10    qed
   30.11  qed
    31.1 --- a/src/HOL/Real/PReal.thy	Tue May 11 14:00:02 2004 +0200
    31.2 +++ b/src/HOL/Real/PReal.thy	Tue May 11 20:11:08 2004 +0200
    31.3 @@ -15,7 +15,7 @@
    31.4  
    31.5  text{*As a special case, the sum of two positives is positive.  One of the
    31.6  premises could be weakened to the relation @{text "\<le>"}.*}
    31.7 -lemma pos_add_strict: "[|0<a; b<c|] ==> b < a + (c::'a::ordered_semiring)"
    31.8 +lemma pos_add_strict: "[|0<a; b<c|] ==> b < a + (c::'a::ordered_semidom)"
    31.9  by (insert add_strict_mono [of 0 a b c], simp)
   31.10  
   31.11  lemma interval_empty_iff:
    32.1 --- a/src/HOL/Real/RealDef.thy	Tue May 11 14:00:02 2004 +0200
    32.2 +++ b/src/HOL/Real/RealDef.thy	Tue May 11 20:11:08 2004 +0200
    32.3 @@ -169,7 +169,7 @@
    32.4  lemma real_add_zero_left: "(0::real) + z = z"
    32.5  by (cases z, simp add: real_add real_zero_def preal_add_ac)
    32.6  
    32.7 -instance real :: plus_ac0
    32.8 +instance real :: comm_monoid_add
    32.9    by (intro_classes,
   32.10        (assumption | 
   32.11         rule real_add_commute real_add_assoc real_add_zero_left)+)
   32.12 @@ -281,9 +281,6 @@
   32.13  instance real :: field
   32.14  proof
   32.15    fix x y z :: real
   32.16 -  show "(x + y) + z = x + (y + z)" by (rule real_add_assoc)
   32.17 -  show "x + y = y + x" by (rule real_add_commute)
   32.18 -  show "0 + x = x" by simp
   32.19    show "- x + x = 0" by (rule real_add_minus_left)
   32.20    show "x - y = x + (-y)" by (simp add: real_diff_def)
   32.21    show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
   32.22 @@ -312,7 +309,7 @@
   32.23          minus_mult_left [symmetric, simp]
   32.24  
   32.25  lemma real_mult_1_right: "z * (1::real) = z"
   32.26 -  by (rule Ring_and_Field.mult_1_right)
   32.27 +  by (rule OrderedGroup.mult_1_right)
   32.28  
   32.29  
   32.30  subsection{*The @{text "\<le>"} Ordering*}
   32.31 @@ -554,11 +551,11 @@
   32.32  by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
   32.33  
   32.34  lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)"
   32.35 -  by (rule Ring_and_Field.add_less_le_mono)
   32.36 +  by (rule OrderedGroup.add_less_le_mono)
   32.37  
   32.38  lemma real_add_le_less_mono:
   32.39       "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
   32.40 -  by (rule Ring_and_Field.add_le_less_mono)
   32.41 +  by (rule OrderedGroup.add_le_less_mono)
   32.42  
   32.43  lemma real_le_square [simp]: "(0::real) \<le> x*x"
   32.44   by (rule Ring_and_Field.zero_le_square)
   32.45 @@ -597,7 +594,7 @@
   32.46  
   32.47  text{*FIXME: delete or at least combine the next two lemmas*}
   32.48  lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"
   32.49 -apply (drule Ring_and_Field.equals_zero_I [THEN sym])
   32.50 +apply (drule OrderedGroup.equals_zero_I [THEN sym])
   32.51  apply (cut_tac x = y in real_le_square) 
   32.52  apply (auto, drule order_antisym, auto)
   32.53  done
   32.54 @@ -848,7 +845,7 @@
   32.55  by (force simp add: Ring_and_Field.abs_less_iff)
   32.56  
   32.57  lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
   32.58 -by (force simp add: Ring_and_Field.abs_le_iff)
   32.59 +by (force simp add: OrderedGroup.abs_le_iff)
   32.60  
   32.61  (*FIXME: used only once, in SEQ.ML*)
   32.62  lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
   32.63 @@ -892,7 +889,7 @@
   32.64  val abs_minus_eqI2 = thm"abs_minus_eqI2";
   32.65  val abs_ge_zero = thm"abs_ge_zero";
   32.66  val abs_idempotent = thm"abs_idempotent";
   32.67 -val abs_zero_iff = thm"abs_zero_iff";
   32.68 +val abs_eq_0 = thm"abs_eq_0";
   32.69  val abs_ge_self = thm"abs_ge_self";
   32.70  val abs_ge_minus_self = thm"abs_ge_minus_self";
   32.71  val abs_mult = thm"abs_mult";
    33.1 --- a/src/HOL/Ring_and_Field.thy	Tue May 11 14:00:02 2004 +0200
    33.2 +++ b/src/HOL/Ring_and_Field.thy	Tue May 11 20:11:08 2004 +0200
    33.3 @@ -2,286 +2,132 @@
    33.4      ID:      $Id$
    33.5      Author:  Gertrud Bauer and Markus Wenzel, TU Muenchen
    33.6               Lawrence C Paulson, University of Cambridge
    33.7 +             Revised and splitted into Ring_and_Field.thy and Group.thy 
    33.8 +             by Steven Obua, TU Muenchen, in May 2004
    33.9      License: GPL (GNU GENERAL PUBLIC LICENSE)
   33.10  *)
   33.11  
   33.12 -header {* Ring and field structures *}
   33.13 +header {* (Ordered) Rings and Fields *}
   33.14  
   33.15 -theory Ring_and_Field = Inductive:
   33.16 +theory Ring_and_Field = OrderedGroup:
   33.17  
   33.18 -subsection {* Abstract algebraic structures *}
   33.19 +text {*
   33.20 +  The theory of partially ordered rings is taken from the books:
   33.21 +  \begin{itemize}
   33.22 +  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
   33.23 +  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
   33.24 +  \end{itemize}
   33.25 +  Most of the used notions can also be looked up in 
   33.26 +  \begin{itemize}
   33.27 +  \item \emph{www.mathworld.com} by Eric Weisstein et. al.
   33.28 +  \item \emph{Algebra I} by van der Waerden, Springer.
   33.29 +  \end{itemize}
   33.30 +*}
   33.31  
   33.32 -subsection {* Types Classes @{text plus_ac0} and @{text times_ac1} *}
   33.33 +axclass semiring \<subseteq> ab_semigroup_add, semigroup_mult
   33.34 +  left_distrib: "(a + b) * c = a * c + b * c"
   33.35 +  right_distrib: "a * (b + c) = a * b + a * c"
   33.36  
   33.37 -axclass plus_ac0 \<subseteq> plus, zero
   33.38 -  commute:     "x + y = y + x"
   33.39 -  assoc:       "(x + y) + z = x + (y + z)"
   33.40 -  zero [simp]: "0 + x = x"
   33.41 +axclass semiring_0 \<subseteq> semiring, comm_monoid_add
   33.42  
   33.43 -lemma plus_ac0_left_commute: "x + (y+z) = y + ((x+z)::'a::plus_ac0)"
   33.44 -by(rule mk_left_commute[of "op +",OF plus_ac0.assoc plus_ac0.commute])
   33.45 +axclass comm_semiring \<subseteq> ab_semigroup_add, ab_semigroup_mult  
   33.46 +  mult_commute: "a * b = b * a"
   33.47 +  distrib: "(a + b) * c = a * c + b * c"
   33.48  
   33.49 -lemma plus_ac0_zero_right [simp]: "x + 0 = (x ::'a::plus_ac0)"
   33.50 -apply (rule plus_ac0.commute [THEN trans])
   33.51 -apply (rule plus_ac0.zero)
   33.52 -done
   33.53 -
   33.54 -lemmas plus_ac0 = plus_ac0.assoc plus_ac0.commute plus_ac0_left_commute
   33.55 -                  plus_ac0.zero plus_ac0_zero_right
   33.56 -
   33.57 -axclass times_ac1 \<subseteq> times, one
   33.58 -  commute:     "x * y = y * x"
   33.59 -  assoc:       "(x * y) * z = x * (y * z)"
   33.60 -  one [simp]:  "1 * x = x"
   33.61 -
   33.62 -theorem times_ac1_left_commute: "(x::'a::times_ac1) * ((y::'a) * z) =
   33.63 -  y * (x * z)"
   33.64 -by(rule mk_left_commute[of "op *",OF times_ac1.assoc times_ac1.commute])
   33.65 -
   33.66 -theorem times_ac1_one_right [simp]: "(x::'a::times_ac1) * 1 = x"
   33.67 -proof -
   33.68 -  have "x * 1 = 1 * x"
   33.69 -    by (rule times_ac1.commute)
   33.70 -  also have "... = x"
   33.71 -    by (rule times_ac1.one)
   33.72 -  finally show ?thesis .
   33.73 +instance comm_semiring \<subseteq> semiring
   33.74 +proof
   33.75 +  fix a b c :: 'a
   33.76 +  show "(a + b) * c = a * c + b * c" by (simp add: distrib)
   33.77 +  have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
   33.78 +  also have "... = b * a + c * a" by (simp only: distrib)
   33.79 +  also have "... = a * b + a * c" by (simp add: mult_ac)
   33.80 +  finally show "a * (b + c) = a * b + a * c" by blast
   33.81  qed
   33.82  
   33.83 -theorems times_ac1 = times_ac1.assoc times_ac1.commute times_ac1_left_commute
   33.84 -  times_ac1.one times_ac1_one_right
   33.85 +axclass comm_semiring_0 \<subseteq> comm_semiring, comm_monoid_add
   33.86  
   33.87 +instance comm_semiring_0 \<subseteq> semiring_0 ..
   33.88  
   33.89 -text{*This class is the same as @{text plus_ac0}, while using the same axiom
   33.90 -names as the other axclasses.*}
   33.91 -axclass abelian_semigroup \<subseteq> zero, plus
   33.92 -  add_assoc: "(a + b) + c = a + (b + c)"
   33.93 -  add_commute: "a + b = b + a"
   33.94 -  add_0 [simp]: "0 + a = a"
   33.95 -
   33.96 -text{*This class underlies both @{text semiring} and @{text ring}*}
   33.97 -axclass almost_semiring \<subseteq> abelian_semigroup, one, times
   33.98 -  mult_assoc: "(a * b) * c = a * (b * c)"
   33.99 -  mult_commute: "a * b = b * a"
  33.100 -  mult_1 [simp]: "1 * a = a"
  33.101 -
  33.102 -  left_distrib: "(a + b) * c = a * c + b * c"
  33.103 +axclass axclass_0_neq_1 \<subseteq> zero, one
  33.104    zero_neq_one [simp]: "0 \<noteq> 1"
  33.105  
  33.106 +axclass semiring_1 \<subseteq> axclass_0_neq_1, semiring_0, monoid_mult
  33.107  
  33.108 -axclass semiring \<subseteq> almost_semiring
  33.109 -  add_left_imp_eq: "a + b = a + c ==> b=c"
  33.110 -    --{*This axiom is needed for semirings only: for rings, etc., it is
  33.111 -        redundant. Including it allows many more of the following results
  33.112 -        to be proved for semirings too.*}
  33.113 +axclass comm_semiring_1 \<subseteq> axclass_0_neq_1, comm_semiring_0, comm_monoid_mult (* previously almost_semiring *)
  33.114  
  33.115 -instance abelian_semigroup \<subseteq> plus_ac0
  33.116 -proof
  33.117 -  fix x y z :: 'a
  33.118 -  show "x + y = y + x" by (rule add_commute)
  33.119 -  show "x + y + z = x + (y + z)" by (rule add_assoc)
  33.120 -  show "0+x = x" by (rule add_0)
  33.121 -qed
  33.122 +instance comm_semiring_1 \<subseteq> semiring_1 ..
  33.123  
  33.124 -instance almost_semiring \<subseteq> times_ac1
  33.125 -proof
  33.126 -  fix x y z :: 'a
  33.127 -  show "x * y = y * x" by (rule mult_commute)
  33.128 -  show "x * y * z = x * (y * z)" by (rule mult_assoc)
  33.129 -  show "1*x = x" by simp
  33.130 -qed
  33.131 +axclass axclass_no_zero_divisors \<subseteq> zero, times
  33.132 +  no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
  33.133  
  33.134 -axclass abelian_group \<subseteq> abelian_semigroup, minus
  33.135 -   left_minus [simp]: "-a + a = 0"
  33.136 -   diff_minus: "a - b = a + -b"
  33.137 +axclass comm_semiring_1_cancel \<subseteq> comm_semiring_1, cancel_ab_semigroup_add (* previously semiring *)
  33.138  
  33.139 -axclass ring \<subseteq> almost_semiring, abelian_group
  33.140 +axclass ring \<subseteq> semiring, ab_group_add
  33.141  
  33.142 -text{*Proving axiom @{text add_left_imp_eq} makes all @{text semiring}
  33.143 -      theorems available to members of @{term ring} *}
  33.144 -instance ring \<subseteq> semiring
  33.145 -proof
  33.146 -  fix a b c :: 'a
  33.147 -  assume "a + b = a + c"
  33.148 -  hence  "-a + a + b = -a + a + c" by (simp only: add_assoc)
  33.149 -  thus "b = c" by simp
  33.150 -qed
  33.151 +instance ring \<subseteq> semiring_0 ..
  33.152  
  33.153 -text{*This class underlies @{text ordered_semiring} and @{text ordered_ring}*}
  33.154 -axclass almost_ordered_semiring \<subseteq> semiring, linorder
  33.155 -  add_left_mono: "a \<le> b ==> c + a \<le> c + b"
  33.156 -  mult_strict_left_mono: "a < b ==> 0 < c ==> c * a < c * b"
  33.157 +axclass comm_ring \<subseteq> comm_semiring_0, ab_group_add
  33.158  
  33.159 -axclass ordered_semiring \<subseteq> almost_ordered_semiring
  33.160 -  zero_less_one [simp]: "0 < 1" --{*This too is needed for semirings only.*}
  33.161 +instance comm_ring \<subseteq> ring ..
  33.162  
  33.163 -axclass ordered_ring \<subseteq> almost_ordered_semiring, ring
  33.164 -  abs_if: "\<bar>a\<bar> = (if a < 0 then -a else a)"
  33.165 +instance comm_ring \<subseteq> comm_semiring_0 ..
  33.166  
  33.167 -axclass field \<subseteq> ring, inverse
  33.168 +axclass ring_1 \<subseteq> ring, semiring_1
  33.169 +
  33.170 +axclass comm_ring_1 \<subseteq> comm_ring, comm_semiring_1 (* previously ring *)
  33.171 +
  33.172 +instance comm_ring_1 \<subseteq> ring_1 ..
  33.173 +
  33.174 +instance comm_ring_1 \<subseteq> comm_semiring_1_cancel ..
  33.175 +
  33.176 +axclass idom \<subseteq> comm_ring_1, axclass_no_zero_divisors
  33.177 +
  33.178 +axclass field \<subseteq> comm_ring_1, inverse
  33.179    left_inverse [simp]: "a \<noteq> 0 ==> inverse a * a = 1"
  33.180    divide_inverse:      "a / b = a * inverse b"
  33.181  
  33.182 -axclass ordered_field \<subseteq> ordered_ring, field
  33.183 +lemma mult_zero_left [simp]: "0 * a = (0::'a::{semiring_0, cancel_semigroup_add})"
  33.184 +proof -
  33.185 +  have "0*a + 0*a = 0*a + 0"
  33.186 +    by (simp add: left_distrib [symmetric])
  33.187 +  thus ?thesis 
  33.188 +    by (simp only: add_left_cancel)
  33.189 +qed
  33.190  
  33.191 +lemma mult_zero_right [simp]: "a * 0 = (0::'a::{semiring_0, cancel_semigroup_add})"
  33.192 +proof -
  33.193 +  have "a*0 + a*0 = a*0 + 0"
  33.194 +    by (simp add: right_distrib [symmetric])
  33.195 +  thus ?thesis 
  33.196 +    by (simp only: add_left_cancel)
  33.197 +qed
  33.198 +
  33.199 +lemma field_mult_eq_0_iff [simp]: "(a*b = (0::'a::field)) = (a = 0 | b = 0)"
  33.200 +proof cases
  33.201 +  assume "a=0" thus ?thesis by simp
  33.202 +next
  33.203 +  assume anz [simp]: "a\<noteq>0"
  33.204 +  { assume "a * b = 0"
  33.205 +    hence "inverse a * (a * b) = 0" by simp
  33.206 +    hence "b = 0"  by (simp (no_asm_use) add: mult_assoc [symmetric])}
  33.207 +  thus ?thesis by force
  33.208 +qed
  33.209 +
  33.210 +instance field \<subseteq> idom
  33.211 +by (intro_classes, simp)
  33.212 +  
  33.213  axclass division_by_zero \<subseteq> zero, inverse
  33.214    inverse_zero [simp]: "inverse 0 = 0"
  33.215  
  33.216 -
  33.217 -subsection {* Derived Rules for Addition *}
  33.218 -
  33.219 -lemma add_0_right [simp]: "a + 0 = (a::'a::plus_ac0)"
  33.220 -proof -
  33.221 -  have "a + 0 = 0 + a" by (rule plus_ac0.commute)
  33.222 -  also have "... = a" by simp
  33.223 -  finally show ?thesis .
  33.224 -qed
  33.225 -
  33.226 -lemma add_left_commute: "a + (b + c) = b + (a + (c::'a::plus_ac0))"
  33.227 -  by (rule mk_left_commute [of "op +", OF plus_ac0.assoc plus_ac0.commute])
  33.228 -
  33.229 -theorems add_ac = add_assoc add_commute add_left_commute
  33.230 -
  33.231 -lemma right_minus [simp]: "a + -(a::'a::abelian_group) = 0"
  33.232 -proof -
  33.233 -  have "a + -a = -a + a" by (simp add: add_ac)
  33.234 -  also have "... = 0" by simp
  33.235 -  finally show ?thesis .
  33.236 -qed
  33.237 -
  33.238 -lemma right_minus_eq: "(a - b = 0) = (a = (b::'a::abelian_group))"
  33.239 -proof
  33.240 -  have "a = a - b + b" by (simp add: diff_minus add_ac)
  33.241 -  also assume "a - b = 0"
  33.242 -  finally show "a = b" by simp
  33.243 -next
  33.244 -  assume "a = b"
  33.245 -  thus "a - b = 0" by (simp add: diff_minus)
  33.246 -qed
  33.247 -
  33.248 -lemma add_left_cancel [simp]:
  33.249 -     "(a + b = a + c) = (b = (c::'a::semiring))"
  33.250 -by (blast dest: add_left_imp_eq) 
  33.251 -
  33.252 -lemma add_right_cancel [simp]:
  33.253 -     "(b + a = c + a) = (b = (c::'a::semiring))"
  33.254 -  by (simp add: add_commute)
  33.255 -
  33.256 -lemma minus_minus [simp]: "- (- (a::'a::abelian_group)) = a" 
  33.257 -apply (rule right_minus_eq [THEN iffD1]) 
  33.258 -apply (simp add: diff_minus) 
  33.259 -done
  33.260 -
  33.261 -lemma equals_zero_I: "a+b = 0 ==> -a = (b::'a::abelian_group)"
  33.262 -apply (rule right_minus_eq [THEN iffD1, symmetric])
  33.263 -apply (simp add: diff_minus add_commute) 
  33.264 -done
  33.265 -
  33.266 -lemma minus_zero [simp]: "- 0 = (0::'a::abelian_group)"
  33.267 -by (simp add: equals_zero_I)
  33.268 -
  33.269 -lemma diff_self [simp]: "a - (a::'a::abelian_group) = 0"
  33.270 -  by (simp add: diff_minus)
  33.271 -
  33.272 -lemma diff_0 [simp]: "(0::'a::abelian_group) - a = -a"
  33.273 -by (simp add: diff_minus)
  33.274 -
  33.275 -lemma diff_0_right [simp]: "a - (0::'a::abelian_group) = a" 
  33.276 -by (simp add: diff_minus)
  33.277 -
  33.278 -lemma diff_minus_eq_add [simp]: "a - - b = a + (b::'a::abelian_group)"
  33.279 -by (simp add: diff_minus)
  33.280 -
  33.281 -lemma neg_equal_iff_equal [simp]: "(-a = -b) = (a = (b::'a::abelian_group))" 
  33.282 -proof 
  33.283 -  assume "- a = - b"
  33.284 -  hence "- (- a) = - (- b)"
  33.285 -    by simp
  33.286 -  thus "a=b" by simp
  33.287 -next
  33.288 -  assume "a=b"
  33.289 -  thus "-a = -b" by simp
  33.290 -qed
  33.291 -
  33.292 -lemma neg_equal_0_iff_equal [simp]: "(-a = 0) = (a = (0::'a::abelian_group))"
  33.293 -by (subst neg_equal_iff_equal [symmetric], simp)
  33.294 -
  33.295 -lemma neg_0_equal_iff_equal [simp]: "(0 = -a) = (0 = (a::'a::abelian_group))"
  33.296 -by (subst neg_equal_iff_equal [symmetric], simp)
  33.297 -
  33.298 -lemma add_minus_self [simp]: "a + b - b = (a::'a::abelian_group)"; 
  33.299 -  by (simp add: diff_minus add_assoc)
  33.300 -
  33.301 -lemma add_minus_self_left [simp]:  "a + (b - a)  = (b::'a::abelian_group)";
  33.302 -by (simp add: diff_minus add_left_commute [of a]) 
  33.303 -
  33.304 -lemma add_minus_self_right  [simp]:  "a + b - a  = (b::'a::abelian_group)";
  33.305 -by (simp add: diff_minus add_left_commute [of a] add_assoc) 
  33.306 -
  33.307 -lemma minus_add_self [simp]: "a - b + b = (a::'a::abelian_group)"; 
  33.308 -by (simp add: diff_minus add_assoc) 
  33.309 -
  33.310 -text{*The next two equations can make the simplifier loop!*}
  33.311 -
  33.312 -lemma equation_minus_iff: "(a = - b) = (b = - (a::'a::abelian_group))"
  33.313 -proof -
  33.314 -  have "(- (-a) = - b) = (- a = b)" by (rule neg_equal_iff_equal)
  33.315 -  thus ?thesis by (simp add: eq_commute)
  33.316 -qed
  33.317 -
  33.318 -lemma minus_equation_iff: "(- a = b) = (- (b::'a::abelian_group) = a)"
  33.319 -proof -
  33.320 -  have "(- a = - (-b)) = (a = -b)" by (rule neg_equal_iff_equal)
  33.321 -  thus ?thesis by (simp add: eq_commute)
  33.322 -qed
  33.323 -
  33.324 -
  33.325 -subsection {* Derived rules for multiplication *}
  33.326 -
  33.327 -lemma mult_1_right [simp]: "a * (1::'a::almost_semiring) = a"
  33.328 -proof -
  33.329 -  have "a * 1 = 1 * a" by (simp add: mult_commute)
  33.330 -  also have "... = a" by simp
  33.331 -  finally show ?thesis .
  33.332 -qed
  33.333 -
  33.334 -lemma mult_left_commute: "a * (b * c) = b * (a * (c::'a::almost_semiring))"
  33.335 -  by (rule mk_left_commute [of "op *", OF mult_assoc mult_commute])
  33.336 -
  33.337 -theorems mult_ac = mult_assoc mult_commute mult_left_commute
  33.338 -
  33.339 -lemma mult_zero_left [simp]: "0 * a = (0::'a::semiring)"
  33.340 -proof -
  33.341 -  have "0*a + 0*a = 0*a + 0"
  33.342 -    by (simp add: left_distrib [symmetric])
  33.343 -  thus ?thesis by (simp only: add_left_cancel)
  33.344 -qed
  33.345 -
  33.346 -lemma mult_zero_right [simp]: "a * 0 = (0::'a::semiring)"
  33.347 -  by (simp add: mult_commute)
  33.348 -
  33.349 -
  33.350  subsection {* Distribution rules *}
  33.351  
  33.352 -lemma right_distrib: "a * (b + c) = a * b + a * (c::'a::almost_semiring)"
  33.353 -proof -
  33.354 -  have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
  33.355 -  also have "... = b * a + c * a" by (simp only: left_distrib)
  33.356 -  also have "... = a * b + a * c" by (simp add: mult_ac)
  33.357 -  finally show ?thesis .
  33.358 -qed
  33.359 -
  33.360  theorems ring_distrib = right_distrib left_distrib
  33.361  
  33.362  text{*For the @{text combine_numerals} simproc*}
  33.363  lemma combine_common_factor:
  33.364 -     "a*e + (b*e + c) = (a+b)*e + (c::'a::almost_semiring)"
  33.365 +     "a*e + (b*e + c) = (a+b)*e + (c::'a::semiring)"
  33.366  by (simp add: left_distrib add_ac)
  33.367  
  33.368 -lemma minus_add_distrib [simp]: "- (a + b) = -a + -(b::'a::abelian_group)"
  33.369 -apply (rule equals_zero_I)
  33.370 -apply (simp add: plus_ac0) 
  33.371 -done
  33.372 -
  33.373  lemma minus_mult_left: "- (a * b) = (-a) * (b::'a::ring)"
  33.374  apply (rule equals_zero_I)
  33.375  apply (simp add: left_distrib [symmetric]) 
  33.376 @@ -303,237 +149,86 @@
  33.377                minus_mult_left [symmetric] minus_mult_right [symmetric]) 
  33.378  
  33.379  lemma left_diff_distrib: "(a - b) * c = a * c - b * (c::'a::ring)"
  33.380 -by (simp add: mult_commute [of _ c] right_diff_distrib) 
  33.381 +by (simp add: left_distrib diff_minus 
  33.382 +              minus_mult_left [symmetric] minus_mult_right [symmetric]) 
  33.383  
  33.384 -lemma minus_diff_eq [simp]: "- (a - b) = b - (a::'a::ring)"
  33.385 -by (simp add: diff_minus add_commute) 
  33.386 +axclass pordered_semiring \<subseteq> semiring_0, pordered_ab_semigroup_add 
  33.387 +  mult_left_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b"
  33.388 +  mult_right_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> a * c <= b * c"
  33.389  
  33.390 +axclass pordered_cancel_semiring \<subseteq> pordered_semiring, cancel_ab_semigroup_add
  33.391  
  33.392 -subsection {* Ordering Rules for Addition *}
  33.393 +axclass ordered_semiring_strict \<subseteq> semiring_0, ordered_cancel_ab_semigroup_add
  33.394 +  mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
  33.395 +  mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
  33.396  
  33.397 -lemma add_right_mono: "a \<le> (b::'a::almost_ordered_semiring) ==> a + c \<le> b + c"
  33.398 -by (simp add: add_commute [of _ c] add_left_mono)
  33.399 -
  33.400 -text {* non-strict, in both arguments *}
  33.401 -lemma add_mono:
  33.402 -     "[|a \<le> b;  c \<le> d|] ==> a + c \<le> b + (d::'a::almost_ordered_semiring)"
  33.403 -  apply (erule add_right_mono [THEN order_trans])
  33.404 -  apply (simp add: add_commute add_left_mono)
  33.405 -  done
  33.406 -
  33.407 -lemma add_strict_left_mono:
  33.408 -     "a < b ==> c + a < c + (b::'a::almost_ordered_semiring)"
  33.409 - by (simp add: order_less_le add_left_mono) 
  33.410 -
  33.411 -lemma add_strict_right_mono:
  33.412 -     "a < b ==> a + c < b + (c::'a::almost_ordered_semiring)"
  33.413 - by (simp add: add_commute [of _ c] add_strict_left_mono)
  33.414 -
  33.415 -text{*Strict monotonicity in both arguments*}
  33.416 -lemma add_strict_mono: "[|a<b; c<d|] ==> a + c < b + (d::'a::almost_ordered_semiring)"
  33.417 -apply (erule add_strict_right_mono [THEN order_less_trans])
  33.418 -apply (erule add_strict_left_mono)
  33.419 +instance ordered_semiring_strict \<subseteq> pordered_cancel_semiring
  33.420 +apply intro_classes
  33.421 +apply (case_tac "a < b & 0 < c")
  33.422 +apply (auto simp add: mult_strict_left_mono order_less_le)
  33.423 +apply (auto simp add: mult_strict_left_mono order_le_less)
  33.424 +apply (simp add: mult_strict_right_mono)
  33.425  done
  33.426  
  33.427 -lemma add_less_le_mono:
  33.428 -     "[| a<b; c\<le>d |] ==> a + c < b + (d::'a::almost_ordered_semiring)"
  33.429 -apply (erule add_strict_right_mono [THEN order_less_le_trans])
  33.430 -apply (erule add_left_mono) 
  33.431 +axclass pordered_comm_semiring \<subseteq> comm_semiring_0, pordered_ab_semigroup_add
  33.432 +  mult_mono: "a <= b \<Longrightarrow> 0 <= c \<Longrightarrow> c * a <= c * b"
  33.433 +
  33.434 +axclass pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring, cancel_ab_semigroup_add
  33.435 +
  33.436 +instance pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring ..
  33.437 +
  33.438 +axclass ordered_comm_semiring_strict \<subseteq> comm_semiring_0, ordered_cancel_ab_semigroup_add
  33.439 +  mult_strict_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
  33.440 +
  33.441 +instance pordered_comm_semiring \<subseteq> pordered_semiring
  33.442 +by (intro_classes, insert mult_mono, simp_all add: mult_commute, blast+)
  33.443 +
  33.444 +instance pordered_cancel_comm_semiring \<subseteq> pordered_cancel_semiring ..
  33.445 +
  33.446 +instance ordered_comm_semiring_strict \<subseteq> ordered_semiring_strict
  33.447 +by (intro_classes, insert mult_strict_mono, simp_all add: mult_commute, blast+)
  33.448 +
  33.449 +instance ordered_comm_semiring_strict \<subseteq> pordered_cancel_comm_semiring
  33.450 +apply (intro_classes)
  33.451 +apply (case_tac "a < b & 0 < c")
  33.452 +apply (auto simp add: mult_strict_left_mono order_less_le)
  33.453 +apply (auto simp add: mult_strict_left_mono order_le_less)
  33.454  done
  33.455  
  33.456 -lemma add_le_less_mono:
  33.457 -     "[| a\<le>b; c<d |] ==> a + c < b + (d::'a::almost_ordered_semiring)"
  33.458 -apply (erule add_right_mono [THEN order_le_less_trans])
  33.459 -apply (erule add_strict_left_mono) 
  33.460 -done
  33.461 +axclass pordered_ring \<subseteq> ring, pordered_semiring 
  33.462  
  33.463 -lemma add_less_imp_less_left:
  33.464 -      assumes less: "c + a < c + b"  shows "a < (b::'a::almost_ordered_semiring)"
  33.465 -proof (rule ccontr)
  33.466 -  assume "~ a < b"
  33.467 -  hence "b \<le> a" by (simp add: linorder_not_less)
  33.468 -  hence "c+b \<le> c+a" by (rule add_left_mono)
  33.469 -  with this and less show False 
  33.470 -    by (simp add: linorder_not_less [symmetric])
  33.471 -qed
  33.472 +instance pordered_ring \<subseteq> pordered_ab_group_add ..
  33.473  
  33.474 -lemma add_less_imp_less_right:
  33.475 -      "a + c < b + c ==> a < (b::'a::almost_ordered_semiring)"
  33.476 -apply (rule add_less_imp_less_left [of c])
  33.477 -apply (simp add: add_commute)  
  33.478 -done
  33.479 +instance pordered_ring \<subseteq> pordered_cancel_semiring ..
  33.480  
  33.481 -lemma add_less_cancel_left [simp]:
  33.482 -    "(c+a < c+b) = (a < (b::'a::almost_ordered_semiring))"
  33.483 -by (blast intro: add_less_imp_less_left add_strict_left_mono) 
  33.484 +axclass lordered_ring \<subseteq> pordered_ring, lordered_ab_group_abs
  33.485  
  33.486 -lemma add_less_cancel_right [simp]:
  33.487 -    "(a+c < b+c) = (a < (b::'a::almost_ordered_semiring))"
  33.488 -by (blast intro: add_less_imp_less_right add_strict_right_mono)
  33.489 +axclass axclass_abs_if \<subseteq> minus, ord, zero
  33.490 +  abs_if: "abs a = (if (a < 0) then (-a) else a)"
  33.491  
  33.492 -lemma add_le_cancel_left [simp]:
  33.493 -    "(c+a \<le> c+b) = (a \<le> (b::'a::almost_ordered_semiring))"
  33.494 -by (simp add: linorder_not_less [symmetric]) 
  33.495 +axclass ordered_ring_strict \<subseteq> ring, ordered_semiring_strict, axclass_abs_if
  33.496  
  33.497 -lemma add_le_cancel_right [simp]:
  33.498 -    "(a+c \<le> b+c) = (a \<le> (b::'a::almost_ordered_semiring))"
  33.499 -by (simp add: linorder_not_less [symmetric]) 
  33.500 +instance ordered_ring_strict \<subseteq> lordered_ab_group ..
  33.501  
  33.502 -lemma add_le_imp_le_left:
  33.503 -      "c + a \<le> c + b ==> a \<le> (b::'a::almost_ordered_semiring)"
  33.504 -by simp
  33.505 +instance ordered_ring_strict \<subseteq> lordered_ring
  33.506 +by (intro_classes, simp add: abs_if join_eq_if)
  33.507  
  33.508 -lemma add_le_imp_le_right:
  33.509 -      "a + c \<le> b + c ==> a \<le> (b::'a::almost_ordered_semiring)"
  33.510 -by simp
  33.511 +axclass pordered_comm_ring \<subseteq> comm_ring, pordered_comm_semiring
  33.512  
  33.513 -lemma add_increasing: "[|0\<le>a; b\<le>c|] ==> b \<le> a + (c::'a::almost_ordered_semiring)"
  33.514 -by (insert add_mono [of 0 a b c], simp)
  33.515 +axclass ordered_semidom \<subseteq> comm_semiring_1_cancel, ordered_comm_semiring_strict (* previously ordered_semiring *)
  33.516 +  zero_less_one [simp]: "0 < 1"
  33.517  
  33.518 +axclass ordered_idom \<subseteq> comm_ring_1, ordered_comm_semiring_strict, axclass_abs_if (* previously ordered_ring *)
  33.519  
  33.520 -subsection {* Ordering Rules for Unary Minus *}
  33.521 +instance ordered_idom \<subseteq> ordered_ring_strict ..
  33.522  
  33.523 -lemma le_imp_neg_le:
  33.524 -      assumes "a \<le> (b::'a::ordered_ring)" shows "-b \<le> -a"
  33.525 -proof -
  33.526 -  have "-a+a \<le> -a+b"
  33.527 -    by (rule add_left_mono) 
  33.528 -  hence "0 \<le> -a+b"
  33.529 -    by simp
  33.530 -  hence "0 + (-b) \<le> (-a + b) + (-b)"
  33.531 -    by (rule add_right_mono) 
  33.532 -  thus ?thesis
  33.533 -    by (simp add: add_assoc)
  33.534 -qed
  33.535 -
  33.536 -lemma neg_le_iff_le [simp]: "(-b \<le> -a) = (a \<le> (b::'a::ordered_ring))"
  33.537 -proof 
  33.538 -  assume "- b \<le> - a"
  33.539 -  hence "- (- a) \<le> - (- b)"
  33.540 -    by (rule le_imp_neg_le)
  33.541 -  thus "a\<le>b" by simp
  33.542 -next
  33.543 -  assume "a\<le>b"
  33.544 -  thus "-b \<le> -a" by (rule le_imp_neg_le)
  33.545 -qed
  33.546 -
  33.547 -lemma neg_le_0_iff_le [simp]: "(-a \<le> 0) = (0 \<le> (a::'a::ordered_ring))"
  33.548 -by (subst neg_le_iff_le [symmetric], simp)
  33.549 -
  33.550 -lemma neg_0_le_iff_le [simp]: "(0 \<le> -a) = (a \<le> (0::'a::ordered_ring))"
  33.551 -by (subst neg_le_iff_le [symmetric], simp)
  33.552 -
  33.553 -lemma neg_less_iff_less [simp]: "(-b < -a) = (a < (b::'a::ordered_ring))"
  33.554 -by (force simp add: order_less_le) 
  33.555 -
  33.556 -lemma neg_less_0_iff_less [simp]: "(-a < 0) = (0 < (a::'a::ordered_ring))"
  33.557 -by (subst neg_less_iff_less [symmetric], simp)
  33.558 -
  33.559 -lemma neg_0_less_iff_less [simp]: "(0 < -a) = (a < (0::'a::ordered_ring))"
  33.560 -by (subst neg_less_iff_less [symmetric], simp)
  33.561 -
  33.562 -text{*The next several equations can make the simplifier loop!*}
  33.563 -
  33.564 -lemma less_minus_iff: "(a < - b) = (b < - (a::'a::ordered_ring))"
  33.565 -proof -
  33.566 -  have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)
  33.567 -  thus ?thesis by simp
  33.568 -qed
  33.569 -
  33.570 -lemma minus_less_iff: "(- a < b) = (- b < (a::'a::ordered_ring))"
  33.571 -proof -
  33.572 -  have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)
  33.573 -  thus ?thesis by simp
  33.574 -qed
  33.575 -
  33.576 -lemma le_minus_iff: "(a \<le> - b) = (b \<le> - (a::'a::ordered_ring))"
  33.577 -apply (simp add: linorder_not_less [symmetric])
  33.578 -apply (rule minus_less_iff) 
  33.579 -done
  33.580 -
  33.581 -lemma minus_le_iff: "(- a \<le> b) = (- b \<le> (a::'a::ordered_ring))"
  33.582 -apply (simp add: linorder_not_less [symmetric])
  33.583 -apply (rule less_minus_iff) 
  33.584 -done
  33.585 -
  33.586 -
  33.587 -subsection{*Subtraction Laws*}
  33.588 -
  33.589 -lemma add_diff_eq: "a + (b - c) = (a + b) - (c::'a::abelian_group)"
  33.590 -by (simp add: diff_minus plus_ac0)
  33.591 -
  33.592 -lemma diff_add_eq: "(a - b) + c = (a + c) - (b::'a::abelian_group)"
  33.593 -by (simp add: diff_minus plus_ac0)
  33.594 -
  33.595 -lemma diff_eq_eq: "(a-b = c) = (a = c + (b::'a::abelian_group))"
  33.596 -by (auto simp add: diff_minus add_assoc)
  33.597 -
  33.598 -lemma eq_diff_eq: "(a = c-b) = (a + (b::'a::abelian_group) = c)"
  33.599 -by (auto simp add: diff_minus add_assoc)
  33.600 -
  33.601 -lemma diff_diff_eq: "(a - b) - c = a - (b + (c::'a::abelian_group))"
  33.602 -by (simp add: diff_minus plus_ac0)
  33.603 -
  33.604 -lemma diff_diff_eq2: "a - (b - c) = (a + c) - (b::'a::abelian_group)"
  33.605 -by (simp add: diff_minus plus_ac0)
  33.606 -
  33.607 -text{*Further subtraction laws for ordered rings*}
  33.608 -
  33.609 -lemma less_iff_diff_less_0: "(a < b) = (a - b < (0::'a::ordered_ring))"
  33.610 -proof -
  33.611 -  have  "(a < b) = (a + (- b) < b + (-b))"  
  33.612 -    by (simp only: add_less_cancel_right)
  33.613 -  also have "... =  (a - b < 0)" by (simp add: diff_minus)
  33.614 -  finally show ?thesis .
  33.615 -qed
  33.616 -
  33.617 -lemma diff_less_eq: "(a-b < c) = (a < c + (b::'a::ordered_ring))"
  33.618 -apply (subst less_iff_diff_less_0)
  33.619 -apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
  33.620 -apply (simp add: diff_minus add_ac)
  33.621 -done
  33.622 -
  33.623 -lemma less_diff_eq: "(a < c-b) = (a + (b::'a::ordered_ring) < c)"
  33.624 -apply (subst less_iff_diff_less_0)
  33.625 -apply (rule less_iff_diff_less_0 [of _ "c-b", THEN ssubst])
  33.626 -apply (simp add: diff_minus add_ac)
  33.627 -done
  33.628 -
  33.629 -lemma diff_le_eq: "(a-b \<le> c) = (a \<le> c + (b::'a::ordered_ring))"
  33.630 -by (simp add: linorder_not_less [symmetric] less_diff_eq)
  33.631 -
  33.632 -lemma le_diff_eq: "(a \<le> c-b) = (a + (b::'a::ordered_ring) \<le> c)"
  33.633 -by (simp add: linorder_not_less [symmetric] diff_less_eq)
  33.634 -
  33.635 -text{*This list of rewrites simplifies (in)equalities by bringing subtractions
  33.636 -  to the top and then moving negative terms to the other side.
  33.637 -  Use with @{text add_ac}*}
  33.638 -lemmas compare_rls =
  33.639 -       diff_minus [symmetric]
  33.640 -       add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
  33.641 -       diff_less_eq less_diff_eq diff_le_eq le_diff_eq
  33.642 -       diff_eq_eq eq_diff_eq
  33.643 -
  33.644 -text{*This list of rewrites decides ring equalities by ordered rewriting.*}
  33.645 -lemmas ring_eq_simps =
  33.646 -  times_ac1.assoc times_ac1.commute times_ac1_left_commute
  33.647 -  left_distrib right_distrib left_diff_distrib right_diff_distrib
  33.648 -  plus_ac0.assoc plus_ac0.commute plus_ac0_left_commute
  33.649 -  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
  33.650 -  diff_eq_eq eq_diff_eq
  33.651 -
  33.652 -subsection{*Lemmas for the @{text cancel_numerals} simproc*}
  33.653 -
  33.654 -lemma eq_iff_diff_eq_0: "(a = b) = (a-b = (0::'a::abelian_group))"
  33.655 -by (simp add: compare_rls)
  33.656 -
  33.657 -lemma le_iff_diff_le_0: "(a \<le> b) = (a-b \<le> (0::'a::ordered_ring))"
  33.658 -by (simp add: compare_rls)
  33.659 +axclass ordered_field \<subseteq> field, ordered_idom
  33.660  
  33.661  lemma eq_add_iff1:
  33.662       "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))"
  33.663 +apply (simp add: diff_minus left_distrib)
  33.664  apply (simp add: diff_minus left_distrib add_ac)
  33.665 -apply (simp add: compare_rls minus_mult_left [symmetric]) 
  33.666 +apply (simp add: compare_rls minus_mult_left [symmetric])
  33.667  done
  33.668  
  33.669  lemma eq_add_iff2:
  33.670 @@ -543,167 +238,199 @@
  33.671  done
  33.672  
  33.673  lemma less_add_iff1:
  33.674 -     "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::ordered_ring))"
  33.675 +     "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::pordered_ring))"
  33.676  apply (simp add: diff_minus left_distrib add_ac)
  33.677  apply (simp add: compare_rls minus_mult_left [symmetric]) 
  33.678  done
  33.679  
  33.680  lemma less_add_iff2:
  33.681 -     "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::ordered_ring))"
  33.682 +     "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::pordered_ring))"
  33.683  apply (simp add: diff_minus left_distrib add_ac)
  33.684  apply (simp add: compare_rls minus_mult_left [symmetric]) 
  33.685  done
  33.686  
  33.687  lemma le_add_iff1:
  33.688 -     "(a*e + c \<le> b*e + d) = ((a-b)*e + c \<le> (d::'a::ordered_ring))"
  33.689 +     "(a*e + c \<le> b*e + d) = ((a-b)*e + c \<le> (d::'a::pordered_ring))"
  33.690  apply (simp add: diff_minus left_distrib add_ac)
  33.691  apply (simp add: compare_rls minus_mult_left [symmetric]) 
  33.692  done
  33.693  
  33.694  lemma le_add_iff2:
  33.695 -     "(a*e + c \<le> b*e + d) = (c \<le> (b-a)*e + (d::'a::ordered_ring))"
  33.696 +     "(a*e + c \<le> b*e + d) = (c \<le> (b-a)*e + (d::'a::pordered_ring))"
  33.697  apply (simp add: diff_minus left_distrib add_ac)
  33.698  apply (simp add: compare_rls minus_mult_left [symmetric]) 
  33.699  done
  33.700  
  33.701 -
  33.702  subsection {* Ordering Rules for Multiplication *}
  33.703  
  33.704 -lemma mult_strict_right_mono:
  33.705 -     "[|a < b; 0 < c|] ==> a * c < b * (c::'a::almost_ordered_semiring)"
  33.706 -by (simp add: mult_commute [of _ c] mult_strict_left_mono)
  33.707 -
  33.708 -lemma mult_left_mono:
  33.709 -     "[|a \<le> b; 0 \<le> c|] ==> c * a \<le> c * (b::'a::almost_ordered_semiring)"
  33.710 -  apply (case_tac "c=0", simp)
  33.711 -  apply (force simp add: mult_strict_left_mono order_le_less) 
  33.712 -  done
  33.713 -
  33.714 -lemma mult_right_mono:
  33.715 -     "[|a \<le> b; 0 \<le> c|] ==> a*c \<le> b * (c::'a::almost_ordered_semiring)"
  33.716 -  by (simp add: mult_left_mono mult_commute [of _ c]) 
  33.717 -
  33.718  lemma mult_left_le_imp_le:
  33.719 -     "[|c*a \<le> c*b; 0 < c|] ==> a \<le> (b::'a::almost_ordered_semiring)"
  33.720 +     "[|c*a \<le> c*b; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
  33.721    by (force simp add: mult_strict_left_mono linorder_not_less [symmetric])
  33.722   
  33.723  lemma mult_right_le_imp_le:
  33.724 -     "[|a*c \<le> b*c; 0 < c|] ==> a \<le> (b::'a::almost_ordered_semiring)"
  33.725 +     "[|a*c \<le> b*c; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
  33.726    by (force simp add: mult_strict_right_mono linorder_not_less [symmetric])
  33.727  
  33.728  lemma mult_left_less_imp_less:
  33.729 -     "[|c*a < c*b; 0 \<le> c|] ==> a < (b::'a::almost_ordered_semiring)"
  33.730 +     "[|c*a < c*b; 0 \<le> c|] ==> a < (b::'a::ordered_semiring_strict)"
  33.731    by (force simp add: mult_left_mono linorder_not_le [symmetric])
  33.732   
  33.733  lemma mult_right_less_imp_less:
  33.734 -     "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::almost_ordered_semiring)"
  33.735 +     "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::ordered_semiring_strict)"
  33.736    by (force simp add: mult_right_mono linorder_not_le [symmetric])
  33.737  
  33.738  lemma mult_strict_left_mono_neg:
  33.739 -     "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring)"
  33.740 +     "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring_strict)"
  33.741  apply (drule mult_strict_left_mono [of _ _ "-c"])
  33.742  apply (simp_all add: minus_mult_left [symmetric]) 
  33.743  done
  33.744  
  33.745 +lemma mult_left_mono_neg:
  33.746 +     "[|b \<le> a; c \<le> 0|] ==> c * a \<le>  c * (b::'a::pordered_ring)"
  33.747 +apply (drule mult_left_mono [of _ _ "-c"])
  33.748 +apply (simp_all add: minus_mult_left [symmetric]) 
  33.749 +done
  33.750 +
  33.751  lemma mult_strict_right_mono_neg:
  33.752 -     "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring)"
  33.753 +     "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring_strict)"
  33.754  apply (drule mult_strict_right_mono [of _ _ "-c"])
  33.755  apply (simp_all add: minus_mult_right [symmetric]) 
  33.756  done
  33.757  
  33.758 +lemma mult_right_mono_neg:
  33.759 +     "[|b \<le> a; c \<le> 0|] ==> a * c \<le>  (b::'a::pordered_ring) * c"
  33.760 +apply (drule mult_right_mono [of _ _ "-c"])
  33.761 +apply (simp)
  33.762 +apply (simp_all add: minus_mult_right [symmetric]) 
  33.763 +done
  33.764  
  33.765  subsection{* Products of Signs *}
  33.766  
  33.767 -lemma mult_pos: "[| (0::'a::almost_ordered_semiring) < a; 0 < b |] ==> 0 < a*b"
  33.768 +lemma mult_pos: "[| (0::'a::ordered_semiring_strict) < a; 0 < b |] ==> 0 < a*b"
  33.769  by (drule mult_strict_left_mono [of 0 b], auto)
  33.770  
  33.771 -lemma mult_pos_neg: "[| (0::'a::almost_ordered_semiring) < a; b < 0 |] ==> a*b < 0"
  33.772 +lemma mult_pos_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; 0 \<le> b |] ==> 0 \<le> a*b"
  33.773 +by (drule mult_left_mono [of 0 b], auto)
  33.774 +
  33.775 +lemma mult_pos_neg: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> a*b < 0"
  33.776  by (drule mult_strict_left_mono [of b 0], auto)
  33.777  
  33.778 -lemma mult_neg: "[| a < (0::'a::ordered_ring); b < 0 |] ==> 0 < a*b"
  33.779 +lemma mult_pos_neg_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> a*b \<le> 0"
  33.780 +by (drule mult_left_mono [of b 0], auto)
  33.781 +
  33.782 +lemma mult_pos_neg2: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> b*a < 0" 
  33.783 +by (drule mult_strict_right_mono[of b 0], auto)
  33.784 +
  33.785 +lemma mult_pos_neg2_le: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> b*a \<le> 0" 
  33.786 +by (drule mult_right_mono[of b 0], auto)
  33.787 +
  33.788 +lemma mult_neg: "[| a < (0::'a::ordered_ring_strict); b < 0 |] ==> 0 < a*b"
  33.789  by (drule mult_strict_right_mono_neg, auto)
  33.790  
  33.791 +lemma mult_neg_le: "[| a \<le> (0::'a::pordered_ring); b \<le> 0 |] ==> 0 \<le> a*b"
  33.792 +by (drule mult_right_mono_neg[of a 0 b ], auto)
  33.793 +
  33.794  lemma zero_less_mult_pos:
  33.795 -     "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::almost_ordered_semiring)"
  33.796 +     "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
  33.797  apply (case_tac "b\<le>0") 
  33.798   apply (auto simp add: order_le_less linorder_not_less)
  33.799  apply (drule_tac mult_pos_neg [of a b]) 
  33.800   apply (auto dest: order_less_not_sym)
  33.801  done
  33.802  
  33.803 +lemma zero_less_mult_pos2:
  33.804 +     "[| 0 < b*a; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
  33.805 +apply (case_tac "b\<le>0") 
  33.806 + apply (auto simp add: order_le_less linorder_not_less)
  33.807 +apply (drule_tac mult_pos_neg2 [of a b]) 
  33.808 + apply (auto dest: order_less_not_sym)
  33.809 +done
  33.810 +
  33.811  lemma zero_less_mult_iff:
  33.812 -     "((0::'a::ordered_ring) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
  33.813 +     "((0::'a::ordered_ring_strict) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
  33.814  apply (auto simp add: order_le_less linorder_not_less mult_pos mult_neg)
  33.815  apply (blast dest: zero_less_mult_pos) 
  33.816 -apply (simp add: mult_commute [of a b]) 
  33.817 -apply (blast dest: zero_less_mult_pos) 
  33.818 +apply (blast dest: zero_less_mult_pos2)
  33.819  done
  33.820  
  33.821  text{*A field has no "zero divisors", and this theorem holds without the
  33.822        assumption of an ordering.  See @{text field_mult_eq_0_iff} below.*}
  33.823 -lemma mult_eq_0_iff [simp]: "(a*b = (0::'a::ordered_ring)) = (a = 0 | b = 0)"
  33.824 +lemma mult_eq_0_iff [simp]: "(a*b = (0::'a::ordered_ring_strict)) = (a = 0 | b = 0)"
  33.825  apply (case_tac "a < 0")
  33.826  apply (auto simp add: linorder_not_less order_le_less linorder_neq_iff)
  33.827  apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono)+
  33.828  done
  33.829  
  33.830  lemma zero_le_mult_iff:
  33.831 -     "((0::'a::ordered_ring) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
  33.832 +     "((0::'a::ordered_ring_strict) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
  33.833  by (auto simp add: eq_commute [of 0] order_le_less linorder_not_less
  33.834                     zero_less_mult_iff)
  33.835  
  33.836  lemma mult_less_0_iff:
  33.837 -     "(a*b < (0::'a::ordered_ring)) = (0 < a & b < 0 | a < 0 & 0 < b)"
  33.838 +     "(a*b < (0::'a::ordered_ring_strict)) = (0 < a & b < 0 | a < 0 & 0 < b)"
  33.839  apply (insert zero_less_mult_iff [of "-a" b]) 
  33.840  apply (force simp add: minus_mult_left[symmetric]) 
  33.841  done
  33.842  
  33.843  lemma mult_le_0_iff:
  33.844 -     "(a*b \<le> (0::'a::ordered_ring)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
  33.845 +     "(a*b \<le> (0::'a::ordered_ring_strict)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
  33.846  apply (insert zero_le_mult_iff [of "-a" b]) 
  33.847  apply (force simp add: minus_mult_left[symmetric]) 
  33.848  done
  33.849  
  33.850 -lemma zero_le_square: "(0::'a::ordered_ring) \<le> a*a"
  33.851 +lemma split_mult_pos_le: "(0 \<le> a & 0 \<le> b) | (a \<le> 0 & b \<le> 0) \<Longrightarrow> 0 \<le> a * (b::_::pordered_ring)"
  33.852 +by (auto simp add: mult_pos_le mult_neg_le)
  33.853 +
  33.854 +lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> (0::_::pordered_cancel_semiring)" 
  33.855 +by (auto simp add: mult_pos_neg_le mult_pos_neg2_le)
  33.856 +
  33.857 +lemma zero_le_square: "(0::'a::ordered_ring_strict) \<le> a*a"
  33.858  by (simp add: zero_le_mult_iff linorder_linear) 
  33.859  
  33.860 -text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semiring}
  33.861 -      theorems available to members of @{term ordered_ring} *}
  33.862 -instance ordered_ring \<subseteq> ordered_semiring
  33.863 +text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semidom}
  33.864 +      theorems available to members of @{term ordered_idom} *}
  33.865 +
  33.866 +instance ordered_idom \<subseteq> ordered_semidom
  33.867  proof
  33.868    have "(0::'a) \<le> 1*1" by (rule zero_le_square)
  33.869    thus "(0::'a) < 1" by (simp add: order_le_less) 
  33.870  qed
  33.871  
  33.872 +instance ordered_ring_strict \<subseteq> axclass_no_zero_divisors 
  33.873 +by (intro_classes, simp)
  33.874 +
  33.875 +instance ordered_idom \<subseteq> idom ..
  33.876 +
  33.877  text{*All three types of comparision involving 0 and 1 are covered.*}
  33.878  
  33.879  declare zero_neq_one [THEN not_sym, simp]
  33.880  
  33.881 -lemma zero_le_one [simp]: "(0::'a::ordered_semiring) \<le> 1"
  33.882 +lemma zero_le_one [simp]: "(0::'a::ordered_semidom) \<le> 1"
  33.883    by (rule zero_less_one [THEN order_less_imp_le]) 
  33.884  
  33.885 -lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semiring) \<le> 0"
  33.886 -by (simp add: linorder_not_le zero_less_one) 
  33.887 +lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semidom) \<le> 0"
  33.888 +by (simp add: linorder_not_le) 
  33.889  
  33.890 -lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semiring) < 0"
  33.891 -by (simp add: linorder_not_less zero_le_one) 
  33.892 -
  33.893 +lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semidom) < 0"
  33.894 +by (simp add: linorder_not_less) 
  33.895  
  33.896  subsection{*More Monotonicity*}
  33.897  
  33.898  lemma mult_left_mono_neg:
  33.899 -     "[|b \<le> a; c \<le> 0|] ==> c * a \<le> c * (b::'a::ordered_ring)"
  33.900 +     "[|b \<le> a; c \<le> 0|] ==> c * a \<le> c * (b::'a::pordered_ring)"
  33.901  apply (drule mult_left_mono [of _ _ "-c"]) 
  33.902  apply (simp_all add: minus_mult_left [symmetric]) 
  33.903  done
  33.904  
  33.905  lemma mult_right_mono_neg:
  33.906 -     "[|b \<le> a; c \<le> 0|] ==> a * c \<le> b * (c::'a::ordered_ring)"
  33.907 -  by (simp add: mult_left_mono_neg mult_commute [of _ c]) 
  33.908 +     "[|b \<le> a; c \<le> 0|] ==> a * c \<le> b * (c::'a::pordered_ring)"
  33.909 +apply (drule mult_right_mono [of _ _ "-c"]) 
  33.910 +apply (simp_all add: minus_mult_right [symmetric]) 
  33.911 +done  
  33.912  
  33.913  text{*Strict monotonicity in both arguments*}
  33.914  lemma mult_strict_mono:
  33.915 -     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_semiring)"
  33.916 +     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
  33.917  apply (case_tac "c=0")
  33.918   apply (simp add: mult_pos) 
  33.919  apply (erule mult_strict_right_mono [THEN order_less_trans])
  33.920 @@ -713,31 +440,30 @@
  33.921  
  33.922  text{*This weaker variant has more natural premises*}
  33.923  lemma mult_strict_mono':
  33.924 -     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_semiring)"
  33.925 +     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
  33.926  apply (rule mult_strict_mono)
  33.927  apply (blast intro: order_le_less_trans)+
  33.928  done
  33.929  
  33.930  lemma mult_mono:
  33.931       "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] 
  33.932 -      ==> a * c  \<le>  b * (d::'a::ordered_semiring)"
  33.933 +      ==> a * c  \<le>  b * (d::'a::pordered_semiring)"
  33.934  apply (erule mult_right_mono [THEN order_trans], assumption)
  33.935  apply (erule mult_left_mono, assumption)
  33.936  done
  33.937  
  33.938 -lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semiring)"
  33.939 +lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semidom)"
  33.940  apply (insert mult_strict_mono [of 1 m 1 n]) 
  33.941  apply (simp add:  order_less_trans [OF zero_less_one]) 
  33.942  done
  33.943  
  33.944 -
  33.945  subsection{*Cancellation Laws for Relationships With a Common Factor*}
  33.946  
  33.947  text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
  33.948     also with the relations @{text "\<le>"} and equality.*}
  33.949  
  33.950  lemma mult_less_cancel_right:
  33.951 -    "(a*c < b*c) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring)))"
  33.952 +    "(a*c < b*c) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
  33.953  apply (case_tac "c = 0")
  33.954  apply (auto simp add: linorder_neq_iff mult_strict_right_mono 
  33.955                        mult_strict_right_mono_neg)
  33.956 @@ -750,20 +476,29 @@
  33.957  done
  33.958  
  33.959  lemma mult_less_cancel_left:
  33.960 -    "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring)))"
  33.961 -by (simp add: mult_commute [of c] mult_less_cancel_right)
  33.962 +    "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
  33.963 +apply (case_tac "c = 0")
  33.964 +apply (auto simp add: linorder_neq_iff mult_strict_left_mono 
  33.965 +                      mult_strict_left_mono_neg)
  33.966 +apply (auto simp add: linorder_not_less 
  33.967 +                      linorder_not_le [symmetric, of "c*a"]
  33.968 +                      linorder_not_le [symmetric, of a])
  33.969 +apply (erule_tac [!] notE)
  33.970 +apply (auto simp add: order_less_imp_le mult_left_mono 
  33.971 +                      mult_left_mono_neg)
  33.972 +done
  33.973  
  33.974  lemma mult_le_cancel_right:
  33.975 -     "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring)))"
  33.976 +     "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
  33.977  by (simp add: linorder_not_less [symmetric] mult_less_cancel_right)
  33.978  
  33.979  lemma mult_le_cancel_left:
  33.980 -     "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring)))"
  33.981 -by (simp add: mult_commute [of c] mult_le_cancel_right)
  33.982 +     "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
  33.983 +by (simp add: linorder_not_less [symmetric] mult_less_cancel_left)
  33.984  
  33.985  lemma mult_less_imp_less_left:
  33.986        assumes less: "c*a < c*b" and nonneg: "0 \<le> c"
  33.987 -      shows "a < (b::'a::ordered_semiring)"
  33.988 +      shows "a < (b::'a::ordered_semiring_strict)"
  33.989  proof (rule ccontr)
  33.990    assume "~ a < b"
  33.991    hence "b \<le> a" by (simp add: linorder_not_less)
  33.992 @@ -773,12 +508,19 @@
  33.993  qed
  33.994  
  33.995  lemma mult_less_imp_less_right:
  33.996 -    "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::ordered_semiring)"
  33.997 -  by (rule mult_less_imp_less_left, simp add: mult_commute)
  33.998 +  assumes less: "a*c < b*c" and nonneg: "0 <= c"
  33.999 +  shows "a < (b::'a::ordered_semiring_strict)"
 33.1000 +proof (rule ccontr)
 33.1001 +  assume "~ a < b"
 33.1002 +  hence "b \<le> a" by (simp add: linorder_not_less)
 33.1003 +  hence "b*c \<le> a*c" by (rule mult_right_mono)
 33.1004 +  with this and less show False 
 33.1005 +    by (simp add: linorder_not_less [symmetric])
 33.1006 +qed  
 33.1007  
 33.1008  text{*Cancellation of equalities with a common factor*}
 33.1009  lemma mult_cancel_right [simp]:
 33.1010 -     "(a*c = b*c) = (c = (0::'a::ordered_ring) | a=b)"
 33.1011 +     "(a*c = b*c) = (c = (0::'a::ordered_ring_strict) | a=b)"
 33.1012  apply (cut_tac linorder_less_linear [of 0 c])
 33.1013  apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono
 33.1014               simp add: linorder_neq_iff)
 33.1015 @@ -787,10 +529,21 @@
 33.1016  text{*These cancellation theorems require an ordering. Versions are proved
 33.1017        below that work for fields without an ordering.*}
 33.1018  lemma mult_cancel_left [simp]:
 33.1019 -     "(c*a = c*b) = (c = (0::'a::ordered_ring) | a=b)"
 33.1020 -by (simp add: mult_commute [of c] mult_cancel_right)
 33.1021 +     "(c*a = c*b) = (c = (0::'a::ordered_ring_strict) | a=b)"
 33.1022 +apply (cut_tac linorder_less_linear [of 0 c])
 33.1023 +apply (force dest: mult_strict_left_mono_neg mult_strict_left_mono
 33.1024 +             simp add: linorder_neq_iff)
 33.1025 +done
 33.1026  
 33.1027 -
 33.1028 +text{*This list of rewrites decides ring equalities by ordered rewriting.*}
 33.1029 +lemmas ring_eq_simps =
 33.1030 +  mult_ac
 33.1031 +  left_distrib right_distrib left_diff_distrib right_diff_distrib
 33.1032 +  add_ac
 33.1033 +  add_diff_eq diff_add_eq diff_diff_eq diff_diff_eq2
 33.1034 +  diff_eq_eq eq_diff_eq
 33.1035 +    
 33.1036 +thm ring_eq_simps
 33.1037  subsection {* Fields *}
 33.1038  
 33.1039  lemma right_inverse [simp]:
 33.1040 @@ -1571,14 +1324,14 @@
 33.1041  
 33.1042  subsection {* Ordered Fields are Dense *}
 33.1043  
 33.1044 -lemma less_add_one: "a < (a+1::'a::ordered_semiring)"
 33.1045 +lemma less_add_one: "a < (a+1::'a::ordered_semidom)"
 33.1046  proof -
 33.1047 -  have "a+0 < (a+1::'a::ordered_semiring)"
 33.1048 +  have "a+0 < (a+1::'a::ordered_semidom)"
 33.1049      by (blast intro: zero_less_one add_strict_left_mono) 
 33.1050    thus ?thesis by simp
 33.1051  qed
 33.1052  
 33.1053 -lemma zero_less_two: "0 < (1+1::'a::ordered_semiring)"
 33.1054 +lemma zero_less_two: "0 < (1+1::'a::ordered_semidom)"
 33.1055    by (blast intro: order_less_trans zero_less_one less_add_one) 
 33.1056  
 33.1057  lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)"
 33.1058 @@ -1590,62 +1343,102 @@
 33.1059  lemma dense: "a < b ==> \<exists>r::'a::ordered_field. a < r & r < b"
 33.1060  by (blast intro!: less_half_sum gt_half_sum)
 33.1061  
 33.1062 -
 33.1063  subsection {* Absolute Value *}
 33.1064  
 33.1065 -lemma abs_zero [simp]: "abs 0 = (0::'a::ordered_ring)"
 33.1066 -by (simp add: abs_if)
 33.1067 -
 33.1068 -lemma abs_one [simp]: "abs 1 = (1::'a::ordered_ring)"
 33.1069 +lemma abs_one [simp]: "abs 1 = (1::'a::ordered_idom)"
 33.1070    by (simp add: abs_if zero_less_one [THEN order_less_not_sym]) 
 33.1071  
 33.1072 -lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_ring)" 
 33.1073 -apply (case_tac "a=0 | b=0", force) 
 33.1074 -apply (auto elim: order_less_asym
 33.1075 -            simp add: abs_if mult_less_0_iff linorder_neq_iff
 33.1076 -                  minus_mult_left [symmetric] minus_mult_right [symmetric])  
 33.1077 -done
 33.1078 +lemma abs_le_mult: "abs (a * b) \<le> (abs a) * (abs (b::'a::lordered_ring))" 
 33.1079 +proof -
 33.1080 +  let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b"
 33.1081 +  let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
 33.1082 +  have a: "(abs a) * (abs b) = ?x"
 33.1083 +    by (simp only: abs_prts[of a] abs_prts[of b] ring_eq_simps)
 33.1084 +  {
 33.1085 +    fix u v :: 'a
 33.1086 +    have bh: "\<lbrakk>u = a; v = b\<rbrakk> \<Longrightarrow> u * v = ?y"
 33.1087 +      apply (subst prts[of u], subst prts[of v])
 33.1088 +      apply (simp add: left_distrib right_distrib add_ac) 
 33.1089 +      done
 33.1090 +  }
 33.1091 +  note b = this[OF refl[of a] refl[of b]]
 33.1092 +  note addm = add_mono[of "0::'a" _ "0::'a", simplified]
 33.1093 +  note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified]
 33.1094 +  have xy: "- ?x <= ?y"
 33.1095 +    apply (simp add: compare_rls)
 33.1096 +    apply (rule add_le_imp_le_left[of "-(pprt a * nprt b + nprt a * pprt b)"])
 33.1097 +    apply (simp add: add_ac)
 33.1098 +    proof -
 33.1099 +      let ?r = "nprt a * nprt b +(nprt a * nprt b + (nprt a * pprt b + (pprt a * nprt b + (pprt a * pprt b + (pprt a * pprt b +
 33.1100 +	(- (nprt a * pprt b) + - (pprt a * nprt b)))))))"
 33.1101 +      let ?rr = "nprt a * nprt b + nprt a * nprt b + ((nprt a * pprt b) + (- (nprt a * pprt b))) + ((pprt a * nprt b) + - (pprt a * nprt b))
 33.1102 +	+ pprt a * pprt b + pprt a * pprt b"
 33.1103 +      have a:"?r = ?rr" by (simp only: add_ac)      
 33.1104 +      have "0 <= ?rr"
 33.1105 +	apply (simp)
 33.1106 +	apply (rule addm)+
 33.1107 +	apply (simp_all add: mult_neg_le mult_pos_le)
 33.1108 +	done
 33.1109 +      with a show "0 <= ?r" by simp
 33.1110 +    qed
 33.1111 +  have yx: "?y <= ?x"
 33.1112 +    apply (simp add: add_ac)
 33.1113 +    apply (simp add: compare_rls)
 33.1114 +    apply (rule add_le_imp_le_right[of _ "-(pprt a * pprt b)"])
 33.1115 +    apply (simp add: add_ac)
 33.1116 +    apply (rule addm2, (simp add: mult_pos_neg_le mult_pos_neg2_le)+)+
 33.1117 +    done
 33.1118 +  have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
 33.1119 +  have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
 33.1120 +  show ?thesis
 33.1121 +    apply (rule abs_leI)
 33.1122 +    apply (simp add: i1)
 33.1123 +    apply (simp add: i2[simplified minus_le_iff])
 33.1124 +    done
 33.1125 +qed
 33.1126  
 33.1127 -lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_ring)"
 33.1128 +lemma abs_eq_mult: 
 33.1129 +  assumes "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)"
 33.1130 +  shows "abs (a*b) = abs a * abs (b::'a::lordered_ring)"
 33.1131 +proof -
 33.1132 +  have s: "(0 <= a*b) | (a*b <= 0)"
 33.1133 +    apply (auto)    
 33.1134 +    apply (rule_tac split_mult_pos_le)
 33.1135 +    apply (rule_tac contrapos_np[of "a*b <= 0"])
 33.1136 +    apply (simp)
 33.1137 +    apply (rule_tac split_mult_neg_le)
 33.1138 +    apply (insert prems)
 33.1139 +    apply (blast)
 33.1140 +    done
 33.1141 +  have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
 33.1142 +    by (simp add: prts[symmetric])
 33.1143 +  show ?thesis
 33.1144 +  proof cases
 33.1145 +    assume "0 <= a * b"
 33.1146 +    then show ?thesis
 33.1147 +      apply (simp_all add: mulprts abs_prts)
 33.1148 +      apply (insert prems)
 33.1149 +      apply (auto simp add: ring_eq_simps iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_zero_pprt]
 33.1150 +	iff2imp[OF le_zero_iff_pprt_id] iff2imp[OF zero_le_iff_nprt_id] order_antisym mult_pos_neg_le[of a b] mult_pos_neg2_le[of b a])
 33.1151 +      done
 33.1152 +  next
 33.1153 +    assume "~(0 <= a*b)"
 33.1154 +    with s have "a*b <= 0" by simp
 33.1155 +    then show ?thesis
 33.1156 +      apply (simp_all add: mulprts abs_prts)
 33.1157 +      apply (insert prems)
 33.1158 +      apply (auto simp add: ring_eq_simps iff2imp[OF zero_le_iff_zero_nprt] iff2imp[OF le_zero_iff_zero_pprt]
 33.1159 +	iff2imp[OF le_zero_iff_pprt_id] iff2imp[OF zero_le_iff_nprt_id] order_antisym mult_pos_le[of a b] mult_neg_le[of a b])
 33.1160 +      done
 33.1161 +  qed
 33.1162 +qed
 33.1163 +
 33.1164 +lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" 
 33.1165 +by (simp add: abs_eq_mult linorder_linear)
 33.1166 +
 33.1167 +lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)"
 33.1168  by (simp add: abs_if) 
 33.1169  
 33.1170 -lemma abs_eq_0 [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
 33.1171 -by (simp add: abs_if)
 33.1172 -
 33.1173 -lemma zero_less_abs_iff [simp]: "(0 < abs a) = (a \<noteq> (0::'a::ordered_ring))"
 33.1174 -by (simp add: abs_if linorder_neq_iff)
 33.1175 -
 33.1176 -lemma abs_not_less_zero [simp]: "~ abs a < (0::'a::ordered_ring)"
 33.1177 -apply (simp add: abs_if)
 33.1178 -by (simp add: abs_if  order_less_not_sym [of a 0])
 33.1179 -
 33.1180 -lemma abs_le_zero_iff [simp]: "(abs a \<le> (0::'a::ordered_ring)) = (a = 0)" 
 33.1181 -by (simp add: order_le_less) 
 33.1182 -
 33.1183 -lemma abs_minus_cancel [simp]: "abs (-a) = abs(a::'a::ordered_ring)"
 33.1184 -apply (auto simp add: abs_if linorder_not_less order_less_not_sym [of 0 a])  
 33.1185 -apply (drule order_antisym, assumption, simp) 
 33.1186 -done
 33.1187 -
 33.1188 -lemma abs_ge_zero [simp]: "(0::'a::ordered_ring) \<le> abs a"
 33.1189 -apply (simp add: abs_if order_less_imp_le)
 33.1190 -apply (simp add: linorder_not_less) 
 33.1191 -done
 33.1192 -
 33.1193 -lemma abs_idempotent [simp]: "abs (abs a) = abs (a::'a::ordered_ring)"
 33.1194 -  by (force elim: order_less_asym simp add: abs_if)
 33.1195 -
 33.1196 -lemma abs_zero_iff [simp]: "(abs a = 0) = (a = (0::'a::ordered_ring))"
 33.1197 -by (simp add: abs_if)
 33.1198 -
 33.1199 -lemma abs_ge_self: "a \<le> abs (a::'a::ordered_ring)"
 33.1200 -apply (simp add: abs_if)
 33.1201 -apply (simp add: order_less_imp_le order_trans [of _ 0])
 33.1202 -done
 33.1203 -
 33.1204 -lemma abs_ge_minus_self: "-a \<le> abs (a::'a::ordered_ring)"
 33.1205 -by (insert abs_ge_self [of "-a"], simp)
 33.1206 -
 33.1207  lemma nonzero_abs_inverse:
 33.1208       "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)"
 33.1209  apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq 
 33.1210 @@ -1670,72 +1463,8 @@
 33.1211  apply (simp add: nonzero_abs_divide) 
 33.1212  done
 33.1213  
 33.1214 -lemma abs_leI: "[|a \<le> b; -a \<le> b|] ==> abs a \<le> (b::'a::ordered_ring)"
 33.1215 -by (simp add: abs_if)
 33.1216 -
 33.1217 -lemma le_minus_self_iff: "(a \<le> -a) = (a \<le> (0::'a::ordered_ring))"
 33.1218 -proof 
 33.1219 -  assume ale: "a \<le> -a"
 33.1220 -  show "a\<le>0"
 33.1221 -    apply (rule classical) 
 33.1222 -    apply (simp add: linorder_not_le) 
 33.1223 -    apply (blast intro: ale order_trans order_less_imp_le
 33.1224 -                        neg_0_le_iff_le [THEN iffD1]) 
 33.1225 -    done
 33.1226 -next
 33.1227 -  assume "a\<le>0"
 33.1228 -  hence "0 \<le> -a" by (simp only: neg_0_le_iff_le)
 33.1229 -  thus "a \<le> -a"  by (blast intro: prems order_trans) 
 33.1230 -qed
 33.1231 -
 33.1232 -lemma minus_le_self_iff: "(-a \<le> a) = (0 \<le> (a::'a::ordered_ring))"
 33.1233 -by (insert le_minus_self_iff [of "-a"], simp)
 33.1234 -
 33.1235 -lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_ring))"
 33.1236 -by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
 33.1237 -
 33.1238 -lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_ring))"
 33.1239 -by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
 33.1240 -
 33.1241 -lemma abs_le_D1: "abs a \<le> b ==> a \<le> (b::'a::ordered_ring)"
 33.1242 -apply (simp add: abs_if split: split_if_asm)
 33.1243 -apply (rule order_trans [of _ "-a"]) 
 33.1244 - apply (simp add: less_minus_self_iff order_less_imp_le, assumption)
 33.1245 -done
 33.1246 -
 33.1247 -lemma abs_le_D2: "abs a \<le> b ==> -a \<le> (b::'a::ordered_ring)"
 33.1248 -by (insert abs_le_D1 [of "-a"], simp)
 33.1249 -
 33.1250 -lemma abs_le_iff: "(abs a \<le> b) = (a \<le> b & -a \<le> (b::'a::ordered_ring))"
 33.1251 -by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
 33.1252 -
 33.1253 -lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_ring))" 
 33.1254 -apply (simp add: order_less_le abs_le_iff)  
 33.1255 -apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff)
 33.1256 -apply (simp add: le_minus_self_iff linorder_neq_iff) 
 33.1257 -done
 33.1258 -(*
 33.1259 -apply (simp add: order_less_le abs_le_iff)  
 33.1260 -apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff) 
 33.1261 - apply (simp add:  linorder_not_less [symmetric])
 33.1262 -apply (simp add: le_minus_self_iff linorder_neq_iff) 
 33.1263 -apply (simp add:  linorder_not_less [symmetric]) 
 33.1264 -done
 33.1265 -*)
 33.1266 -
 33.1267 -lemma abs_triangle_ineq: "abs (a+b) \<le> abs a + abs (b::'a::ordered_ring)"
 33.1268 -by (force simp add: abs_le_iff abs_ge_self abs_ge_minus_self add_mono)
 33.1269 -
 33.1270 -lemma abs_diff_triangle_ineq:
 33.1271 -     "\<bar>(a::'a::ordered_ring) + b - (c+d)\<bar> \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>"
 33.1272 -proof -
 33.1273 -  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
 33.1274 -  also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
 33.1275 -  finally show ?thesis .
 33.1276 -qed
 33.1277 -
 33.1278  lemma abs_mult_less:
 33.1279 -     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_ring)"
 33.1280 +     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_idom)"
 33.1281  proof -
 33.1282    assume ac: "abs a < c"
 33.1283    hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
 33.1284 @@ -1743,272 +1472,221 @@
 33.1285    thus ?thesis by (simp add: ac cpos mult_strict_mono) 
 33.1286  qed
 33.1287  
 33.1288 +lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_idom))"
 33.1289 +by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
 33.1290 +
 33.1291 +lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_idom))"
 33.1292 +by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
 33.1293 +
 33.1294 +lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_idom))" 
 33.1295 +apply (simp add: order_less_le abs_le_iff)  
 33.1296 +apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff)
 33.1297 +apply (simp add: le_minus_self_iff linorder_neq_iff) 
 33.1298 +done
 33.1299 +
 33.1300  text{*Moving this up spoils many proofs using @{text mult_le_cancel_right}*}
 33.1301  declare times_divide_eq_left [simp]
 33.1302  
 33.1303 -ML
 33.1304 -{*
 33.1305 -val add_assoc = thm"add_assoc";
 33.1306 -val add_commute = thm"add_commute";
 33.1307 -val mult_assoc = thm"mult_assoc";
 33.1308 -val mult_commute = thm"mult_commute";
 33.1309 -val zero_neq_one = thm"zero_neq_one";
 33.1310 -val diff_minus = thm"diff_minus";
 33.1311 -val abs_if = thm"abs_if";
 33.1312 -val divide_inverse = thm"divide_inverse";
 33.1313 -val inverse_zero = thm"inverse_zero";
 33.1314 -val divide_zero = thm"divide_zero";
 33.1315 -
 33.1316 -val add_0 = thm"add_0";
 33.1317 -val add_0_right = thm"add_0_right";
 33.1318 -val add_zero_left = thm"add_0";
 33.1319 -val add_zero_right = thm"add_0_right";
 33.1320 -
 33.1321 -val add_left_commute = thm"add_left_commute";
 33.1322 -val left_minus = thm"left_minus";
 33.1323 -val right_minus = thm"right_minus";
 33.1324 -val right_minus_eq = thm"right_minus_eq";
 33.1325 -val add_left_cancel = thm"add_left_cancel";
 33.1326 -val add_right_cancel = thm"add_right_cancel";
 33.1327 -val minus_minus = thm"minus_minus";
 33.1328 -val equals_zero_I = thm"equals_zero_I";
 33.1329 -val minus_zero = thm"minus_zero";
 33.1330 -val diff_self = thm"diff_self";
 33.1331 -val diff_0 = thm"diff_0";
 33.1332 -val diff_0_right = thm"diff_0_right";
 33.1333 -val diff_minus_eq_add = thm"diff_minus_eq_add";
 33.1334 -val neg_equal_iff_equal = thm"neg_equal_iff_equal";
 33.1335 -val neg_equal_0_iff_equal = thm"neg_equal_0_iff_equal";
 33.1336 -val neg_0_equal_iff_equal = thm"neg_0_equal_iff_equal";
 33.1337 -val equation_minus_iff = thm"equation_minus_iff";
 33.1338 -val minus_equation_iff = thm"minus_equation_iff";
 33.1339 -val mult_1 = thm"mult_1";
 33.1340 -val mult_1_right = thm"mult_1_right";
 33.1341 -val mult_left_commute = thm"mult_left_commute";
 33.1342 -val mult_zero_left = thm"mult_zero_left";
 33.1343 -val mult_zero_right = thm"mult_zero_right";
 33.1344 +ML {*
 33.1345  val left_distrib = thm "left_distrib";
 33.1346 -val right_distrib = thm"right_distrib";
 33.1347 -val combine_common_factor = thm"combine_common_factor";
 33.1348 -val minus_add_distrib = thm"minus_add_distrib";
 33.1349 -val minus_mult_left = thm"minus_mult_left";
 33.1350 -val minus_mult_right = thm"minus_mult_right";
 33.1351 -val minus_mult_minus = thm"minus_mult_minus";
 33.1352 -val minus_mult_commute = thm"minus_mult_commute";
 33.1353 -val right_diff_distrib = thm"right_diff_distrib";
 33.1354 -val left_diff_distrib = thm"left_diff_distrib";
 33.1355 -val minus_diff_eq = thm"minus_diff_eq";
 33.1356 -val add_left_mono = thm"add_left_mono";
 33.1357 -val add_right_mono = thm"add_right_mono";
 33.1358 -val add_mono = thm"add_mono";
 33.1359 -val add_strict_left_mono = thm"add_strict_left_mono";
 33.1360 -val add_strict_right_mono = thm"add_strict_right_mono";
 33.1361 -val add_strict_mono = thm"add_strict_mono";
 33.1362 -val add_less_le_mono = thm"add_less_le_mono";
 33.1363 -val add_le_less_mono = thm"add_le_less_mono";
 33.1364 -val add_less_imp_less_left = thm"add_less_imp_less_left";
 33.1365 -val add_less_imp_less_right = thm"add_less_imp_less_right";
 33.1366 -val add_less_cancel_left = thm"add_less_cancel_left";
 33.1367 -val add_less_cancel_right = thm"add_less_cancel_right";
 33.1368 -val add_le_cancel_left = thm"add_le_cancel_left";
 33.1369 -val add_le_cancel_right = thm"add_le_cancel_right";
 33.1370 -val add_le_imp_le_left = thm"add_le_imp_le_left";
 33.1371 -val add_le_imp_le_right = thm"add_le_imp_le_right";
 33.1372 -val le_imp_neg_le = thm"le_imp_neg_le";
 33.1373 -val neg_le_iff_le = thm"neg_le_iff_le";
 33.1374 -val neg_le_0_iff_le = thm"neg_le_0_iff_le";
 33.1375 -val neg_0_le_iff_le = thm"neg_0_le_iff_le";
 33.1376 -val neg_less_iff_less = thm"neg_less_iff_less";
 33.1377 -val neg_less_0_iff_less = thm"neg_less_0_iff_less";
 33.1378 -val neg_0_less_iff_less = thm"neg_0_less_iff_less";
 33.1379 -val less_minus_iff = thm"less_minus_iff";
 33.1380 -val minus_less_iff = thm"minus_less_iff";
 33.1381 -val le_minus_iff = thm"le_minus_iff";
 33.1382 -val minus_le_iff = thm"minus_le_iff";
 33.1383 -val add_diff_eq = thm"add_diff_eq";
 33.1384 -val diff_add_eq = thm"diff_add_eq";
 33.1385 -val diff_eq_eq = thm"diff_eq_eq";
 33.1386 -val eq_diff_eq = thm"eq_diff_eq";
 33.1387 -val diff_diff_eq = thm"diff_diff_eq";
 33.1388 -val diff_diff_eq2 = thm"diff_diff_eq2";
 33.1389 -val less_iff_diff_less_0 = thm"less_iff_diff_less_0";
 33.1390 -val diff_less_eq = thm"diff_less_eq";
 33.1391 -val less_diff_eq = thm"less_diff_eq";
 33.1392 -val diff_le_eq = thm"diff_le_eq";
 33.1393 -val le_diff_eq = thm"le_diff_eq";
 33.1394 -val eq_iff_diff_eq_0 = thm"eq_iff_diff_eq_0";
 33.1395 -val le_iff_diff_le_0 = thm"le_iff_diff_le_0";
 33.1396 -val eq_add_iff1 = thm"eq_add_iff1";
 33.1397 -val eq_add_iff2 = thm"eq_add_iff2";
 33.1398 -val less_add_iff1 = thm"less_add_iff1";
 33.1399 -val less_add_iff2 = thm"less_add_iff2";
 33.1400 -val le_add_iff1 = thm"le_add_iff1";
 33.1401 -val le_add_iff2 = thm"le_add_iff2";
 33.1402 -val mult_strict_left_mono = thm"mult_strict_left_mono";
 33.1403 -val mult_strict_right_mono = thm"mult_strict_right_mono";
 33.1404 -val mult_left_mono = thm"mult_left_mono";
 33.1405 -val mult_right_mono = thm"mult_right_mono";
 33.1406 -val mult_left_le_imp_le = thm"mult_left_le_imp_le";
 33.1407 -val mult_right_le_imp_le = thm"mult_right_le_imp_le";
 33.1408 -val mult_left_less_imp_less = thm"mult_left_less_imp_less";
 33.1409 -val mult_right_less_imp_less = thm"mult_right_less_imp_less";
 33.1410 -val mult_strict_left_mono_neg = thm"mult_strict_left_mono_neg";
 33.1411 -val mult_strict_right_mono_neg = thm"mult_strict_right_mono_neg";
 33.1412 -val mult_pos = thm"mult_pos";
 33.1413 -val mult_pos_neg = thm"mult_pos_neg";
 33.1414 -val mult_neg = thm"mult_neg";
 33.1415 -val zero_less_mult_pos = thm"zero_less_mult_pos";
 33.1416 -val zero_less_mult_iff = thm"zero_less_mult_iff";
 33.1417 -val mult_eq_0_iff = thm"mult_eq_0_iff";
 33.1418 -val zero_le_mult_iff = thm"zero_le_mult_iff";
 33.1419 -val mult_less_0_iff = thm"mult_less_0_iff";
 33.1420 -val mult_le_0_iff = thm"mult_le_0_iff";
 33.1421 -val zero_le_square = thm"zero_le_square";
 33.1422 -val zero_less_one = thm"zero_less_one";
 33.1423 -val zero_le_one = thm"zero_le_one";
 33.1424 -val not_one_less_zero = thm"not_one_less_zero";
 33.1425 -val not_one_le_zero = thm"not_one_le_zero";
 33.1426 -val mult_left_mono_neg = thm"mult_left_mono_neg";
 33.1427 -val mult_right_mono_neg = thm"mult_right_mono_neg";
 33.1428 -val mult_strict_mono = thm"mult_strict_mono";
 33.1429 -val mult_strict_mono' = thm"mult_strict_mono'";
 33.1430 -val mult_mono = thm"mult_mono";
 33.1431 -val mult_less_cancel_right = thm"mult_less_cancel_right";
 33.1432 -val mult_less_cancel_left = thm"mult_less_cancel_left";
 33.1433 -val mult_le_cancel_right = thm"mult_le_cancel_right";
 33.1434 -val mult_le_cancel_left = thm"mult_le_cancel_left";
 33.1435 -val mult_less_imp_less_left = thm"mult_less_imp_less_left";
 33.1436 -val mult_less_imp_less_right = thm"mult_less_imp_less_right";
 33.1437 -val mult_cancel_right = thm"mult_cancel_right";
 33.1438 -val mult_cancel_left = thm"mult_cancel_left";
 33.1439 +val right_distrib = thm "right_distrib";
 33.1440 +val mult_commute = thm "mult_commute";
 33.1441 +val distrib = thm "distrib";
 33.1442 +val zero_neq_one = thm "zero_neq_one";
 33.1443 +val no_zero_divisors = thm "no_zero_divisors";
 33.1444  val left_inverse = thm "left_inverse";
 33.1445 -val right_inverse = thm"right_inverse";
 33.1446 -val right_inverse_eq = thm"right_inverse_eq";
 33.1447 -val nonzero_inverse_eq_divide = thm"nonzero_inverse_eq_divide";
 33.1448 -val divide_self = thm"divide_self";
 33.1449 -val inverse_divide = thm"inverse_divide";
 33.1450 -val divide_zero_left = thm"divide_zero_left";
 33.1451 -val inverse_eq_divide = thm"inverse_eq_divide";
 33.1452 -val add_divide_distrib = thm"add_divide_distrib";
 33.1453 -val field_mult_eq_0_iff = thm"field_mult_eq_0_iff";
 33.1454 -val field_mult_cancel_right = thm"field_mult_cancel_right";
 33.1455 -val field_mult_cancel_left = thm"field_mult_cancel_left";
 33.1456 -val nonzero_imp_inverse_nonzero = thm"nonzero_imp_inverse_nonzero";
 33.1457 -val inverse_zero_imp_zero = thm"inverse_zero_imp_zero";
 33.1458 -val inverse_nonzero_imp_nonzero = thm"inverse_nonzero_imp_nonzero";
 33.1459 -val inverse_nonzero_iff_nonzero = thm"inverse_nonzero_iff_nonzero";
 33.1460 -val nonzero_inverse_minus_eq = thm"nonzero_inverse_minus_eq";
 33.1461 -val inverse_minus_eq = thm"inverse_minus_eq";
 33.1462 -val nonzero_inverse_eq_imp_eq = thm"nonzero_inverse_eq_imp_eq";
 33.1463 -val inverse_eq_imp_eq = thm"inverse_eq_imp_eq";
 33.1464 -val inverse_eq_iff_eq = thm"inverse_eq_iff_eq";
 33.1465 -val nonzero_inverse_inverse_eq = thm"nonzero_inverse_inverse_eq";
 33.1466 -val inverse_inverse_eq = thm"inverse_inverse_eq";
 33.1467 -val inverse_1 = thm"inverse_1";
 33.1468 -val nonzero_inverse_mult_distrib = thm"nonzero_inverse_mult_distrib";
 33.1469 -val inverse_mult_distrib = thm"inverse_mult_distrib";
 33.1470 -val inverse_add = thm"inverse_add";
 33.1471 -val nonzero_mult_divide_cancel_left = thm"nonzero_mult_divide_cancel_left";
 33.1472 -val mult_divide_cancel_left = thm"mult_divide_cancel_left";
 33.1473 -val nonzero_mult_divide_cancel_right = thm"nonzero_mult_divide_cancel_right";
 33.1474 -val mult_divide_cancel_right = thm"mult_divide_cancel_right";
 33.1475 -val mult_divide_cancel_eq_if = thm"mult_divide_cancel_eq_if";
 33.1476 -val divide_1 = thm"divide_1";
 33.1477 -val times_divide_eq_right = thm"times_divide_eq_right";
 33.1478 -val times_divide_eq_left = thm"times_divide_eq_left";
 33.1479 -val divide_divide_eq_right = thm"divide_divide_eq_right";
 33.1480 -val divide_divide_eq_left = thm"divide_divide_eq_left";
 33.1481 -val nonzero_minus_divide_left = thm"nonzero_minus_divide_left";
 33.1482 -val nonzero_minus_divide_right = thm"nonzero_minus_divide_right";
 33.1483 -val nonzero_minus_divide_divide = thm"nonzero_minus_divide_divide";
 33.1484 -val minus_divide_left = thm"minus_divide_left";
 33.1485 -val minus_divide_right = thm"minus_divide_right";
 33.1486 -val minus_divide_divide = thm"minus_divide_divide";
 33.1487 -val positive_imp_inverse_positive = thm"positive_imp_inverse_positive";
 33.1488 -val negative_imp_inverse_negative = thm"negative_imp_inverse_negative";
 33.1489 -val inverse_le_imp_le = thm"inverse_le_imp_le";
 33.1490 -val inverse_positive_imp_positive = thm"inverse_positive_imp_positive";
 33.1491 -val inverse_positive_iff_positive = thm"inverse_positive_iff_positive";
 33.1492 -val inverse_negative_imp_negative = thm"inverse_negative_imp_negative";
 33.1493 -val inverse_negative_iff_negative = thm"inverse_negative_iff_negative";
 33.1494 -val inverse_nonnegative_iff_nonnegative = thm"inverse_nonnegative_iff_nonnegative";
 33.1495 -val inverse_nonpositive_iff_nonpositive = thm"inverse_nonpositive_iff_nonpositive";
 33.1496 -val less_imp_inverse_less = thm"less_imp_inverse_less";
 33.1497 -val inverse_less_imp_less = thm"inverse_less_imp_less";
 33.1498 -val inverse_less_iff_less = thm"inverse_less_iff_less";
 33.1499 -val le_imp_inverse_le = thm"le_imp_inverse_le";
 33.1500 -val inverse_le_iff_le = thm"inverse_le_iff_le";
 33.1501 -val inverse_le_imp_le_neg = thm"inverse_le_imp_le_neg";
 33.1502 -val less_imp_inverse_less_neg = thm"less_imp_inverse_less_neg";
 33.1503 -val inverse_less_imp_less_neg = thm"inverse_less_imp_less_neg";
 33.1504 -val inverse_less_iff_less_neg = thm"inverse_less_iff_less_neg";
 33.1505 -val le_imp_inverse_le_neg = thm"le_imp_inverse_le_neg";
 33.1506 -val inverse_le_iff_le_neg = thm"inverse_le_iff_le_neg";
 33.1507 -val zero_less_divide_iff = thm"zero_less_divide_iff";
 33.1508 -val divide_less_0_iff = thm"divide_less_0_iff";
 33.1509 -val zero_le_divide_iff = thm"zero_le_divide_iff";
 33.1510 -val divide_le_0_iff = thm"divide_le_0_iff";
 33.1511 -val divide_eq_0_iff = thm"divide_eq_0_iff";
 33.1512 -val pos_le_divide_eq = thm"pos_le_divide_eq";
 33.1513 -val neg_le_divide_eq = thm"neg_le_divide_eq";
 33.1514 -val le_divide_eq = thm"le_divide_eq";
 33.1515 -val pos_divide_le_eq = thm"pos_divide_le_eq";
 33.1516 -val neg_divide_le_eq = thm"neg_divide_le_eq";
 33.1517 -val divide_le_eq = thm"divide_le_eq";
 33.1518 -val pos_less_divide_eq = thm"pos_less_divide_eq";
 33.1519 -val neg_less_divide_eq = thm"neg_less_divide_eq";
 33.1520 -val less_divide_eq = thm"less_divide_eq";
 33.1521 -val pos_divide_less_eq = thm"pos_divide_less_eq";
 33.1522 -val neg_divide_less_eq = thm"neg_divide_less_eq";
 33.1523 -val divide_less_eq = thm"divide_less_eq";
 33.1524 -val nonzero_eq_divide_eq = thm"nonzero_eq_divide_eq";
 33.1525 -val eq_divide_eq = thm"eq_divide_eq";
 33.1526 -val nonzero_divide_eq_eq = thm"nonzero_divide_eq_eq";
 33.1527 -val divide_eq_eq = thm"divide_eq_eq";
 33.1528 -val divide_cancel_right = thm"divide_cancel_right";
 33.1529 -val divide_cancel_left = thm"divide_cancel_left";
 33.1530 -val divide_strict_right_mono = thm"divide_strict_right_mono";
 33.1531 -val divide_right_mono = thm"divide_right_mono";
 33.1532 -val divide_strict_left_mono = thm"divide_strict_left_mono";
 33.1533 -val divide_left_mono = thm"divide_left_mono";
 33.1534 -val divide_strict_left_mono_neg = thm"divide_strict_left_mono_neg";
 33.1535 -val divide_strict_right_mono_neg = thm"divide_strict_right_mono_neg";
 33.1536 -val zero_less_two = thm"zero_less_two";
 33.1537 -val less_half_sum = thm"less_half_sum";
 33.1538 -val gt_half_sum = thm"gt_half_sum";
 33.1539 -val dense = thm"dense";
 33.1540 -val abs_zero = thm"abs_zero";
 33.1541 -val abs_one = thm"abs_one";
 33.1542 -val abs_mult = thm"abs_mult";
 33.1543 -val abs_mult_self = thm"abs_mult_self";
 33.1544 -val abs_eq_0 = thm"abs_eq_0";
 33.1545 -val zero_less_abs_iff = thm"zero_less_abs_iff";
 33.1546 -val abs_not_less_zero = thm"abs_not_less_zero";
 33.1547 -val abs_le_zero_iff = thm"abs_le_zero_iff";
 33.1548 -val abs_minus_cancel = thm"abs_minus_cancel";
 33.1549 -val abs_ge_zero = thm"abs_ge_zero";
 33.1550 -val abs_idempotent = thm"abs_idempotent";
 33.1551 -val abs_zero_iff = thm"abs_zero_iff";
 33.1552 -val abs_ge_self = thm"abs_ge_self";
 33.1553 -val abs_ge_minus_self = thm"abs_ge_minus_self";
 33.1554 -val nonzero_abs_inverse = thm"nonzero_abs_inverse";
 33.1555 -val abs_inverse = thm"abs_inverse";
 33.1556 -val nonzero_abs_divide = thm"nonzero_abs_divide";
 33.1557 -val abs_divide = thm"abs_divide";
 33.1558 -val abs_leI = thm"abs_leI";
 33.1559 -val le_minus_self_iff = thm"le_minus_self_iff";
 33.1560 -val minus_le_self_iff = thm"minus_le_self_iff";
 33.1561 -val eq_minus_self_iff = thm"eq_minus_self_iff";
 33.1562 -val less_minus_self_iff = thm"less_minus_self_iff";
 33.1563 -val abs_le_D1 = thm"abs_le_D1";
 33.1564 -val abs_le_D2 = thm"abs_le_D2";
 33.1565 -val abs_le_iff = thm"abs_le_iff";
 33.1566 -val abs_less_iff = thm"abs_less_iff";
 33.1567 -val abs_triangle_ineq = thm"abs_triangle_ineq";
 33.1568 -val abs_mult_less = thm"abs_mult_less";
 33.1569 -
 33.1570 -val compare_rls = thms"compare_rls";
 33.1571 +val divide_inverse = thm "divide_inverse";
 33.1572 +val mult_zero_left = thm "mult_zero_left";
 33.1573 +val mult_zero_right = thm "mult_zero_right";
 33.1574 +val field_mult_eq_0_iff = thm "field_mult_eq_0_iff";
 33.1575 +val inverse_zero = thm "inverse_zero";
 33.1576 +val ring_distrib = thms "ring_distrib";
 33.1577 +val combine_common_factor = thm "combine_common_factor";
 33.1578 +val minus_mult_left = thm "minus_mult_left";
 33.1579 +val minus_mult_right = thm "minus_mult_right";
 33.1580 +val minus_mult_minus = thm "minus_mult_minus";
 33.1581 +val minus_mult_commute = thm "minus_mult_commute";
 33.1582 +val right_diff_distrib = thm "right_diff_distrib";
 33.1583 +val left_diff_distrib = thm "left_diff_distrib";
 33.1584 +val mult_left_mono = thm "mult_left_mono";
 33.1585 +val mult_right_mono = thm "mult_right_mono";
 33.1586 +val mult_strict_left_mono = thm "mult_strict_left_mono";
 33.1587 +val mult_strict_right_mono = thm "mult_strict_right_mono";
 33.1588 +val mult_mono = thm "mult_mono";
 33.1589 +val mult_strict_mono = thm "mult_strict_mono";
 33.1590 +val abs_if = thm "abs_if";
 33.1591 +val zero_less_one = thm "zero_less_one";
 33.1592 +val eq_add_iff1 = thm "eq_add_iff1";
 33.1593 +val eq_add_iff2 = thm "eq_add_iff2";
 33.1594 +val less_add_iff1 = thm "less_add_iff1";
 33.1595 +val less_add_iff2 = thm "less_add_iff2";
 33.1596 +val le_add_iff1 = thm "le_add_iff1";
 33.1597 +val le_add_iff2 = thm "le_add_iff2";
 33.1598 +val mult_left_le_imp_le = thm "mult_left_le_imp_le";
 33.1599 +val mult_right_le_imp_le = thm "mult_right_le_imp_le";
 33.1600 +val mult_left_less_imp_less = thm "mult_left_less_imp_less";
 33.1601 +val mult_right_less_imp_less = thm "mult_right_less_imp_less";
 33.1602 +val mult_strict_left_mono_neg = thm "mult_strict_left_mono_neg";
 33.1603 +val mult_left_mono_neg = thm "mult_left_mono_neg";
 33.1604 +val mult_strict_right_mono_neg = thm "mult_strict_right_mono_neg";
 33.1605 +val mult_right_mono_neg = thm "mult_right_mono_neg";
 33.1606 +val mult_pos = thm "mult_pos";
 33.1607 +val mult_pos_le = thm "mult_pos_le";
 33.1608 +val mult_pos_neg = thm "mult_pos_neg";
 33.1609 +val mult_pos_neg_le = thm "mult_pos_neg_le";
 33.1610 +val mult_pos_neg2 = thm "mult_pos_neg2";
 33.1611 +val mult_pos_neg2_le = thm "mult_pos_neg2_le";
 33.1612 +val mult_neg = thm "mult_neg";
 33.1613 +val mult_neg_le = thm "mult_neg_le";
 33.1614 +val zero_less_mult_pos = thm "zero_less_mult_pos";
 33.1615 +val zero_less_mult_pos2 = thm "zero_less_mult_pos2";
 33.1616 +val zero_less_mult_iff = thm "zero_less_mult_iff";
 33.1617 +val mult_eq_0_iff = thm "mult_eq_0_iff";
 33.1618 +val zero_le_mult_iff = thm "zero_le_mult_iff";
 33.1619 +val mult_less_0_iff = thm "mult_less_0_iff";
 33.1620 +val mult_le_0_iff = thm "mult_le_0_iff";
 33.1621 +val split_mult_pos_le = thm "split_mult_pos_le";
 33.1622 +val split_mult_neg_le = thm "split_mult_neg_le";
 33.1623 +val zero_le_square = thm "zero_le_square";
 33.1624 +val zero_le_one = thm "zero_le_one";
 33.1625 +val not_one_le_zero = thm "not_one_le_zero";
 33.1626 +val not_one_less_zero = thm "not_one_less_zero";
 33.1627 +val mult_left_mono_neg = thm "mult_left_mono_neg";
 33.1628 +val mult_right_mono_neg = thm "mult_right_mono_neg";
 33.1629 +val mult_strict_mono = thm "mult_strict_mono";
 33.1630 +val mult_strict_mono' = thm "mult_strict_mono'";
 33.1631 +val mult_mono = thm "mult_mono";
 33.1632 +val less_1_mult = thm "less_1_mult";
 33.1633 +val mult_less_cancel_right = thm "mult_less_cancel_right";
 33.1634 +val mult_less_cancel_left = thm "mult_less_cancel_left";
 33.1635 +val mult_le_cancel_right = thm "mult_le_cancel_right";
 33.1636 +val mult_le_cancel_left = thm "mult_le_cancel_left";
 33.1637 +val mult_less_imp_less_left = thm "mult_less_imp_less_left";
 33.1638 +val mult_less_imp_less_right = thm "mult_less_imp_less_right";
 33.1639 +val mult_cancel_right = thm "mult_cancel_right";
 33.1640 +val mult_cancel_left = thm "mult_cancel_left";
 33.1641 +val ring_eq_simps = thms "ring_eq_simps";
 33.1642 +val right_inverse = thm "right_inverse";
 33.1643 +val right_inverse_eq = thm "right_inverse_eq";
 33.1644 +val nonzero_inverse_eq_divide = thm "nonzero_inverse_eq_divide";
 33.1645 +val divide_self = thm "divide_self";
 33.1646 +val divide_zero = thm "divide_zero";
 33.1647 +val divide_zero_left = thm "divide_zero_left";
 33.1648 +val inverse_eq_divide = thm "inverse_eq_divide";
 33.1649 +val add_divide_distrib = thm "add_divide_distrib";
 33.1650 +val field_mult_eq_0_iff = thm "field_mult_eq_0_iff";
 33.1651 +val field_mult_cancel_right_lemma = thm "field_mult_cancel_right_lemma";
 33.1652 +val field_mult_cancel_right = thm "field_mult_cancel_right";
 33.1653 +val field_mult_cancel_left = thm "field_mult_cancel_left";
 33.1654 +val nonzero_imp_inverse_nonzero = thm "nonzero_imp_inverse_nonzero";
 33.1655 +val inverse_zero_imp_zero = thm "inverse_zero_imp_zero";
 33.1656 +val inverse_nonzero_imp_nonzero = thm "inverse_nonzero_imp_nonzero";
 33.1657 +val inverse_nonzero_iff_nonzero = thm "inverse_nonzero_iff_nonzero";
 33.1658 +val nonzero_inverse_minus_eq = thm "nonzero_inverse_minus_eq";
 33.1659 +val inverse_minus_eq = thm "inverse_minus_eq";
 33.1660 +val nonzero_inverse_eq_imp_eq = thm "nonzero_inverse_eq_imp_eq";
 33.1661 +val inverse_eq_imp_eq = thm "inverse_eq_imp_eq";
 33.1662 +val inverse_eq_iff_eq = thm "inverse_eq_iff_eq";
 33.1663 +val nonzero_inverse_inverse_eq = thm "nonzero_inverse_inverse_eq";
 33.1664 +val inverse_inverse_eq = thm "inverse_inverse_eq";
 33.1665 +val inverse_1 = thm "inverse_1";
 33.1666 +val nonzero_inverse_mult_distrib = thm "nonzero_inverse_mult_distrib";
 33.1667 +val inverse_mult_distrib = thm "inverse_mult_distrib";
 33.1668 +val inverse_add = thm "inverse_add";
 33.1669 +val inverse_divide = thm "inverse_divide";
 33.1670 +val nonzero_mult_divide_cancel_left = thm "nonzero_mult_divide_cancel_left";
 33.1671 +val mult_divide_cancel_left = thm "mult_divide_cancel_left";
 33.1672 +val nonzero_mult_divide_cancel_right = thm "nonzero_mult_divide_cancel_right";
 33.1673 +val mult_divide_cancel_right = thm "mult_divide_cancel_right";
 33.1674 +val mult_divide_cancel_eq_if = thm "mult_divide_cancel_eq_if";
 33.1675 +val divide_1 = thm "divide_1";
 33.1676 +val times_divide_eq_right = thm "times_divide_eq_right";
 33.1677 +val times_divide_eq_left = thm "times_divide_eq_left";
 33.1678 +val divide_divide_eq_right = thm "divide_divide_eq_right";
 33.1679 +val divide_divide_eq_left = thm "divide_divide_eq_left";
 33.1680 +val nonzero_minus_divide_left = thm "nonzero_minus_divide_left";
 33.1681 +val nonzero_minus_divide_right = thm "nonzero_minus_divide_right";
 33.1682 +val nonzero_minus_divide_divide = thm "nonzero_minus_divide_divide";
 33.1683 +val minus_divide_left = thm "minus_divide_left";
 33.1684 +val minus_divide_right = thm "minus_divide_right";
 33.1685 +val minus_divide_divide = thm "minus_divide_divide";
 33.1686 +val diff_divide_distrib = thm "diff_divide_distrib";
 33.1687 +val positive_imp_inverse_positive = thm "positive_imp_inverse_positive";
 33.1688 +val negative_imp_inverse_negative = thm "negative_imp_inverse_negative";
 33.1689 +val inverse_le_imp_le = thm "inverse_le_imp_le";
 33.1690 +val inverse_positive_imp_positive = thm "inverse_positive_imp_positive";
 33.1691 +val inverse_positive_iff_positive = thm "inverse_positive_iff_positive";
 33.1692 +val inverse_negative_imp_negative = thm "inverse_negative_imp_negative";
 33.1693 +val inverse_negative_iff_negative = thm "inverse_negative_iff_negative";
 33.1694 +val inverse_nonnegative_iff_nonnegative = thm "inverse_nonnegative_iff_nonnegative";
 33.1695 +val inverse_nonpositive_iff_nonpositive = thm "inverse_nonpositive_iff_nonpositive";
 33.1696 +val less_imp_inverse_less = thm "less_imp_inverse_less";
 33.1697 +val inverse_less_imp_less = thm "inverse_less_imp_less";
 33.1698 +val inverse_less_iff_less = thm "inverse_less_iff_less";
 33.1699 +val le_imp_inverse_le = thm "le_imp_inverse_le";
 33.1700 +val inverse_le_iff_le = thm "inverse_le_iff_le";
 33.1701 +val inverse_le_imp_le_neg = thm "inverse_le_imp_le_neg";
 33.1702 +val less_imp_inverse_less_neg = thm "less_imp_inverse_less_neg";
 33.1703 +val inverse_less_imp_less_neg = thm "inverse_less_imp_less_neg";
 33.1704 +val inverse_less_iff_less_neg = thm "inverse_less_iff_less_neg";
 33.1705 +val le_imp_inverse_le_neg = thm "le_imp_inverse_le_neg";
 33.1706 +val inverse_le_iff_le_neg = thm "inverse_le_iff_le_neg";
 33.1707 +val one_less_inverse_iff = thm "one_less_inverse_iff";
 33.1708 +val inverse_eq_1_iff = thm "inverse_eq_1_iff";
 33.1709 +val one_le_inverse_iff = thm "one_le_inverse_iff";
 33.1710 +val inverse_less_1_iff = thm "inverse_less_1_iff";
 33.1711 +val inverse_le_1_iff = thm "inverse_le_1_iff";
 33.1712 +val zero_less_divide_iff = thm "zero_less_divide_iff";
 33.1713 +val divide_less_0_iff = thm "divide_less_0_iff";
 33.1714 +val zero_le_divide_iff = thm "zero_le_divide_iff";
 33.1715 +val divide_le_0_iff = thm "divide_le_0_iff";
 33.1716 +val divide_eq_0_iff = thm "divide_eq_0_iff";
 33.1717 +val pos_le_divide_eq = thm "pos_le_divide_eq";
 33.1718 +val neg_le_divide_eq = thm "neg_le_divide_eq";
 33.1719 +val le_divide_eq = thm "le_divide_eq";
 33.1720 +val pos_divide_le_eq = thm "pos_divide_le_eq";
 33.1721 +val neg_divide_le_eq = thm "neg_divide_le_eq";
 33.1722 +val divide_le_eq = thm "divide_le_eq";
 33.1723 +val pos_less_divide_eq = thm "pos_less_divide_eq";
 33.1724 +val neg_less_divide_eq = thm "neg_less_divide_eq";
 33.1725 +val less_divide_eq = thm "less_divide_eq";
 33.1726 +val pos_divide_less_eq = thm "pos_divide_less_eq";
 33.1727 +val neg_divide_less_eq = thm "neg_divide_less_eq";
 33.1728 +val divide_less_eq = thm "divide_less_eq";
 33.1729 +val nonzero_eq_divide_eq = thm "nonzero_eq_divide_eq";
 33.1730 +val eq_divide_eq = thm "eq_divide_eq";
 33.1731 +val nonzero_divide_eq_eq = thm "nonzero_divide_eq_eq";
 33.1732 +val divide_eq_eq = thm "divide_eq_eq";
 33.1733 +val divide_cancel_right = thm "divide_cancel_right";
 33.1734 +val divide_cancel_left = thm "divide_cancel_left";
 33.1735 +val divide_eq_1_iff = thm "divide_eq_1_iff";
 33.1736 +val one_eq_divide_iff = thm "one_eq_divide_iff";
 33.1737 +val zero_eq_1_divide_iff = thm "zero_eq_1_divide_iff";
 33.1738 +val one_divide_eq_0_iff = thm "one_divide_eq_0_iff";
 33.1739 +val divide_strict_right_mono = thm "divide_strict_right_mono";
 33.1740 +val divide_right_mono = thm "divide_right_mono";
 33.1741 +val divide_strict_left_mono = thm "divide_strict_left_mono";
 33.1742 +val divide_left_mono = thm "divide_left_mono";
 33.1743 +val divide_strict_left_mono_neg = thm "divide_strict_left_mono_neg";
 33.1744 +val divide_strict_right_mono_neg = thm "divide_strict_right_mono_neg";
 33.1745 +val less_add_one = thm "less_add_one";
 33.1746 +val zero_less_two = thm "zero_less_two";
 33.1747 +val less_half_sum = thm "less_half_sum";
 33.1748 +val gt_half_sum = thm "gt_half_sum";
 33.1749 +val dense = thm "dense";
 33.1750 +val abs_one = thm "abs_one";
 33.1751 +val abs_le_mult = thm "abs_le_mult";
 33.1752 +val abs_eq_mult = thm "abs_eq_mult";
 33.1753 +val abs_mult = thm "abs_mult";
 33.1754 +val abs_mult_self = thm "abs_mult_self";
 33.1755 +val nonzero_abs_inverse = thm "nonzero_abs_inverse";
 33.1756 +val abs_inverse = thm "abs_inverse";
 33.1757 +val nonzero_abs_divide = thm "nonzero_abs_divide";
 33.1758 +val abs_divide = thm "abs_divide";
 33.1759 +val abs_mult_less = thm "abs_mult_less";
 33.1760 +val eq_minus_self_iff = thm "eq_minus_self_iff";
 33.1761 +val less_minus_self_iff = thm "less_minus_self_iff";
 33.1762 +val abs_less_iff = thm "abs_less_iff";
 33.1763  *}
 33.1764  
 33.1765 -
 33.1766  end
    34.1 --- a/src/HOL/UNITY/Comp/AllocBase.thy	Tue May 11 14:00:02 2004 +0200
    34.2 +++ b/src/HOL/UNITY/Comp/AllocBase.thy	Tue May 11 20:11:08 2004 +0200
    34.3 @@ -54,7 +54,7 @@
    34.4  
    34.5  lemma bag_of_append [simp]: "bag_of (l@l') = bag_of l + bag_of l'"
    34.6  apply (induct_tac "l", simp)
    34.7 - apply (simp add: plus_ac0)
    34.8 + apply (simp add: add_ac)
    34.9  done
   34.10  
   34.11  lemma mono_bag_of: "mono (bag_of :: 'a list => ('a::order) multiset)"
   34.12 @@ -80,7 +80,7 @@
   34.13        (\<Sum>i: A Int lessThan (length l). {# l!i #})"
   34.14  apply (rule_tac xs = l in rev_induct, simp)
   34.15  apply (simp add: sublist_append Int_insert_right lessThan_Suc nth_append 
   34.16 -                 bag_of_sublist_lemma plus_ac0)
   34.17 +                 bag_of_sublist_lemma add_ac)
   34.18  done
   34.19  
   34.20  
    35.1 --- a/src/HOL/arith_data.ML	Tue May 11 14:00:02 2004 +0200
    35.2 +++ b/src/HOL/arith_data.ML	Tue May 11 20:11:08 2004 +0200
    35.3 @@ -386,10 +386,10 @@
    35.4  val add_mono_thms_ordered_semiring = map (fn s => prove_goal (the_context ()) s
    35.5   (fn prems => [cut_facts_tac prems 1,
    35.6                 blast_tac (claset() addIs [add_mono]) 1]))
    35.7 -["(i <= j) & (k <= l) ==> i + k <= j + (l::'a::ordered_semiring)",
    35.8 - "(i  = j) & (k <= l) ==> i + k <= j + (l::'a::ordered_semiring)",
    35.9 - "(i <= j) & (k  = l) ==> i + k <= j + (l::'a::ordered_semiring)",
   35.10 - "(i  = j) & (k  = l) ==> i + k  = j + (l::'a::ordered_semiring)"
   35.11 +["(i <= j) & (k <= l) ==> i + k <= j + (l::'a::ordered_semidom)",
   35.12 + "(i  = j) & (k <= l) ==> i + k <= j + (l::'a::ordered_semidom)",
   35.13 + "(i <= j) & (k  = l) ==> i + k <= j + (l::'a::ordered_semidom)",
   35.14 + "(i  = j) & (k  = l) ==> i + k  = j + (l::'a::ordered_semidom)"
   35.15  ];
   35.16  
   35.17  in
    36.1 --- a/src/HOL/ex/Lagrange.thy	Tue May 11 14:00:02 2004 +0200
    36.2 +++ b/src/HOL/ex/Lagrange.thy	Tue May 11 20:11:08 2004 +0200
    36.3 @@ -23,7 +23,7 @@
    36.4  
    36.5  (*once a slow step, but now (2001) just three seconds!*)
    36.6  lemma Lagrange_lemma:
    36.7 - "!!x1::'a::ring.
    36.8 + "!!x1::'a::comm_ring_1.
    36.9    (sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) =
   36.10    sq(x1*y1 - x2*y2 - x3*y3 - x4*y4)  +
   36.11    sq(x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
   36.12 @@ -34,7 +34,7 @@
   36.13  
   36.14  (* A challenge by John Harrison. Takes about 4 mins on a 3GHz machine.
   36.15  
   36.16 -lemma "!!p1::'a::ring.
   36.17 +lemma "!!p1::'a::comm_ring_1.
   36.18   (sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * 
   36.19   (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) 
   36.20    = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) +