--- a/NEWS Mon Feb 08 15:25:00 2010 +0100
+++ b/NEWS Mon Feb 08 17:13:45 2010 +0100
@@ -12,6 +12,14 @@
*** HOL ***
+* New set of rules "ac_simps" provides combined assoc / commute rewrites
+for all interpretations of the appropriate generic locales.
+
+* Renamed theory "OrderedGroup" to "Groups" and split theory "Ring_and_Field"
+into theories "Rings" and "Fields"; for more appropriate and more
+consistent names suitable for name prefixes within the HOL theories.
+INCOMPATIBILITY.
+
* More consistent naming of type classes involving orderings (and lattices):
lower_semilattice ~> semilattice_inf
--- a/src/HOL/Algebra/poly/UnivPoly2.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy Mon Feb 08 17:13:45 2010 +0100
@@ -139,7 +139,7 @@
end
-instance up :: ("{times, comm_monoid_add}") Ring_and_Field.dvd ..
+instance up :: ("{times, comm_monoid_add}") Rings.dvd ..
instantiation up :: ("{times, one, comm_monoid_add, uminus, minus}") inverse
begin
--- a/src/HOL/Decision_Procs/Decision_Procs.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Decision_Procs/Decision_Procs.thy Mon Feb 08 17:13:45 2010 +0100
@@ -2,7 +2,8 @@
theory Decision_Procs
imports
- Commutative_Ring Cooper Ferrack MIR Approximation Dense_Linear_Order Parametric_Ferrante_Rackoff
+ Commutative_Ring Cooper Ferrack MIR Approximation Dense_Linear_Order
+ Parametric_Ferrante_Rackoff
Commutative_Ring_Complete
"ex/Commutative_Ring_Ex" "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex"
begin
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Mon Feb 08 17:13:45 2010 +0100
@@ -637,7 +637,7 @@
using eq_diff_eq[where a= x and b=t and c=0] by simp
-interpretation class_dense_linlinordered_field: constr_dense_linorder
+interpretation class_dense_linordered_field: constr_dense_linorder
"op <=" "op <"
"\<lambda> x y. 1/2 * ((x::'a::{linordered_field,number_ring}) + y)"
proof (unfold_locales, dlo, dlo, auto)
@@ -871,7 +871,7 @@
addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
in
-Ferrante_Rackoff_Data.funs @{thm "class_dense_linlinordered_field.ferrack_axiom"}
+Ferrante_Rackoff_Data.funs @{thm "class_dense_linordered_field.ferrack_axiom"}
{isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
end
*}
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Feb 08 17:13:45 2010 +0100
@@ -7,9 +7,9 @@
theory Parametric_Ferrante_Rackoff
imports Reflected_Multivariate_Polynomial
"~~/src/HOL/Decision_Procs/Dense_Linear_Order"
+ Efficient_Nat
begin
-
subsection {* Terms *}
datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm
@@ -2594,13 +2594,6 @@
by (simp add: Let_def stupid)
-
-(*
-lemmas [code func] = polysub_def
-lemmas [code func del] = Zero_nat_def
-code_gen "frpar" in SML to FRParTest
-*)
-
section{* Second implemenation: Case splits not local *}
lemma fr_eq2: assumes lp: "islin p"
@@ -2945,14 +2938,7 @@
show ?thesis unfolding frpar2_def by (auto simp add: prep)
qed
-code_module FRPar
- contains
- frpar = "frpar"
- frpar2 = "frpar2"
- test = "%x . frpar (E(Lt (Mul 1\<^sub>p (Bound 0))))"
-
-ML{*
-
+ML {*
structure ReflectedFRPar =
struct
@@ -2983,92 +2969,93 @@
fun num_of_term m t =
case t of
- Const(@{const_name Algebras.uminus},_)$t => FRPar.Neg (num_of_term m t)
- | Const(@{const_name Algebras.plus},_)$a$b => FRPar.Add (num_of_term m a, num_of_term m b)
- | Const(@{const_name Algebras.minus},_)$a$b => FRPar.Sub (num_of_term m a, num_of_term m b)
- | Const(@{const_name Algebras.times},_)$a$b => FRPar.Mul (num_of_term m a, num_of_term m b)
- | Const(@{const_name Power.power},_)$a$n => FRPar.Pw (num_of_term m a, dest_nat n)
- | Const(@{const_name Algebras.divide},_)$a$b => FRPar.C (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
- | _ => (FRPar.C (HOLogic.dest_number t |> snd,1)
- handle TERM _ => FRPar.Bound (AList.lookup (op aconv) m t |> the));
+ Const(@{const_name Algebras.uminus},_)$t => @{code poly.Neg} (num_of_term m t)
+ | Const(@{const_name Algebras.plus},_)$a$b => @{code poly.Add} (num_of_term m a, num_of_term m b)
+ | Const(@{const_name Algebras.minus},_)$a$b => @{code poly.Sub} (num_of_term m a, num_of_term m b)
+ | Const(@{const_name Algebras.times},_)$a$b => @{code poly.Mul} (num_of_term m a, num_of_term m b)
+ | Const(@{const_name Power.power},_)$a$n => @{code poly.Pw} (num_of_term m a, dest_nat n)
+ | Const(@{const_name Algebras.divide},_)$a$b => @{code poly.C} (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
+ | _ => (@{code poly.C} (HOLogic.dest_number t |> snd,1)
+ handle TERM _ => @{code poly.Bound} (AList.lookup (op aconv) m t |> the));
fun tm_of_term m m' t =
case t of
- Const(@{const_name Algebras.uminus},_)$t => FRPar.tm_Neg (tm_of_term m m' t)
- | Const(@{const_name Algebras.plus},_)$a$b => FRPar.tm_Add (tm_of_term m m' a, tm_of_term m m' b)
- | Const(@{const_name Algebras.minus},_)$a$b => FRPar.tm_Sub (tm_of_term m m' a, tm_of_term m m' b)
- | Const(@{const_name Algebras.times},_)$a$b => FRPar.tm_Mul (num_of_term m' a, tm_of_term m m' b)
- | _ => (FRPar.CP (num_of_term m' t)
- handle TERM _ => FRPar.tm_Bound (AList.lookup (op aconv) m t |> the)
- | Option => FRPar.tm_Bound (AList.lookup (op aconv) m t |> the));
+ Const(@{const_name Algebras.uminus},_)$t => @{code Neg} (tm_of_term m m' t)
+ | Const(@{const_name Algebras.plus},_)$a$b => @{code Add} (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name Algebras.minus},_)$a$b => @{code Sub} (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name Algebras.times},_)$a$b => @{code Mul} (num_of_term m' a, tm_of_term m m' b)
+ | _ => (@{code CP} (num_of_term m' t)
+ handle TERM _ => @{code Bound} (AList.lookup (op aconv) m t |> the)
+ | Option => @{code Bound} (AList.lookup (op aconv) m t |> the));
fun term_of_num T m t =
case t of
- FRPar.C (a,b) => (if b = 1 then num T a else if b=0 then (rz T)
+ @{code poly.C} (a,b) => (if b = 1 then num T a else if b=0 then (rz T)
else (divt T) $ num T a $ num T b)
-| FRPar.Bound i => AList.lookup (op = : int*int -> bool) m i |> the
-| FRPar.Add(a,b) => (plust T)$(term_of_num T m a)$(term_of_num T m b)
-| FRPar.Mul(a,b) => (timest T)$(term_of_num T m a)$(term_of_num T m b)
-| FRPar.Sub(a,b) => (minust T)$(term_of_num T m a)$(term_of_num T m b)
-| FRPar.Neg a => (uminust T)$(term_of_num T m a)
-| FRPar.Pw(a,n) => (powt T)$(term_of_num T m t)$(HOLogic.mk_number HOLogic.natT n)
-| FRPar.CN(c,n,p) => term_of_num T m (FRPar.Add(c,FRPar.Mul(FRPar.Bound n, p)))
+| @{code poly.Bound} i => AList.lookup (op = : int*int -> bool) m i |> the
+| @{code poly.Add} (a,b) => (plust T)$(term_of_num T m a)$(term_of_num T m b)
+| @{code poly.Mul} (a,b) => (timest T)$(term_of_num T m a)$(term_of_num T m b)
+| @{code poly.Sub} (a,b) => (minust T)$(term_of_num T m a)$(term_of_num T m b)
+| @{code poly.Neg} a => (uminust T)$(term_of_num T m a)
+| @{code poly.Pw} (a,n) => (powt T)$(term_of_num T m t)$(HOLogic.mk_number HOLogic.natT n)
+| @{code poly.CN} (c,n,p) => term_of_num T m (@{code poly.Add} (c, @{code poly.Mul} (@{code poly.Bound} n, p)))
| _ => error "term_of_num: Unknown term";
fun term_of_tm T m m' t =
case t of
- FRPar.CP p => term_of_num T m' p
-| FRPar.tm_Bound i => AList.lookup (op = : int*int -> bool) m i |> the
-| FRPar.tm_Add(a,b) => (plust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
-| FRPar.tm_Mul(a,b) => (timest T)$(term_of_num T m' a)$(term_of_tm T m m' b)
-| FRPar.tm_Sub(a,b) => (minust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
-| FRPar.tm_Neg a => (uminust T)$(term_of_tm T m m' a)
-| FRPar.CNP(n,c,p) => term_of_tm T m m' (FRPar.tm_Add(FRPar.tm_Mul(c, FRPar.tm_Bound n), p))
+ @{code CP} p => term_of_num T m' p
+| @{code Bound} i => AList.lookup (op = : int*int -> bool) m i |> the
+| @{code Add} (a,b) => (plust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| @{code Mul} (a,b) => (timest T)$(term_of_num T m' a)$(term_of_tm T m m' b)
+| @{code Sub} (a,b) => (minust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| @{code Neg} a => (uminust T)$(term_of_tm T m m' a)
+| @{code CNP} (n,c,p) => term_of_tm T m m' (@{code Add}
+ (@{code Mul} (c, @{code Bound} n), p))
| _ => error "term_of_tm: Unknown term";
fun fm_of_term m m' fm =
case fm of
- Const("True",_) => FRPar.T
- | Const("False",_) => FRPar.F
- | Const("Not",_)$p => FRPar.NOT (fm_of_term m m' p)
- | Const("op &",_)$p$q => FRPar.And(fm_of_term m m' p, fm_of_term m m' q)
- | Const("op |",_)$p$q => FRPar.Or(fm_of_term m m' p, fm_of_term m m' q)
- | Const("op -->",_)$p$q => FRPar.Imp(fm_of_term m m' p, fm_of_term m m' q)
+ Const("True",_) => @{code T}
+ | Const("False",_) => @{code F}
+ | Const("Not",_)$p => @{code NOT} (fm_of_term m m' p)
+ | Const("op &",_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
+ | Const("op |",_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
+ | Const("op -->",_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
| Const("op =",ty)$p$q =>
- if domain_type ty = bT then FRPar.Iff(fm_of_term m m' p, fm_of_term m m' q)
- else FRPar.Eq (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ if domain_type ty = bT then @{code Iff} (fm_of_term m m' p, fm_of_term m m' q)
+ else @{code Eq} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
| Const(@{const_name Algebras.less},_)$p$q =>
- FRPar.Lt (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ @{code Lt} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
| Const(@{const_name Algebras.less_eq},_)$p$q =>
- FRPar.Le (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ @{code Le} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
| Const("Ex",_)$Abs(xn,xT,p) =>
let val (xn', p') = variant_abs (xn,xT,p)
val x = Free(xn',xT)
fun incr i = i + 1
val m0 = (x,0):: (map (apsnd incr) m)
- in FRPar.E (fm_of_term m0 m' p') end
+ in @{code E} (fm_of_term m0 m' p') end
| Const("All",_)$Abs(xn,xT,p) =>
let val (xn', p') = variant_abs (xn,xT,p)
val x = Free(xn',xT)
fun incr i = i + 1
val m0 = (x,0):: (map (apsnd incr) m)
- in FRPar.A (fm_of_term m0 m' p') end
+ in @{code A} (fm_of_term m0 m' p') end
| _ => error "fm_of_term";
fun term_of_fm T m m' t =
case t of
- FRPar.T => Const("True",bT)
- | FRPar.F => Const("False",bT)
- | FRPar.NOT p => nott $ (term_of_fm T m m' p)
- | FRPar.And (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
- | FRPar.Or (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
- | FRPar.Imp (p,q) => impt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
- | FRPar.Iff (p,q) => ifft $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
- | FRPar.Lt p => (llt T) $ (term_of_tm T m m' p) $ (rz T)
- | FRPar.Le p => (lle T) $ (term_of_tm T m m' p) $ (rz T)
- | FRPar.Eq p => (eqt T) $ (term_of_tm T m m' p) $ (rz T)
- | FRPar.NEq p => nott $ ((eqt T) $ (term_of_tm T m m' p) $ (rz T))
+ @{code T} => Const("True",bT)
+ | @{code F} => Const("False",bT)
+ | @{code NOT} p => nott $ (term_of_fm T m m' p)
+ | @{code And} (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | @{code Or} (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | @{code Imp} (p,q) => impt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | @{code Iff} (p,q) => ifft $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | @{code Lt} p => (llt T) $ (term_of_tm T m m' p) $ (rz T)
+ | @{code Le} p => (lle T) $ (term_of_tm T m m' p) $ (rz T)
+ | @{code Eq} p => (eqt T) $ (term_of_tm T m m' p) $ (rz T)
+ | @{code NEq} p => nott $ ((eqt T) $ (term_of_tm T m m' p) $ (rz T))
| _ => error "term_of_fm: quantifiers!!!!???";
fun frpar_oracle (T,m, m', fm) =
@@ -3077,7 +3064,7 @@
val im = 0 upto (length m - 1)
val im' = 0 upto (length m' - 1)
in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')
- (FRPar.frpar (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ (@{code frpar} (fm_of_term (m ~~ im) (m' ~~ im') t))))
end;
fun frpar_oracle2 (T,m, m', fm) =
@@ -3086,7 +3073,7 @@
val im = 0 upto (length m - 1)
val im' = 0 upto (length m' - 1)
in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')
- (FRPar.frpar2 (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ (@{code frpar2} (fm_of_term (m ~~ im) (m' ~~ im') t))))
end;
end;
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Feb 08 17:13:45 2010 +0100
@@ -2,14 +2,13 @@
Author: Amine Chaieb
*)
-header {* Implementation and verification of mutivariate polynomials Library *}
-
+header {* Implementation and verification of multivariate polynomials *}
theory Reflected_Multivariate_Polynomial
-imports Parity Abstract_Rat Efficient_Nat List Polynomial_List
+imports Complex_Main Abstract_Rat Polynomial_List
begin
- (* Impelementation *)
+ (* Implementation *)
subsection{* Datatype of polynomial expressions *}
--- a/src/HOL/Decision_Procs/mir_tac.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Mon Feb 08 17:13:45 2010 +0100
@@ -33,7 +33,7 @@
@{thm "real_of_nat_number_of"},
@{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"},
@{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"},
- @{thm "Ring_and_Field.divide_zero"},
+ @{thm "Fields.divide_zero"},
@{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"},
@{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
@{thm "diff_def"}, @{thm "minus_divide_left"}]
--- a/src/HOL/Divides.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Divides.thy Mon Feb 08 17:13:45 2010 +0100
@@ -657,7 +657,7 @@
val trans = trans;
val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
- (@{thm monoid_add_class.add_0_left} :: @{thm monoid_add_class.add_0_right} :: @{thms add_ac}))
+ (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac}))
end)
@@ -1655,8 +1655,8 @@
lemmas arithmetic_simps =
arith_simps
add_special
- OrderedGroup.add_0_left
- OrderedGroup.add_0_right
+ add_0_left
+ add_0_right
mult_zero_left
mult_zero_right
mult_1_left
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Fields.thy Mon Feb 08 17:13:45 2010 +0100
@@ -0,0 +1,1044 @@
+(* Title: HOL/Fields.thy
+ Author: Gertrud Bauer
+ Author: Steven Obua
+ Author: Tobias Nipkow
+ Author: Lawrence C Paulson
+ Author: Markus Wenzel
+ Author: Jeremy Avigad
+*)
+
+header {* Fields *}
+
+theory Fields
+imports Rings
+begin
+
+class field = comm_ring_1 + inverse +
+ assumes field_inverse: "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
+ assumes divide_inverse: "a / b = a * inverse b"
+begin
+
+subclass division_ring
+proof
+ fix a :: 'a
+ assume "a \<noteq> 0"
+ thus "inverse a * a = 1" by (rule field_inverse)
+ thus "a * inverse a = 1" by (simp only: mult_commute)
+qed
+
+subclass idom ..
+
+lemma right_inverse_eq: "b \<noteq> 0 \<Longrightarrow> a / b = 1 \<longleftrightarrow> a = b"
+proof
+ assume neq: "b \<noteq> 0"
+ {
+ hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
+ also assume "a / b = 1"
+ finally show "a = b" by simp
+ next
+ assume "a = b"
+ with neq show "a / b = 1" by (simp add: divide_inverse)
+ }
+qed
+
+lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
+by (simp add: divide_inverse)
+
+lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
+by (simp add: divide_inverse)
+
+lemma divide_zero_left [simp]: "0 / a = 0"
+by (simp add: divide_inverse)
+
+lemma inverse_eq_divide: "inverse a = 1 / a"
+by (simp add: divide_inverse)
+
+lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
+by (simp add: divide_inverse algebra_simps)
+
+text{*There is no slick version using division by zero.*}
+lemma inverse_add:
+ "[| a \<noteq> 0; b \<noteq> 0 |]
+ ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
+by (simp add: division_ring_inverse_add mult_ac)
+
+lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]:
+assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
+proof -
+ have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
+ by (simp add: divide_inverse nonzero_inverse_mult_distrib)
+ also have "... = a * inverse b * (inverse c * c)"
+ by (simp only: mult_ac)
+ also have "... = a * inverse b" by simp
+ finally show ?thesis by (simp add: divide_inverse)
+qed
+
+lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]:
+ "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
+by (simp add: mult_commute [of _ c])
+
+lemma divide_1 [simp]: "a / 1 = a"
+by (simp add: divide_inverse)
+
+lemma times_divide_eq_right: "a * (b / c) = (a * b) / c"
+by (simp add: divide_inverse mult_assoc)
+
+lemma times_divide_eq_left: "(b / c) * a = (b * a) / c"
+by (simp add: divide_inverse mult_ac)
+
+text {* These are later declared as simp rules. *}
+lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left
+
+lemma add_frac_eq:
+ assumes "y \<noteq> 0" and "z \<noteq> 0"
+ shows "x / y + w / z = (x * z + w * y) / (y * z)"
+proof -
+ have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)"
+ using assms by simp
+ also have "\<dots> = (x * z + y * w) / (y * z)"
+ by (simp only: add_divide_distrib)
+ finally show ?thesis
+ by (simp only: mult_commute)
+qed
+
+text{*Special Cancellation Simprules for Division*}
+
+lemma nonzero_mult_divide_cancel_right [simp, noatp]:
+ "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
+using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
+
+lemma nonzero_mult_divide_cancel_left [simp, noatp]:
+ "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
+using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
+
+lemma nonzero_divide_mult_cancel_right [simp, noatp]:
+ "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
+using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
+
+lemma nonzero_divide_mult_cancel_left [simp, noatp]:
+ "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
+using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
+
+lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]:
+ "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
+using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
+
+lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]:
+ "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
+using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
+
+lemma minus_divide_left: "- (a / b) = (-a) / b"
+by (simp add: divide_inverse)
+
+lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
+by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
+by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma divide_minus_left [simp, noatp]: "(-a) / b = - (a / b)"
+by (simp add: divide_inverse)
+
+lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
+by (simp add: diff_minus add_divide_distrib)
+
+lemma add_divide_eq_iff:
+ "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
+by (simp add: add_divide_distrib)
+
+lemma divide_add_eq_iff:
+ "z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
+by (simp add: add_divide_distrib)
+
+lemma diff_divide_eq_iff:
+ "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
+by (simp add: diff_divide_distrib)
+
+lemma divide_diff_eq_iff:
+ "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
+by (simp add: diff_divide_distrib)
+
+lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
+proof -
+ assume [simp]: "c \<noteq> 0"
+ have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
+ also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
+proof -
+ assume [simp]: "c \<noteq> 0"
+ have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
+ also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
+by simp
+
+lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
+by (erule subst, simp)
+
+lemmas field_eq_simps[noatp] = algebra_simps
+ (* pull / out*)
+ add_divide_eq_iff divide_add_eq_iff
+ diff_divide_eq_iff divide_diff_eq_iff
+ (* multiply eqn *)
+ nonzero_eq_divide_eq nonzero_divide_eq_eq
+(* is added later:
+ times_divide_eq_left times_divide_eq_right
+*)
+
+text{*An example:*}
+lemma "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f\<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
+apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
+ apply(simp add:field_eq_simps)
+apply(simp)
+done
+
+lemma diff_frac_eq:
+ "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y - w / z = (x * z - w * y) / (y * z)"
+by (simp add: field_eq_simps times_divide_eq)
+
+lemma frac_eq_eq:
+ "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (x / y = w / z) = (x * z = w * y)"
+by (simp add: field_eq_simps times_divide_eq)
+
+end
+
+class division_by_zero = zero + inverse +
+ assumes inverse_zero [simp]: "inverse 0 = 0"
+
+lemma divide_zero [simp]:
+ "a / 0 = (0::'a::{field,division_by_zero})"
+by (simp add: divide_inverse)
+
+lemma divide_self_if [simp]:
+ "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
+by simp
+
+class linordered_field = field + linordered_idom
+
+lemma inverse_nonzero_iff_nonzero [simp]:
+ "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
+by (force dest: inverse_zero_imp_zero)
+
+lemma inverse_minus_eq [simp]:
+ "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
+proof cases
+ assume "a=0" thus ?thesis by (simp add: inverse_zero)
+next
+ assume "a\<noteq>0"
+ thus ?thesis by (simp add: nonzero_inverse_minus_eq)
+qed
+
+lemma inverse_eq_imp_eq:
+ "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
+apply (cases "a=0 | b=0")
+ apply (force dest!: inverse_zero_imp_zero
+ simp add: eq_commute [of "0::'a"])
+apply (force dest!: nonzero_inverse_eq_imp_eq)
+done
+
+lemma inverse_eq_iff_eq [simp]:
+ "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))"
+by (force dest!: inverse_eq_imp_eq)
+
+lemma inverse_inverse_eq [simp]:
+ "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
+ proof cases
+ assume "a=0" thus ?thesis by simp
+ next
+ assume "a\<noteq>0"
+ thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
+ qed
+
+text{*This version builds in division by zero while also re-orienting
+ the right-hand side.*}
+lemma inverse_mult_distrib [simp]:
+ "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
+ proof cases
+ assume "a \<noteq> 0 & b \<noteq> 0"
+ thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute)
+ next
+ assume "~ (a \<noteq> 0 & b \<noteq> 0)"
+ thus ?thesis by force
+ qed
+
+lemma inverse_divide [simp]:
+ "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
+by (simp add: divide_inverse mult_commute)
+
+
+subsection {* Calculations with fractions *}
+
+text{* There is a whole bunch of simp-rules just for class @{text
+field} but none for class @{text field} and @{text nonzero_divides}
+because the latter are covered by a simproc. *}
+
+lemma mult_divide_mult_cancel_left:
+ "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
+apply (cases "b = 0")
+apply (simp_all add: nonzero_mult_divide_mult_cancel_left)
+done
+
+lemma mult_divide_mult_cancel_right:
+ "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
+apply (cases "b = 0")
+apply (simp_all add: nonzero_mult_divide_mult_cancel_right)
+done
+
+lemma divide_divide_eq_right [simp,noatp]:
+ "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
+by (simp add: divide_inverse mult_ac)
+
+lemma divide_divide_eq_left [simp,noatp]:
+ "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
+by (simp add: divide_inverse mult_assoc)
+
+
+subsubsection{*Special Cancellation Simprules for Division*}
+
+lemma mult_divide_mult_cancel_left_if[simp,noatp]:
+fixes c :: "'a :: {field,division_by_zero}"
+shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
+by (simp add: mult_divide_mult_cancel_left)
+
+
+subsection {* Division and Unary Minus *}
+
+lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
+by (simp add: divide_inverse)
+
+lemma divide_minus_right [simp, noatp]:
+ "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
+by (simp add: divide_inverse)
+
+lemma minus_divide_divide:
+ "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
+apply (cases "b=0", simp)
+apply (simp add: nonzero_minus_divide_divide)
+done
+
+lemma eq_divide_eq:
+ "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
+by (simp add: nonzero_eq_divide_eq)
+
+lemma divide_eq_eq:
+ "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
+by (force simp add: nonzero_divide_eq_eq)
+
+
+subsection {* Ordered Fields *}
+
+lemma positive_imp_inverse_positive:
+assumes a_gt_0: "0 < a" shows "0 < inverse (a::'a::linordered_field)"
+proof -
+ have "0 < a * inverse a"
+ by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one)
+ thus "0 < inverse a"
+ by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff)
+qed
+
+lemma negative_imp_inverse_negative:
+ "a < 0 ==> inverse a < (0::'a::linordered_field)"
+by (insert positive_imp_inverse_positive [of "-a"],
+ simp add: nonzero_inverse_minus_eq order_less_imp_not_eq)
+
+lemma inverse_le_imp_le:
+assumes invle: "inverse a \<le> inverse b" and apos: "0 < a"
+shows "b \<le> (a::'a::linordered_field)"
+proof (rule classical)
+ assume "~ b \<le> a"
+ hence "a < b" by (simp add: linorder_not_le)
+ hence bpos: "0 < b" by (blast intro: apos order_less_trans)
+ hence "a * inverse a \<le> a * inverse b"
+ by (simp add: apos invle order_less_imp_le mult_left_mono)
+ hence "(a * inverse a) * b \<le> (a * inverse b) * b"
+ by (simp add: bpos order_less_imp_le mult_right_mono)
+ thus "b \<le> a" by (simp add: mult_assoc apos bpos order_less_imp_not_eq2)
+qed
+
+lemma inverse_positive_imp_positive:
+assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
+shows "0 < (a::'a::linordered_field)"
+proof -
+ have "0 < inverse (inverse a)"
+ using inv_gt_0 by (rule positive_imp_inverse_positive)
+ thus "0 < a"
+ using nz by (simp add: nonzero_inverse_inverse_eq)
+qed
+
+lemma inverse_positive_iff_positive [simp]:
+ "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))"
+apply (cases "a = 0", simp)
+apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
+done
+
+lemma inverse_negative_imp_negative:
+assumes inv_less_0: "inverse a < 0" and nz: "a \<noteq> 0"
+shows "a < (0::'a::linordered_field)"
+proof -
+ have "inverse (inverse a) < 0"
+ using inv_less_0 by (rule negative_imp_inverse_negative)
+ thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq)
+qed
+
+lemma inverse_negative_iff_negative [simp]:
+ "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))"
+apply (cases "a = 0", simp)
+apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
+done
+
+lemma inverse_nonnegative_iff_nonnegative [simp]:
+ "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_by_zero}))"
+by (simp add: linorder_not_less [symmetric])
+
+lemma inverse_nonpositive_iff_nonpositive [simp]:
+ "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_by_zero}))"
+by (simp add: linorder_not_less [symmetric])
+
+lemma linordered_field_no_lb: "\<forall> x. \<exists>y. y < (x::'a::linordered_field)"
+proof
+ fix x::'a
+ have m1: "- (1::'a) < 0" by simp
+ from add_strict_right_mono[OF m1, where c=x]
+ have "(- 1) + x < x" by simp
+ thus "\<exists>y. y < x" by blast
+qed
+
+lemma linordered_field_no_ub: "\<forall> x. \<exists>y. y > (x::'a::linordered_field)"
+proof
+ fix x::'a
+ have m1: " (1::'a) > 0" by simp
+ from add_strict_right_mono[OF m1, where c=x]
+ have "1 + x > x" by simp
+ thus "\<exists>y. y > x" by blast
+qed
+
+subsection{*Anti-Monotonicity of @{term inverse}*}
+
+lemma less_imp_inverse_less:
+assumes less: "a < b" and apos: "0 < a"
+shows "inverse b < inverse (a::'a::linordered_field)"
+proof (rule ccontr)
+ assume "~ inverse b < inverse a"
+ hence "inverse a \<le> inverse b" by (simp add: linorder_not_less)
+ hence "~ (a < b)"
+ by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
+ thus False by (rule notE [OF _ less])
+qed
+
+lemma inverse_less_imp_less:
+ "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)"
+apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"])
+apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq)
+done
+
+text{*Both premises are essential. Consider -1 and 1.*}
+lemma inverse_less_iff_less [simp,noatp]:
+ "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
+by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less)
+
+lemma le_imp_inverse_le:
+ "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
+by (force simp add: order_le_less less_imp_inverse_less)
+
+lemma inverse_le_iff_le [simp,noatp]:
+ "[|0 < a; 0 < b|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
+by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le)
+
+
+text{*These results refer to both operands being negative. The opposite-sign
+case is trivial, since inverse preserves signs.*}
+lemma inverse_le_imp_le_neg:
+ "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::linordered_field)"
+apply (rule classical)
+apply (subgoal_tac "a < 0")
+ prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans)
+apply (insert inverse_le_imp_le [of "-b" "-a"])
+apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
+done
+
+lemma less_imp_inverse_less_neg:
+ "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)"
+apply (subgoal_tac "a < 0")
+ prefer 2 apply (blast intro: order_less_trans)
+apply (insert less_imp_inverse_less [of "-b" "-a"])
+apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
+done
+
+lemma inverse_less_imp_less_neg:
+ "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)"
+apply (rule classical)
+apply (subgoal_tac "a < 0")
+ prefer 2
+ apply (force simp add: linorder_not_less intro: order_le_less_trans)
+apply (insert inverse_less_imp_less [of "-b" "-a"])
+apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
+done
+
+lemma inverse_less_iff_less_neg [simp,noatp]:
+ "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
+apply (insert inverse_less_iff_less [of "-b" "-a"])
+apply (simp del: inverse_less_iff_less
+ add: order_less_imp_not_eq nonzero_inverse_minus_eq)
+done
+
+lemma le_imp_inverse_le_neg:
+ "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
+by (force simp add: order_le_less less_imp_inverse_less_neg)
+
+lemma inverse_le_iff_le_neg [simp,noatp]:
+ "[|a < 0; b < 0|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
+by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg)
+
+
+subsection{*Inverses and the Number One*}
+
+lemma one_less_inverse_iff:
+ "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))"
+proof cases
+ assume "0 < x"
+ with inverse_less_iff_less [OF zero_less_one, of x]
+ show ?thesis by simp
+next
+ assume notless: "~ (0 < x)"
+ have "~ (1 < inverse x)"
+ proof
+ assume "1 < inverse x"
+ also with notless have "... \<le> 0" by (simp add: linorder_not_less)
+ also have "... < 1" by (rule zero_less_one)
+ finally show False by auto
+ qed
+ with notless show ?thesis by simp
+qed
+
+lemma inverse_eq_1_iff [simp]:
+ "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
+by (insert inverse_eq_iff_eq [of x 1], simp)
+
+lemma one_le_inverse_iff:
+ "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_by_zero}))"
+by (force simp add: order_le_less one_less_inverse_iff zero_less_one
+ eq_commute [of 1])
+
+lemma inverse_less_1_iff:
+ "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_by_zero}))"
+by (simp add: linorder_not_le [symmetric] one_le_inverse_iff)
+
+lemma inverse_le_1_iff:
+ "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_by_zero}))"
+by (simp add: linorder_not_less [symmetric] one_less_inverse_iff)
+
+
+subsection{*Simplification of Inequalities Involving Literal Divisors*}
+
+lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
+proof -
+ assume less: "0<c"
+ hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
+ by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
+ also have "... = (a*c \<le> b)"
+ by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
+proof -
+ assume less: "c<0"
+ hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
+ by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
+ also have "... = (b \<le> a*c)"
+ by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma le_divide_eq:
+ "(a \<le> b/c) =
+ (if 0 < c then a*c \<le> b
+ else if c < 0 then b \<le> a*c
+ else a \<le> (0::'a::{linordered_field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff)
+done
+
+lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
+proof -
+ assume less: "0<c"
+ hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
+ by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
+ also have "... = (b \<le> a*c)"
+ by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
+proof -
+ assume less: "c<0"
+ hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
+ by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
+ also have "... = (a*c \<le> b)"
+ by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma divide_le_eq:
+ "(b/c \<le> a) =
+ (if 0 < c then b \<le> a*c
+ else if c < 0 then a*c \<le> b
+ else 0 \<le> (a::'a::{linordered_field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff)
+done
+
+lemma pos_less_divide_eq:
+ "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)"
+proof -
+ assume less: "0<c"
+ hence "(a < b/c) = (a*c < (b/c)*c)"
+ by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
+ also have "... = (a*c < b)"
+ by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma neg_less_divide_eq:
+ "c < (0::'a::linordered_field) ==> (a < b/c) = (b < a*c)"
+proof -
+ assume less: "c<0"
+ hence "(a < b/c) = ((b/c)*c < a*c)"
+ by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
+ also have "... = (b < a*c)"
+ by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma less_divide_eq:
+ "(a < b/c) =
+ (if 0 < c then a*c < b
+ else if c < 0 then b < a*c
+ else a < (0::'a::{linordered_field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff)
+done
+
+lemma pos_divide_less_eq:
+ "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)"
+proof -
+ assume less: "0<c"
+ hence "(b/c < a) = ((b/c)*c < a*c)"
+ by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
+ also have "... = (b < a*c)"
+ by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma neg_divide_less_eq:
+ "c < (0::'a::linordered_field) ==> (b/c < a) = (a*c < b)"
+proof -
+ assume less: "c<0"
+ hence "(b/c < a) = (a*c < (b/c)*c)"
+ by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
+ also have "... = (a*c < b)"
+ by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
+ finally show ?thesis .
+qed
+
+lemma divide_less_eq:
+ "(b/c < a) =
+ (if 0 < c then b < a*c
+ else if c < 0 then a*c < b
+ else 0 < (a::'a::{linordered_field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff)
+done
+
+
+subsection{*Field simplification*}
+
+text{* Lemmas @{text field_simps} multiply with denominators in in(equations)
+if they can be proved to be non-zero (for equations) or positive/negative
+(for inequations). Can be too aggressive and is therefore separate from the
+more benign @{text algebra_simps}. *}
+
+lemmas field_simps[noatp] = field_eq_simps
+ (* multiply ineqn *)
+ pos_divide_less_eq neg_divide_less_eq
+ pos_less_divide_eq neg_less_divide_eq
+ pos_divide_le_eq neg_divide_le_eq
+ pos_le_divide_eq neg_le_divide_eq
+
+text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
+of positivity/negativity needed for @{text field_simps}. Have not added @{text
+sign_simps} to @{text field_simps} because the former can lead to case
+explosions. *}
+
+lemmas sign_simps[noatp] = group_simps
+ zero_less_mult_iff mult_less_0_iff
+
+(* Only works once linear arithmetic is installed:
+text{*An example:*}
+lemma fixes a b c d e f :: "'a::linordered_field"
+shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
+ ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
+ ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
+apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
+ prefer 2 apply(simp add:sign_simps)
+apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
+ prefer 2 apply(simp add:sign_simps)
+apply(simp add:field_simps)
+done
+*)
+
+
+subsection{*Division and Signs*}
+
+lemma zero_less_divide_iff:
+ "((0::'a::{linordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
+by (simp add: divide_inverse zero_less_mult_iff)
+
+lemma divide_less_0_iff:
+ "(a/b < (0::'a::{linordered_field,division_by_zero})) =
+ (0 < a & b < 0 | a < 0 & 0 < b)"
+by (simp add: divide_inverse mult_less_0_iff)
+
+lemma zero_le_divide_iff:
+ "((0::'a::{linordered_field,division_by_zero}) \<le> a/b) =
+ (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
+by (simp add: divide_inverse zero_le_mult_iff)
+
+lemma divide_le_0_iff:
+ "(a/b \<le> (0::'a::{linordered_field,division_by_zero})) =
+ (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
+by (simp add: divide_inverse mult_le_0_iff)
+
+lemma divide_eq_0_iff [simp,noatp]:
+ "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
+by (simp add: divide_inverse)
+
+lemma divide_pos_pos:
+ "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y"
+by(simp add:field_simps)
+
+
+lemma divide_nonneg_pos:
+ "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y"
+by(simp add:field_simps)
+
+lemma divide_neg_pos:
+ "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0"
+by(simp add:field_simps)
+
+lemma divide_nonpos_pos:
+ "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0"
+by(simp add:field_simps)
+
+lemma divide_pos_neg:
+ "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0"
+by(simp add:field_simps)
+
+lemma divide_nonneg_neg:
+ "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0"
+by(simp add:field_simps)
+
+lemma divide_neg_neg:
+ "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y"
+by(simp add:field_simps)
+
+lemma divide_nonpos_neg:
+ "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y"
+by(simp add:field_simps)
+
+
+subsection{*Cancellation Laws for Division*}
+
+lemma divide_cancel_right [simp,noatp]:
+ "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (simp add: divide_inverse)
+done
+
+lemma divide_cancel_left [simp,noatp]:
+ "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
+apply (cases "c=0", simp)
+apply (simp add: divide_inverse)
+done
+
+
+subsection {* Division and the Number One *}
+
+text{*Simplify expressions equated with 1*}
+lemma divide_eq_1_iff [simp,noatp]:
+ "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
+apply (cases "b=0", simp)
+apply (simp add: right_inverse_eq)
+done
+
+lemma one_eq_divide_iff [simp,noatp]:
+ "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
+by (simp add: eq_commute [of 1])
+
+lemma zero_eq_1_divide_iff [simp,noatp]:
+ "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)"
+apply (cases "a=0", simp)
+apply (auto simp add: nonzero_eq_divide_eq)
+done
+
+lemma one_divide_eq_0_iff [simp,noatp]:
+ "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)"
+apply (cases "a=0", simp)
+apply (insert zero_neq_one [THEN not_sym])
+apply (auto simp add: nonzero_divide_eq_eq)
+done
+
+text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
+lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified]
+lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified]
+lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified]
+lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified]
+
+declare zero_less_divide_1_iff [simp,noatp]
+declare divide_less_0_1_iff [simp,noatp]
+declare zero_le_divide_1_iff [simp,noatp]
+declare divide_le_0_1_iff [simp,noatp]
+
+
+subsection {* Ordering Rules for Division *}
+
+lemma divide_strict_right_mono:
+ "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)"
+by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono
+ positive_imp_inverse_positive)
+
+lemma divide_right_mono:
+ "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_by_zero})"
+by (force simp add: divide_strict_right_mono order_le_less)
+
+lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b
+ ==> c <= 0 ==> b / c <= a / c"
+apply (drule divide_right_mono [of _ _ "- c"])
+apply auto
+done
+
+lemma divide_strict_right_mono_neg:
+ "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)"
+apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
+apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric])
+done
+
+text{*The last premise ensures that @{term a} and @{term b}
+ have the same sign*}
+lemma divide_strict_left_mono:
+ "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono)
+
+lemma divide_left_mono:
+ "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::linordered_field)"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono)
+
+lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b
+ ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
+ apply (drule divide_left_mono [of _ _ "- c"])
+ apply (auto simp add: mult_commute)
+done
+
+lemma divide_strict_left_mono_neg:
+ "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg)
+
+
+text{*Simplify quotients that are compared with the value 1.*}
+
+lemma le_divide_eq_1 [noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
+by (auto simp add: le_divide_eq)
+
+lemma divide_le_eq_1 [noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
+by (auto simp add: divide_le_eq)
+
+lemma less_divide_eq_1 [noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
+by (auto simp add: less_divide_eq)
+
+lemma divide_less_eq_1 [noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
+by (auto simp add: divide_less_eq)
+
+
+subsection{*Conditional Simplification Rules: No Case Splits*}
+
+lemma le_divide_eq_1_pos [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
+by (auto simp add: le_divide_eq)
+
+lemma le_divide_eq_1_neg [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
+by (auto simp add: le_divide_eq)
+
+lemma divide_le_eq_1_pos [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
+by (auto simp add: divide_le_eq)
+
+lemma divide_le_eq_1_neg [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
+by (auto simp add: divide_le_eq)
+
+lemma less_divide_eq_1_pos [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
+by (auto simp add: less_divide_eq)
+
+lemma less_divide_eq_1_neg [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
+by (auto simp add: less_divide_eq)
+
+lemma divide_less_eq_1_pos [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
+by (auto simp add: divide_less_eq)
+
+lemma divide_less_eq_1_neg [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
+by (auto simp add: divide_less_eq)
+
+lemma eq_divide_eq_1 [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(1 = b/a) = ((a \<noteq> 0 & a = b))"
+by (auto simp add: eq_divide_eq)
+
+lemma divide_eq_eq_1 [simp,noatp]:
+ fixes a :: "'a :: {linordered_field,division_by_zero}"
+ shows "(b/a = 1) = ((a \<noteq> 0 & a = b))"
+by (auto simp add: divide_eq_eq)
+
+
+subsection {* Reasoning about inequalities with division *}
+
+lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==>
+ x / y <= z"
+by (subst pos_divide_le_eq, assumption+)
+
+lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==>
+ z <= x / y"
+by(simp add:field_simps)
+
+lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==>
+ x / y < z"
+by(simp add:field_simps)
+
+lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==>
+ z < x / y"
+by(simp add:field_simps)
+
+lemma frac_le: "(0::'a::linordered_field) <= x ==>
+ x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w"
+ apply (rule mult_imp_div_pos_le)
+ apply simp
+ apply (subst times_divide_eq_left)
+ apply (rule mult_imp_le_div_pos, assumption)
+ apply (rule mult_mono)
+ apply simp_all
+done
+
+lemma frac_less: "(0::'a::linordered_field) <= x ==>
+ x < y ==> 0 < w ==> w <= z ==> x / z < y / w"
+ apply (rule mult_imp_div_pos_less)
+ apply simp
+ apply (subst times_divide_eq_left)
+ apply (rule mult_imp_less_div_pos, assumption)
+ apply (erule mult_less_le_imp_less)
+ apply simp_all
+done
+
+lemma frac_less2: "(0::'a::linordered_field) < x ==>
+ x <= y ==> 0 < w ==> w < z ==> x / z < y / w"
+ apply (rule mult_imp_div_pos_less)
+ apply simp_all
+ apply (subst times_divide_eq_left)
+ apply (rule mult_imp_less_div_pos, assumption)
+ apply (erule mult_le_less_imp_less)
+ apply simp_all
+done
+
+text{*It's not obvious whether these should be simprules or not.
+ Their effect is to gather terms into one big fraction, like
+ a*b*c / x*y*z. The rationale for that is unclear, but many proofs
+ seem to need them.*}
+
+declare times_divide_eq [simp]
+
+
+subsection {* Ordered Fields are Dense *}
+
+lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)"
+by (simp add: field_simps zero_less_two)
+
+lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b"
+by (simp add: field_simps zero_less_two)
+
+instance linordered_field < dense_linorder
+proof
+ fix x y :: 'a
+ have "x < x + 1" by simp
+ then show "\<exists>y. x < y" ..
+ have "x - 1 < x" by simp
+ then show "\<exists>y. y < x" ..
+ show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
+qed
+
+
+subsection {* Absolute Value *}
+
+lemma nonzero_abs_inverse:
+ "a \<noteq> 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)"
+apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq
+ negative_imp_inverse_negative)
+apply (blast intro: positive_imp_inverse_positive elim: order_less_asym)
+done
+
+lemma abs_inverse [simp]:
+ "abs (inverse (a::'a::{linordered_field,division_by_zero})) =
+ inverse (abs a)"
+apply (cases "a=0", simp)
+apply (simp add: nonzero_abs_inverse)
+done
+
+lemma nonzero_abs_divide:
+ "b \<noteq> 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b"
+by (simp add: divide_inverse abs_mult nonzero_abs_inverse)
+
+lemma abs_divide [simp]:
+ "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b"
+apply (cases "b=0", simp)
+apply (simp add: nonzero_abs_divide)
+done
+
+lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==>
+ abs x / y = abs (x / y)"
+ apply (subst abs_divide)
+ apply (simp add: order_less_imp_le)
+done
+
+code_modulename SML
+ Fields Arith
+
+code_modulename OCaml
+ Fields Arith
+
+code_modulename Haskell
+ Fields Arith
+
+end
--- a/src/HOL/Groebner_Basis.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Groebner_Basis.thy Mon Feb 08 17:13:45 2010 +0100
@@ -621,7 +621,7 @@
val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
@{thm "divide_Numeral1"},
- @{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"},
+ @{thm "Fields.divide_zero"}, @{thm "divide_Numeral0"},
@{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"},
@{thm "mult_num_frac"}, @{thm "mult_frac_num"},
@{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"},
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Groups.thy Mon Feb 08 17:13:45 2010 +0100
@@ -0,0 +1,1159 @@
+(* Title: HOL/Groups.thy
+ Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad
+*)
+
+header {* Groups, also combined with orderings *}
+
+theory Groups
+imports Lattices
+uses "~~/src/Provers/Arith/abel_cancel.ML"
+begin
+
+text {*
+ The theory of partially ordered groups is taken from the books:
+ \begin{itemize}
+ \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
+ \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+ \end{itemize}
+ Most of the used notions can also be looked up in
+ \begin{itemize}
+ \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+ \item \emph{Algebra I} by van der Waerden, Springer.
+ \end{itemize}
+*}
+
+ML {*
+structure Algebra_Simps = Named_Thms(
+ val name = "algebra_simps"
+ val description = "algebra simplification rules"
+)
+*}
+
+setup Algebra_Simps.setup
+
+text{* The rewrites accumulated in @{text algebra_simps} deal with the
+classical algebraic structures of groups, rings and family. They simplify
+terms by multiplying everything out (in case of a ring) and bringing sums and
+products into a canonical form (by ordered rewriting). As a result it decides
+group and ring equalities but also helps with inequalities.
+
+Of course it also works for fields, but it knows nothing about multiplicative
+inverses or division. This is catered for by @{text field_simps}. *}
+
+subsection {* Semigroups and Monoids *}
+
+class semigroup_add = plus +
+ assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
+
+sublocale semigroup_add < plus!: semigroup plus proof
+qed (fact add_assoc)
+
+class ab_semigroup_add = semigroup_add +
+ assumes add_commute [algebra_simps]: "a + b = b + a"
+
+sublocale ab_semigroup_add < plus!: abel_semigroup plus proof
+qed (fact add_commute)
+
+context ab_semigroup_add
+begin
+
+lemmas add_left_commute [algebra_simps] = plus.left_commute
+
+theorems add_ac = add_assoc add_commute add_left_commute
+
+end
+
+theorems add_ac = add_assoc add_commute add_left_commute
+
+class semigroup_mult = times +
+ assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)"
+
+sublocale semigroup_mult < times!: semigroup times proof
+qed (fact mult_assoc)
+
+class ab_semigroup_mult = semigroup_mult +
+ assumes mult_commute [algebra_simps]: "a * b = b * a"
+
+sublocale ab_semigroup_mult < times!: abel_semigroup times proof
+qed (fact mult_commute)
+
+context ab_semigroup_mult
+begin
+
+lemmas mult_left_commute [algebra_simps] = times.left_commute
+
+theorems mult_ac = mult_assoc mult_commute mult_left_commute
+
+end
+
+theorems mult_ac = mult_assoc mult_commute mult_left_commute
+
+class ab_semigroup_idem_mult = ab_semigroup_mult +
+ assumes mult_idem: "x * x = x"
+
+sublocale ab_semigroup_idem_mult < times!: semilattice times proof
+qed (fact mult_idem)
+
+context ab_semigroup_idem_mult
+begin
+
+lemmas mult_left_idem = times.left_idem
+
+end
+
+class monoid_add = zero + semigroup_add +
+ assumes add_0_left [simp]: "0 + a = a"
+ and add_0_right [simp]: "a + 0 = a"
+
+lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"
+by (rule eq_commute)
+
+class comm_monoid_add = zero + ab_semigroup_add +
+ assumes add_0: "0 + a = a"
+begin
+
+subclass monoid_add
+ proof qed (insert add_0, simp_all add: add_commute)
+
+end
+
+class monoid_mult = one + semigroup_mult +
+ assumes mult_1_left [simp]: "1 * a = a"
+ assumes mult_1_right [simp]: "a * 1 = a"
+
+lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"
+by (rule eq_commute)
+
+class comm_monoid_mult = one + ab_semigroup_mult +
+ assumes mult_1: "1 * a = a"
+begin
+
+subclass monoid_mult
+ proof qed (insert mult_1, simp_all add: mult_commute)
+
+end
+
+class cancel_semigroup_add = semigroup_add +
+ assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
+ assumes add_right_imp_eq: "b + a = c + a \<Longrightarrow> b = c"
+begin
+
+lemma add_left_cancel [simp]:
+ "a + b = a + c \<longleftrightarrow> b = c"
+by (blast dest: add_left_imp_eq)
+
+lemma add_right_cancel [simp]:
+ "b + a = c + a \<longleftrightarrow> b = c"
+by (blast dest: add_right_imp_eq)
+
+end
+
+class cancel_ab_semigroup_add = ab_semigroup_add +
+ assumes add_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
+begin
+
+subclass cancel_semigroup_add
+proof
+ fix a b c :: 'a
+ assume "a + b = a + c"
+ then show "b = c" by (rule add_imp_eq)
+next
+ fix a b c :: 'a
+ assume "b + a = c + a"
+ then have "a + b = a + c" by (simp only: add_commute)
+ then show "b = c" by (rule add_imp_eq)
+qed
+
+end
+
+class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add
+
+
+subsection {* Groups *}
+
+class group_add = minus + uminus + monoid_add +
+ assumes left_minus [simp]: "- a + a = 0"
+ assumes diff_minus: "a - b = a + (- b)"
+begin
+
+lemma minus_unique:
+ assumes "a + b = 0" shows "- a = b"
+proof -
+ have "- a = - a + (a + b)" using assms by simp
+ also have "\<dots> = b" by (simp add: add_assoc [symmetric])
+ finally show ?thesis .
+qed
+
+lemmas equals_zero_I = minus_unique (* legacy name *)
+
+lemma minus_zero [simp]: "- 0 = 0"
+proof -
+ have "0 + 0 = 0" by (rule add_0_right)
+ thus "- 0 = 0" by (rule minus_unique)
+qed
+
+lemma minus_minus [simp]: "- (- a) = a"
+proof -
+ have "- a + a = 0" by (rule left_minus)
+ thus "- (- a) = a" by (rule minus_unique)
+qed
+
+lemma right_minus [simp]: "a + - a = 0"
+proof -
+ have "a + - a = - (- a) + - a" by simp
+ also have "\<dots> = 0" by (rule left_minus)
+ finally show ?thesis .
+qed
+
+lemma minus_add_cancel: "- a + (a + b) = b"
+by (simp add: add_assoc [symmetric])
+
+lemma add_minus_cancel: "a + (- a + b) = b"
+by (simp add: add_assoc [symmetric])
+
+lemma minus_add: "- (a + b) = - b + - a"
+proof -
+ have "(a + b) + (- b + - a) = 0"
+ by (simp add: add_assoc add_minus_cancel)
+ thus "- (a + b) = - b + - a"
+ by (rule minus_unique)
+qed
+
+lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"
+proof
+ assume "a - b = 0"
+ have "a = (a - b) + b" by (simp add:diff_minus add_assoc)
+ also have "\<dots> = b" using `a - b = 0` by simp
+ finally show "a = b" .
+next
+ assume "a = b" thus "a - b = 0" by (simp add: diff_minus)
+qed
+
+lemma diff_self [simp]: "a - a = 0"
+by (simp add: diff_minus)
+
+lemma diff_0 [simp]: "0 - a = - a"
+by (simp add: diff_minus)
+
+lemma diff_0_right [simp]: "a - 0 = a"
+by (simp add: diff_minus)
+
+lemma diff_minus_eq_add [simp]: "a - - b = a + b"
+by (simp add: diff_minus)
+
+lemma neg_equal_iff_equal [simp]:
+ "- a = - b \<longleftrightarrow> a = b"
+proof
+ assume "- a = - b"
+ hence "- (- a) = - (- b)" by simp
+ thus "a = b" by simp
+next
+ assume "a = b"
+ thus "- a = - b" by simp
+qed
+
+lemma neg_equal_0_iff_equal [simp]:
+ "- a = 0 \<longleftrightarrow> a = 0"
+by (subst neg_equal_iff_equal [symmetric], simp)
+
+lemma neg_0_equal_iff_equal [simp]:
+ "0 = - a \<longleftrightarrow> 0 = a"
+by (subst neg_equal_iff_equal [symmetric], simp)
+
+text{*The next two equations can make the simplifier loop!*}
+
+lemma equation_minus_iff:
+ "a = - b \<longleftrightarrow> b = - a"
+proof -
+ have "- (- a) = - b \<longleftrightarrow> - a = b" by (rule neg_equal_iff_equal)
+ thus ?thesis by (simp add: eq_commute)
+qed
+
+lemma minus_equation_iff:
+ "- a = b \<longleftrightarrow> - b = a"
+proof -
+ have "- a = - (- b) \<longleftrightarrow> a = -b" by (rule neg_equal_iff_equal)
+ thus ?thesis by (simp add: eq_commute)
+qed
+
+lemma diff_add_cancel: "a - b + b = a"
+by (simp add: diff_minus add_assoc)
+
+lemma add_diff_cancel: "a + b - b = a"
+by (simp add: diff_minus add_assoc)
+
+declare diff_minus[symmetric, algebra_simps]
+
+lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
+proof
+ assume "a = - b" then show "a + b = 0" by simp
+next
+ assume "a + b = 0"
+ moreover have "a + (b + - b) = (a + b) + - b"
+ by (simp only: add_assoc)
+ ultimately show "a = - b" by simp
+qed
+
+end
+
+class ab_group_add = minus + uminus + comm_monoid_add +
+ assumes ab_left_minus: "- a + a = 0"
+ assumes ab_diff_minus: "a - b = a + (- b)"
+begin
+
+subclass group_add
+ proof qed (simp_all add: ab_left_minus ab_diff_minus)
+
+subclass cancel_comm_monoid_add
+proof
+ fix a b c :: 'a
+ assume "a + b = a + c"
+ then have "- a + a + b = - a + a + c"
+ unfolding add_assoc by simp
+ then show "b = c" by simp
+qed
+
+lemma uminus_add_conv_diff[algebra_simps]:
+ "- a + b = b - a"
+by (simp add:diff_minus add_commute)
+
+lemma minus_add_distrib [simp]:
+ "- (a + b) = - a + - b"
+by (rule minus_unique) (simp add: add_ac)
+
+lemma minus_diff_eq [simp]:
+ "- (a - b) = b - a"
+by (simp add: diff_minus add_commute)
+
+lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c"
+by (simp add: diff_minus add_ac)
+
+lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b"
+by (simp add: diff_minus add_ac)
+
+lemma diff_eq_eq[algebra_simps]: "a - b = c \<longleftrightarrow> a = c + b"
+by (auto simp add: diff_minus add_assoc)
+
+lemma eq_diff_eq[algebra_simps]: "a = c - b \<longleftrightarrow> a + b = c"
+by (auto simp add: diff_minus add_assoc)
+
+lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)"
+by (simp add: diff_minus add_ac)
+
+lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b"
+by (simp add: diff_minus add_ac)
+
+lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
+by (simp add: algebra_simps)
+
+lemma diff_eq_0_iff_eq [simp, noatp]: "a - b = 0 \<longleftrightarrow> a = b"
+by (simp add: algebra_simps)
+
+end
+
+subsection {* (Partially) Ordered Groups *}
+
+class ordered_ab_semigroup_add = order + ab_semigroup_add +
+ assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
+begin
+
+lemma add_right_mono:
+ "a \<le> b \<Longrightarrow> a + c \<le> b + c"
+by (simp add: add_commute [of _ c] add_left_mono)
+
+text {* non-strict, in both arguments *}
+lemma add_mono:
+ "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d"
+ apply (erule add_right_mono [THEN order_trans])
+ apply (simp add: add_commute add_left_mono)
+ done
+
+end
+
+class ordered_cancel_ab_semigroup_add =
+ ordered_ab_semigroup_add + cancel_ab_semigroup_add
+begin
+
+lemma add_strict_left_mono:
+ "a < b \<Longrightarrow> c + a < c + b"
+by (auto simp add: less_le add_left_mono)
+
+lemma add_strict_right_mono:
+ "a < b \<Longrightarrow> a + c < b + c"
+by (simp add: add_commute [of _ c] add_strict_left_mono)
+
+text{*Strict monotonicity in both arguments*}
+lemma add_strict_mono:
+ "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
+apply (erule add_strict_right_mono [THEN less_trans])
+apply (erule add_strict_left_mono)
+done
+
+lemma add_less_le_mono:
+ "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c < b + d"
+apply (erule add_strict_right_mono [THEN less_le_trans])
+apply (erule add_left_mono)
+done
+
+lemma add_le_less_mono:
+ "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
+apply (erule add_right_mono [THEN le_less_trans])
+apply (erule add_strict_left_mono)
+done
+
+end
+
+class ordered_ab_semigroup_add_imp_le =
+ ordered_cancel_ab_semigroup_add +
+ assumes add_le_imp_le_left: "c + a \<le> c + b \<Longrightarrow> a \<le> b"
+begin
+
+lemma add_less_imp_less_left:
+ assumes less: "c + a < c + b" shows "a < b"
+proof -
+ from less have le: "c + a <= c + b" by (simp add: order_le_less)
+ have "a <= b"
+ apply (insert le)
+ apply (drule add_le_imp_le_left)
+ by (insert le, drule add_le_imp_le_left, assumption)
+ moreover have "a \<noteq> b"
+ proof (rule ccontr)
+ assume "~(a \<noteq> b)"
+ then have "a = b" by simp
+ then have "c + a = c + b" by simp
+ with less show "False"by simp
+ qed
+ ultimately show "a < b" by (simp add: order_le_less)
+qed
+
+lemma add_less_imp_less_right:
+ "a + c < b + c \<Longrightarrow> a < b"
+apply (rule add_less_imp_less_left [of c])
+apply (simp add: add_commute)
+done
+
+lemma add_less_cancel_left [simp]:
+ "c + a < c + b \<longleftrightarrow> a < b"
+by (blast intro: add_less_imp_less_left add_strict_left_mono)
+
+lemma add_less_cancel_right [simp]:
+ "a + c < b + c \<longleftrightarrow> a < b"
+by (blast intro: add_less_imp_less_right add_strict_right_mono)
+
+lemma add_le_cancel_left [simp]:
+ "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
+by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)
+
+lemma add_le_cancel_right [simp]:
+ "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
+by (simp add: add_commute [of a c] add_commute [of b c])
+
+lemma add_le_imp_le_right:
+ "a + c \<le> b + c \<Longrightarrow> a \<le> b"
+by simp
+
+lemma max_add_distrib_left:
+ "max x y + z = max (x + z) (y + z)"
+ unfolding max_def by auto
+
+lemma min_add_distrib_left:
+ "min x y + z = min (x + z) (y + z)"
+ unfolding min_def by auto
+
+end
+
+subsection {* Support for reasoning about signs *}
+
+class ordered_comm_monoid_add =
+ ordered_cancel_ab_semigroup_add + comm_monoid_add
+begin
+
+lemma add_pos_nonneg:
+ assumes "0 < a" and "0 \<le> b" shows "0 < a + b"
+proof -
+ have "0 + 0 < a + b"
+ using assms by (rule add_less_le_mono)
+ then show ?thesis by simp
+qed
+
+lemma add_pos_pos:
+ assumes "0 < a" and "0 < b" shows "0 < a + b"
+by (rule add_pos_nonneg) (insert assms, auto)
+
+lemma add_nonneg_pos:
+ assumes "0 \<le> a" and "0 < b" shows "0 < a + b"
+proof -
+ have "0 + 0 < a + b"
+ using assms by (rule add_le_less_mono)
+ then show ?thesis by simp
+qed
+
+lemma add_nonneg_nonneg:
+ assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"
+proof -
+ have "0 + 0 \<le> a + b"
+ using assms by (rule add_mono)
+ then show ?thesis by simp
+qed
+
+lemma add_neg_nonpos:
+ assumes "a < 0" and "b \<le> 0" shows "a + b < 0"
+proof -
+ have "a + b < 0 + 0"
+ using assms by (rule add_less_le_mono)
+ then show ?thesis by simp
+qed
+
+lemma add_neg_neg:
+ assumes "a < 0" and "b < 0" shows "a + b < 0"
+by (rule add_neg_nonpos) (insert assms, auto)
+
+lemma add_nonpos_neg:
+ assumes "a \<le> 0" and "b < 0" shows "a + b < 0"
+proof -
+ have "a + b < 0 + 0"
+ using assms by (rule add_le_less_mono)
+ then show ?thesis by simp
+qed
+
+lemma add_nonpos_nonpos:
+ assumes "a \<le> 0" and "b \<le> 0" shows "a + b \<le> 0"
+proof -
+ have "a + b \<le> 0 + 0"
+ using assms by (rule add_mono)
+ then show ?thesis by simp
+qed
+
+lemmas add_sign_intros =
+ add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg
+ add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos
+
+lemma add_nonneg_eq_0_iff:
+ assumes x: "0 \<le> x" and y: "0 \<le> y"
+ shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+proof (intro iffI conjI)
+ have "x = x + 0" by simp
+ also have "x + 0 \<le> x + y" using y by (rule add_left_mono)
+ also assume "x + y = 0"
+ also have "0 \<le> x" using x .
+ finally show "x = 0" .
+next
+ have "y = 0 + y" by simp
+ also have "0 + y \<le> x + y" using x by (rule add_right_mono)
+ also assume "x + y = 0"
+ also have "0 \<le> y" using y .
+ finally show "y = 0" .
+next
+ assume "x = 0 \<and> y = 0"
+ then show "x + y = 0" by simp
+qed
+
+end
+
+class ordered_ab_group_add =
+ ab_group_add + ordered_ab_semigroup_add
+begin
+
+subclass ordered_cancel_ab_semigroup_add ..
+
+subclass ordered_ab_semigroup_add_imp_le
+proof
+ fix a b c :: 'a
+ assume "c + a \<le> c + b"
+ hence "(-c) + (c + a) \<le> (-c) + (c + b)" by (rule add_left_mono)
+ hence "((-c) + c) + a \<le> ((-c) + c) + b" by (simp only: add_assoc)
+ thus "a \<le> b" by simp
+qed
+
+subclass ordered_comm_monoid_add ..
+
+lemma max_diff_distrib_left:
+ shows "max x y - z = max (x - z) (y - z)"
+by (simp add: diff_minus, rule max_add_distrib_left)
+
+lemma min_diff_distrib_left:
+ shows "min x y - z = min (x - z) (y - z)"
+by (simp add: diff_minus, rule min_add_distrib_left)
+
+lemma le_imp_neg_le:
+ assumes "a \<le> b" shows "-b \<le> -a"
+proof -
+ have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono)
+ hence "0 \<le> -a+b" by simp
+ hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)
+ thus ?thesis by (simp add: add_assoc)
+qed
+
+lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
+proof
+ assume "- b \<le> - a"
+ hence "- (- a) \<le> - (- b)" by (rule le_imp_neg_le)
+ thus "a\<le>b" by simp
+next
+ assume "a\<le>b"
+ thus "-b \<le> -a" by (rule le_imp_neg_le)
+qed
+
+lemma neg_le_0_iff_le [simp]: "- a \<le> 0 \<longleftrightarrow> 0 \<le> a"
+by (subst neg_le_iff_le [symmetric], simp)
+
+lemma neg_0_le_iff_le [simp]: "0 \<le> - a \<longleftrightarrow> a \<le> 0"
+by (subst neg_le_iff_le [symmetric], simp)
+
+lemma neg_less_iff_less [simp]: "- b < - a \<longleftrightarrow> a < b"
+by (force simp add: less_le)
+
+lemma neg_less_0_iff_less [simp]: "- a < 0 \<longleftrightarrow> 0 < a"
+by (subst neg_less_iff_less [symmetric], simp)
+
+lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0"
+by (subst neg_less_iff_less [symmetric], simp)
+
+text{*The next several equations can make the simplifier loop!*}
+
+lemma less_minus_iff: "a < - b \<longleftrightarrow> b < - a"
+proof -
+ have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)
+ thus ?thesis by simp
+qed
+
+lemma minus_less_iff: "- a < b \<longleftrightarrow> - b < a"
+proof -
+ have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)
+ thus ?thesis by simp
+qed
+
+lemma le_minus_iff: "a \<le> - b \<longleftrightarrow> b \<le> - a"
+proof -
+ have mm: "!! a (b::'a). (-(-a)) < -b \<Longrightarrow> -(-b) < -a" by (simp only: minus_less_iff)
+ have "(- (- a) <= -b) = (b <= - a)"
+ apply (auto simp only: le_less)
+ apply (drule mm)
+ apply (simp_all)
+ apply (drule mm[simplified], assumption)
+ done
+ then show ?thesis by simp
+qed
+
+lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
+by (auto simp add: le_less minus_less_iff)
+
+lemma less_iff_diff_less_0: "a < b \<longleftrightarrow> a - b < 0"
+proof -
+ have "(a < b) = (a + (- b) < b + (-b))"
+ by (simp only: add_less_cancel_right)
+ also have "... = (a - b < 0)" by (simp add: diff_minus)
+ finally show ?thesis .
+qed
+
+lemma diff_less_eq[algebra_simps]: "a - b < c \<longleftrightarrow> a < c + b"
+apply (subst less_iff_diff_less_0 [of a])
+apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
+apply (simp add: diff_minus add_ac)
+done
+
+lemma less_diff_eq[algebra_simps]: "a < c - b \<longleftrightarrow> a + b < c"
+apply (subst less_iff_diff_less_0 [of "plus a b"])
+apply (subst less_iff_diff_less_0 [of a])
+apply (simp add: diff_minus add_ac)
+done
+
+lemma diff_le_eq[algebra_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
+by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
+
+lemma le_diff_eq[algebra_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
+by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
+
+lemma le_iff_diff_le_0: "a \<le> b \<longleftrightarrow> a - b \<le> 0"
+by (simp add: algebra_simps)
+
+text{*Legacy - use @{text algebra_simps} *}
+lemmas group_simps[noatp] = algebra_simps
+
+end
+
+text{*Legacy - use @{text algebra_simps} *}
+lemmas group_simps[noatp] = algebra_simps
+
+class linordered_ab_semigroup_add =
+ linorder + ordered_ab_semigroup_add
+
+class linordered_cancel_ab_semigroup_add =
+ linorder + ordered_cancel_ab_semigroup_add
+begin
+
+subclass linordered_ab_semigroup_add ..
+
+subclass ordered_ab_semigroup_add_imp_le
+proof
+ fix a b c :: 'a
+ assume le: "c + a <= c + b"
+ show "a <= b"
+ proof (rule ccontr)
+ assume w: "~ a \<le> b"
+ hence "b <= a" by (simp add: linorder_not_le)
+ hence le2: "c + b <= c + a" by (rule add_left_mono)
+ have "a = b"
+ apply (insert le)
+ apply (insert le2)
+ apply (drule antisym, simp_all)
+ done
+ with w show False
+ by (simp add: linorder_not_le [symmetric])
+ qed
+qed
+
+end
+
+class linordered_ab_group_add = linorder + ordered_ab_group_add
+begin
+
+subclass linordered_cancel_ab_semigroup_add ..
+
+lemma neg_less_eq_nonneg [simp]:
+ "- a \<le> a \<longleftrightarrow> 0 \<le> a"
+proof
+ assume A: "- a \<le> a" show "0 \<le> a"
+ proof (rule classical)
+ assume "\<not> 0 \<le> a"
+ then have "a < 0" by auto
+ with A have "- a < 0" by (rule le_less_trans)
+ then show ?thesis by auto
+ qed
+next
+ assume A: "0 \<le> a" show "- a \<le> a"
+ proof (rule order_trans)
+ show "- a \<le> 0" using A by (simp add: minus_le_iff)
+ next
+ show "0 \<le> a" using A .
+ qed
+qed
+
+lemma neg_less_nonneg [simp]:
+ "- a < a \<longleftrightarrow> 0 < a"
+proof
+ assume A: "- a < a" show "0 < a"
+ proof (rule classical)
+ assume "\<not> 0 < a"
+ then have "a \<le> 0" by auto
+ with A have "- a < 0" by (rule less_le_trans)
+ then show ?thesis by auto
+ qed
+next
+ assume A: "0 < a" show "- a < a"
+ proof (rule less_trans)
+ show "- a < 0" using A by (simp add: minus_le_iff)
+ next
+ show "0 < a" using A .
+ qed
+qed
+
+lemma less_eq_neg_nonpos [simp]:
+ "a \<le> - a \<longleftrightarrow> a \<le> 0"
+proof
+ assume A: "a \<le> - a" show "a \<le> 0"
+ proof (rule classical)
+ assume "\<not> a \<le> 0"
+ then have "0 < a" by auto
+ then have "0 < - a" using A by (rule less_le_trans)
+ then show ?thesis by auto
+ qed
+next
+ assume A: "a \<le> 0" show "a \<le> - a"
+ proof (rule order_trans)
+ show "0 \<le> - a" using A by (simp add: minus_le_iff)
+ next
+ show "a \<le> 0" using A .
+ qed
+qed
+
+lemma equal_neg_zero [simp]:
+ "a = - a \<longleftrightarrow> a = 0"
+proof
+ assume "a = 0" then show "a = - a" by simp
+next
+ assume A: "a = - a" show "a = 0"
+ proof (cases "0 \<le> a")
+ case True with A have "0 \<le> - a" by auto
+ with le_minus_iff have "a \<le> 0" by simp
+ with True show ?thesis by (auto intro: order_trans)
+ next
+ case False then have B: "a \<le> 0" by auto
+ with A have "- a \<le> 0" by auto
+ with B show ?thesis by (auto intro: order_trans)
+ qed
+qed
+
+lemma neg_equal_zero [simp]:
+ "- a = a \<longleftrightarrow> a = 0"
+ by (auto dest: sym)
+
+lemma double_zero [simp]:
+ "a + a = 0 \<longleftrightarrow> a = 0"
+proof
+ assume assm: "a + a = 0"
+ then have a: "- a = a" by (rule minus_unique)
+ then show "a = 0" by (simp add: neg_equal_zero)
+qed simp
+
+lemma double_zero_sym [simp]:
+ "0 = a + a \<longleftrightarrow> a = 0"
+ by (rule, drule sym) simp_all
+
+lemma zero_less_double_add_iff_zero_less_single_add [simp]:
+ "0 < a + a \<longleftrightarrow> 0 < a"
+proof
+ assume "0 < a + a"
+ then have "0 - a < a" by (simp only: diff_less_eq)
+ then have "- a < a" by simp
+ then show "0 < a" by (simp add: neg_less_nonneg)
+next
+ assume "0 < a"
+ with this have "0 + 0 < a + a"
+ by (rule add_strict_mono)
+ then show "0 < a + a" by simp
+qed
+
+lemma zero_le_double_add_iff_zero_le_single_add [simp]:
+ "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
+ by (auto simp add: le_less)
+
+lemma double_add_less_zero_iff_single_add_less_zero [simp]:
+ "a + a < 0 \<longleftrightarrow> a < 0"
+proof -
+ have "\<not> a + a < 0 \<longleftrightarrow> \<not> a < 0"
+ by (simp add: not_less)
+ then show ?thesis by simp
+qed
+
+lemma double_add_le_zero_iff_single_add_le_zero [simp]:
+ "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
+proof -
+ have "\<not> a + a \<le> 0 \<longleftrightarrow> \<not> a \<le> 0"
+ by (simp add: not_le)
+ then show ?thesis by simp
+qed
+
+lemma le_minus_self_iff:
+ "a \<le> - a \<longleftrightarrow> a \<le> 0"
+proof -
+ from add_le_cancel_left [of "- a" "a + a" 0]
+ have "a \<le> - a \<longleftrightarrow> a + a \<le> 0"
+ by (simp add: add_assoc [symmetric])
+ thus ?thesis by simp
+qed
+
+lemma minus_le_self_iff:
+ "- a \<le> a \<longleftrightarrow> 0 \<le> a"
+proof -
+ from add_le_cancel_left [of "- a" 0 "a + a"]
+ have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"
+ by (simp add: add_assoc [symmetric])
+ thus ?thesis by simp
+qed
+
+lemma minus_max_eq_min:
+ "- max x y = min (-x) (-y)"
+ by (auto simp add: max_def min_def)
+
+lemma minus_min_eq_max:
+ "- min x y = max (-x) (-y)"
+ by (auto simp add: max_def min_def)
+
+end
+
+-- {* FIXME localize the following *}
+
+lemma add_increasing:
+ fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
+ shows "[|0\<le>a; b\<le>c|] ==> b \<le> a + c"
+by (insert add_mono [of 0 a b c], simp)
+
+lemma add_increasing2:
+ fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
+ shows "[|0\<le>c; b\<le>a|] ==> b \<le> a + c"
+by (simp add:add_increasing add_commute[of a])
+
+lemma add_strict_increasing:
+ fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
+ shows "[|0<a; b\<le>c|] ==> b < a + c"
+by (insert add_less_le_mono [of 0 a b c], simp)
+
+lemma add_strict_increasing2:
+ fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
+ shows "[|0\<le>a; b<c|] ==> b < a + c"
+by (insert add_le_less_mono [of 0 a b c], simp)
+
+
+class ordered_ab_group_add_abs = ordered_ab_group_add + abs +
+ assumes abs_ge_zero [simp]: "\<bar>a\<bar> \<ge> 0"
+ and abs_ge_self: "a \<le> \<bar>a\<bar>"
+ and abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"
+ and abs_minus_cancel [simp]: "\<bar>-a\<bar> = \<bar>a\<bar>"
+ and abs_triangle_ineq: "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
+begin
+
+lemma abs_minus_le_zero: "- \<bar>a\<bar> \<le> 0"
+ unfolding neg_le_0_iff_le by simp
+
+lemma abs_of_nonneg [simp]:
+ assumes nonneg: "0 \<le> a" shows "\<bar>a\<bar> = a"
+proof (rule antisym)
+ from nonneg le_imp_neg_le have "- a \<le> 0" by simp
+ from this nonneg have "- a \<le> a" by (rule order_trans)
+ then show "\<bar>a\<bar> \<le> a" by (auto intro: abs_leI)
+qed (rule abs_ge_self)
+
+lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"
+by (rule antisym)
+ (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"])
+
+lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"
+proof -
+ have "\<bar>a\<bar> = 0 \<Longrightarrow> a = 0"
+ proof (rule antisym)
+ assume zero: "\<bar>a\<bar> = 0"
+ with abs_ge_self show "a \<le> 0" by auto
+ from zero have "\<bar>-a\<bar> = 0" by simp
+ with abs_ge_self [of "uminus a"] have "- a \<le> 0" by auto
+ with neg_le_0_iff_le show "0 \<le> a" by auto
+ qed
+ then show ?thesis by auto
+qed
+
+lemma abs_zero [simp]: "\<bar>0\<bar> = 0"
+by simp
+
+lemma abs_0_eq [simp, noatp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
+proof -
+ have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac)
+ thus ?thesis by simp
+qed
+
+lemma abs_le_zero_iff [simp]: "\<bar>a\<bar> \<le> 0 \<longleftrightarrow> a = 0"
+proof
+ assume "\<bar>a\<bar> \<le> 0"
+ then have "\<bar>a\<bar> = 0" by (rule antisym) simp
+ thus "a = 0" by simp
+next
+ assume "a = 0"
+ thus "\<bar>a\<bar> \<le> 0" by simp
+qed
+
+lemma zero_less_abs_iff [simp]: "0 < \<bar>a\<bar> \<longleftrightarrow> a \<noteq> 0"
+by (simp add: less_le)
+
+lemma abs_not_less_zero [simp]: "\<not> \<bar>a\<bar> < 0"
+proof -
+ have a: "\<And>x y. x \<le> y \<Longrightarrow> \<not> y < x" by auto
+ show ?thesis by (simp add: a)
+qed
+
+lemma abs_ge_minus_self: "- a \<le> \<bar>a\<bar>"
+proof -
+ have "- a \<le> \<bar>-a\<bar>" by (rule abs_ge_self)
+ then show ?thesis by simp
+qed
+
+lemma abs_minus_commute:
+ "\<bar>a - b\<bar> = \<bar>b - a\<bar>"
+proof -
+ have "\<bar>a - b\<bar> = \<bar>- (a - b)\<bar>" by (simp only: abs_minus_cancel)
+ also have "... = \<bar>b - a\<bar>" by simp
+ finally show ?thesis .
+qed
+
+lemma abs_of_pos: "0 < a \<Longrightarrow> \<bar>a\<bar> = a"
+by (rule abs_of_nonneg, rule less_imp_le)
+
+lemma abs_of_nonpos [simp]:
+ assumes "a \<le> 0" shows "\<bar>a\<bar> = - a"
+proof -
+ let ?b = "- a"
+ have "- ?b \<le> 0 \<Longrightarrow> \<bar>- ?b\<bar> = - (- ?b)"
+ unfolding abs_minus_cancel [of "?b"]
+ unfolding neg_le_0_iff_le [of "?b"]
+ unfolding minus_minus by (erule abs_of_nonneg)
+ then show ?thesis using assms by auto
+qed
+
+lemma abs_of_neg: "a < 0 \<Longrightarrow> \<bar>a\<bar> = - a"
+by (rule abs_of_nonpos, rule less_imp_le)
+
+lemma abs_le_D1: "\<bar>a\<bar> \<le> b \<Longrightarrow> a \<le> b"
+by (insert abs_ge_self, blast intro: order_trans)
+
+lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b"
+by (insert abs_le_D1 [of "uminus a"], simp)
+
+lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b"
+by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
+
+lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
+ apply (simp add: algebra_simps)
+ apply (subgoal_tac "abs a = abs (plus b (minus a b))")
+ apply (erule ssubst)
+ apply (rule abs_triangle_ineq)
+ apply (rule arg_cong[of _ _ abs])
+ apply (simp add: algebra_simps)
+done
+
+lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>"
+ apply (subst abs_le_iff)
+ apply auto
+ apply (rule abs_triangle_ineq2)
+ apply (subst abs_minus_commute)
+ apply (rule abs_triangle_ineq2)
+done
+
+lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
+proof -
+ have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl)
+ also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq)
+ finally show ?thesis by simp
+qed
+
+lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
+proof -
+ have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
+ also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
+ finally show ?thesis .
+qed
+
+lemma abs_add_abs [simp]:
+ "\<bar>\<bar>a\<bar> + \<bar>b\<bar>\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" (is "?L = ?R")
+proof (rule antisym)
+ show "?L \<ge> ?R" by(rule abs_ge_self)
+next
+ have "?L \<le> \<bar>\<bar>a\<bar>\<bar> + \<bar>\<bar>b\<bar>\<bar>" by(rule abs_triangle_ineq)
+ also have "\<dots> = ?R" by simp
+ finally show "?L \<le> ?R" .
+qed
+
+end
+
+text {* Needed for abelian cancellation simprocs: *}
+
+lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)"
+apply (subst add_left_commute)
+apply (subst add_left_cancel)
+apply simp
+done
+
+lemma add_cancel_end: "(x + (y + z) = y) = (x = - (z::'a::ab_group_add))"
+apply (subst add_cancel_21[of _ _ _ 0, simplified])
+apply (simp add: add_right_cancel[symmetric, of "x" "-z" "z", simplified])
+done
+
+lemma less_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \<Longrightarrow> (x < y) = (x' < y')"
+by (simp add: less_iff_diff_less_0[of x y] less_iff_diff_less_0[of x' y'])
+
+lemma le_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \<Longrightarrow> (y <= x) = (y' <= x')"
+apply (simp add: le_iff_diff_le_0[of y x] le_iff_diff_le_0[of y' x'])
+apply (simp add: neg_le_iff_le[symmetric, of "y-x" 0] neg_le_iff_le[symmetric, of "y'-x'" 0])
+done
+
+lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \<Longrightarrow> (x = y) = (x' = y')"
+by (simp only: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y'])
+
+lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)"
+by (simp add: diff_minus)
+
+lemma le_add_right_mono:
+ assumes
+ "a <= b + (c::'a::ordered_ab_group_add)"
+ "c <= d"
+ shows "a <= b + d"
+ apply (rule_tac order_trans[where y = "b+c"])
+ apply (simp_all add: prems)
+ done
+
+
+subsection {* Tools setup *}
+
+lemma add_mono_thms_linordered_semiring [noatp]:
+ fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
+ shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
+ and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
+ and "i \<le> j \<and> k = l \<Longrightarrow> i + k \<le> j + l"
+ and "i = j \<and> k = l \<Longrightarrow> i + k = j + l"
+by (rule add_mono, clarify+)+
+
+lemma add_mono_thms_linordered_field [noatp]:
+ fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"
+ shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"
+ and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"
+ and "i < j \<and> k \<le> l \<Longrightarrow> i + k < j + l"
+ and "i \<le> j \<and> k < l \<Longrightarrow> i + k < j + l"
+ and "i < j \<and> k < l \<Longrightarrow> i + k < j + l"
+by (auto intro: add_strict_right_mono add_strict_left_mono
+ add_less_le_mono add_le_less_mono add_strict_mono)
+
+text{*Simplification of @{term "x-y < 0"}, etc.*}
+lemmas diff_less_0_iff_less [simp, noatp] = less_iff_diff_less_0 [symmetric]
+lemmas diff_le_0_iff_le [simp, noatp] = le_iff_diff_le_0 [symmetric]
+
+ML {*
+structure ab_group_add_cancel = Abel_Cancel
+(
+
+(* term order for abelian groups *)
+
+fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a')
+ [@{const_name Algebras.zero}, @{const_name Algebras.plus},
+ @{const_name Algebras.uminus}, @{const_name Algebras.minus}]
+ | agrp_ord _ = ~1;
+
+fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
+
+local
+ val ac1 = mk_meta_eq @{thm add_assoc};
+ val ac2 = mk_meta_eq @{thm add_commute};
+ val ac3 = mk_meta_eq @{thm add_left_commute};
+ fun solve_add_ac thy _ (_ $ (Const (@{const_name Algebras.plus},_) $ _ $ _) $ _) =
+ SOME ac1
+ | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Algebras.plus},_) $ y $ z)) =
+ if termless_agrp (y, x) then SOME ac3 else NONE
+ | solve_add_ac thy _ (_ $ x $ y) =
+ if termless_agrp (y, x) then SOME ac2 else NONE
+ | solve_add_ac thy _ _ = NONE
+in
+ val add_ac_proc = Simplifier.simproc @{theory}
+ "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac;
+end;
+
+val eq_reflection = @{thm eq_reflection};
+
+val T = @{typ "'a::ab_group_add"};
+
+val cancel_ss = HOL_basic_ss settermless termless_agrp
+ addsimprocs [add_ac_proc] addsimps
+ [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_def},
+ @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero},
+ @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel},
+ @{thm minus_add_cancel}];
+
+val sum_pats = [@{cterm "x + y::'a::ab_group_add"}, @{cterm "x - y::'a::ab_group_add"}];
+
+val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}];
+
+val dest_eqI =
+ fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of;
+
+);
+*}
+
+ML {*
+ Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv];
+*}
+
+code_modulename SML
+ Groups Arith
+
+code_modulename OCaml
+ Groups Arith
+
+code_modulename Haskell
+ Groups Arith
+
+end
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Mon Feb 08 17:13:45 2010 +0100
@@ -629,6 +629,8 @@
return a
done"
+code_reserved SML upto
+
ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
export_code qsort in SML_imp module_name QSort
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Feb 08 17:13:45 2010 +0100
@@ -986,6 +986,8 @@
return zs
done)"
+code_reserved SML upto
+
ML {* @{code test_1} () *}
ML {* @{code test_2} () *}
ML {* @{code test_3} () *}
--- a/src/HOL/Import/HOL/arithmetic.imp Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Import/HOL/arithmetic.imp Mon Feb 08 17:13:45 2010 +0100
@@ -162,12 +162,12 @@
"LESS_OR" > "Nat.Suc_leI"
"LESS_NOT_SUC" > "HOL4Base.arithmetic.LESS_NOT_SUC"
"LESS_MULT_MONO" > "Nat.Suc_mult_less_cancel1"
- "LESS_MULT2" > "Ring_and_Field.mult_pos_pos"
+ "LESS_MULT2" > "Rings.mult_pos_pos"
"LESS_MONO_REV" > "Nat.Suc_less_SucD"
"LESS_MONO_MULT" > "Nat.mult_le_mono1"
"LESS_MONO_EQ" > "Nat.Suc_less_eq"
- "LESS_MONO_ADD_INV" > "OrderedGroup.add_less_imp_less_right"
- "LESS_MONO_ADD_EQ" > "OrderedGroup.add_less_cancel_right"
+ "LESS_MONO_ADD_INV" > "Groups.add_less_imp_less_right"
+ "LESS_MONO_ADD_EQ" > "Groups.add_less_cancel_right"
"LESS_MONO_ADD" > "Nat.add_less_mono1"
"LESS_MOD" > "Divides.mod_less"
"LESS_LESS_SUC" > "HOL4Base.arithmetic.LESS_LESS_SUC"
@@ -180,7 +180,7 @@
"LESS_EQ_SUC_REFL" > "HOL4Base.arithmetic.LESS_EQ_SUC_REFL"
"LESS_EQ_SUB_LESS" > "HOL4Base.arithmetic.LESS_EQ_SUB_LESS"
"LESS_EQ_REFL" > "Finite_Set.max.f_below.below_refl"
- "LESS_EQ_MONO_ADD_EQ" > "OrderedGroup.add_le_cancel_right"
+ "LESS_EQ_MONO_ADD_EQ" > "Groups.add_le_cancel_right"
"LESS_EQ_MONO" > "Nat.Suc_le_mono"
"LESS_EQ_LESS_TRANS" > "Nat.le_less_trans"
"LESS_EQ_LESS_EQ_MONO" > "Nat.add_le_mono"
--- a/src/HOL/Import/HOL/divides.imp Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Import/HOL/divides.imp Mon Feb 08 17:13:45 2010 +0100
@@ -9,16 +9,16 @@
"divides_def" > "HOL4Compat.divides_def"
"ONE_DIVIDES_ALL" > "HOL4Base.divides.ONE_DIVIDES_ALL"
"NOT_LT_DIV" > "NatSimprocs.nat_dvd_not_less"
- "DIVIDES_TRANS" > "Ring_and_Field.dvd_trans"
- "DIVIDES_SUB" > "Ring_and_Field.dvd_diff"
- "DIVIDES_REFL" > "Ring_and_Field.dvd_refl"
+ "DIVIDES_TRANS" > "Rings.dvd_trans"
+ "DIVIDES_SUB" > "Rings.dvd_diff"
+ "DIVIDES_REFL" > "Rings.dvd_refl"
"DIVIDES_MULT_LEFT" > "HOL4Base.divides.DIVIDES_MULT_LEFT"
"DIVIDES_MULT" > "Divides.dvd_mult2"
"DIVIDES_LE" > "Divides.dvd_imp_le"
"DIVIDES_FACT" > "HOL4Base.divides.DIVIDES_FACT"
"DIVIDES_ANTISYM" > "Divides.dvd_antisym"
"DIVIDES_ADD_2" > "HOL4Base.divides.DIVIDES_ADD_2"
- "DIVIDES_ADD_1" > "Ring_and_Field.dvd_add"
+ "DIVIDES_ADD_1" > "Rings.dvd_add"
"ALL_DIVIDES_0" > "Divides.dvd_0_right"
end
--- a/src/HOL/Import/HOL/prob_extra.imp Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Import/HOL/prob_extra.imp Mon Feb 08 17:13:45 2010 +0100
@@ -23,9 +23,9 @@
"REAL_SUP_LE_X" > "HOL4Prob.prob_extra.REAL_SUP_LE_X"
"REAL_SUP_EXISTS_UNIQUE" > "HOL4Prob.prob_extra.REAL_SUP_EXISTS_UNIQUE"
"REAL_POW" > "RealPow.realpow_real_of_nat"
- "REAL_LE_INV_LE" > "Ring_and_Field.le_imp_inverse_le"
+ "REAL_LE_INV_LE" > "Rings.le_imp_inverse_le"
"REAL_LE_EQ" > "Set.basic_trans_rules_26"
- "REAL_INVINV_ALL" > "Ring_and_Field.inverse_inverse_eq"
+ "REAL_INVINV_ALL" > "Rings.inverse_inverse_eq"
"REAL_INF_MIN" > "HOL4Prob.prob_extra.REAL_INF_MIN"
"RAND_THM" > "HOL.arg_cong"
"POW_HALF_TWICE" > "HOL4Prob.prob_extra.POW_HALF_TWICE"
--- a/src/HOL/Import/HOL/real.imp Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Import/HOL/real.imp Mon Feb 08 17:13:45 2010 +0100
@@ -25,13 +25,13 @@
"sumc" > "HOL4Real.real.sumc"
"sum_def" > "HOL4Real.real.sum_def"
"sum" > "HOL4Real.real.sum"
- "real_sub" > "OrderedGroup.diff_def"
+ "real_sub" > "Groups.diff_def"
"real_of_num" > "HOL4Compat.real_of_num"
"real_lte" > "HOL4Compat.real_lte"
"real_lt" > "Orderings.linorder_not_le"
"real_gt" > "HOL4Compat.real_gt"
"real_ge" > "HOL4Compat.real_ge"
- "real_div" > "Ring_and_Field.field_class.divide_inverse"
+ "real_div" > "Rings.field_class.divide_inverse"
"pow" > "HOL4Compat.pow"
"abs" > "HOL4Compat.abs"
"SUP_LEMMA3" > "HOL4Real.real.SUP_LEMMA3"
@@ -74,24 +74,24 @@
"REAL_SUB_TRIANGLE" > "HOL4Real.real.REAL_SUB_TRIANGLE"
"REAL_SUB_SUB2" > "HOL4Real.real.REAL_SUB_SUB2"
"REAL_SUB_SUB" > "HOL4Real.real.REAL_SUB_SUB"
- "REAL_SUB_RZERO" > "OrderedGroup.diff_0_right"
- "REAL_SUB_RNEG" > "OrderedGroup.diff_minus_eq_add"
- "REAL_SUB_REFL" > "OrderedGroup.diff_self"
- "REAL_SUB_RDISTRIB" > "Ring_and_Field.ring_eq_simps_3"
+ "REAL_SUB_RZERO" > "Groups.diff_0_right"
+ "REAL_SUB_RNEG" > "Groups.diff_minus_eq_add"
+ "REAL_SUB_REFL" > "Groups.diff_self"
+ "REAL_SUB_RDISTRIB" > "Rings.ring_eq_simps_3"
"REAL_SUB_NEG2" > "HOL4Real.real.REAL_SUB_NEG2"
- "REAL_SUB_LZERO" > "OrderedGroup.diff_0"
+ "REAL_SUB_LZERO" > "Groups.diff_0"
"REAL_SUB_LT" > "HOL4Real.real.REAL_SUB_LT"
"REAL_SUB_LNEG" > "HOL4Real.real.REAL_SUB_LNEG"
"REAL_SUB_LE" > "HOL4Real.real.REAL_SUB_LE"
- "REAL_SUB_LDISTRIB" > "Ring_and_Field.ring_eq_simps_4"
+ "REAL_SUB_LDISTRIB" > "Rings.ring_eq_simps_4"
"REAL_SUB_INV2" > "HOL4Real.real.REAL_SUB_INV2"
"REAL_SUB_ADD2" > "HOL4Real.real.REAL_SUB_ADD2"
- "REAL_SUB_ADD" > "OrderedGroup.diff_add_cancel"
- "REAL_SUB_ABS" > "OrderedGroup.abs_triangle_ineq2"
- "REAL_SUB_0" > "OrderedGroup.diff_eq_0_iff_eq"
+ "REAL_SUB_ADD" > "Groups.diff_add_cancel"
+ "REAL_SUB_ABS" > "Groups.abs_triangle_ineq2"
+ "REAL_SUB_0" > "Groups.diff_eq_0_iff_eq"
"REAL_RNEG_UNIQ" > "RealDef.real_add_eq_0_iff"
- "REAL_RINV_UNIQ" > "Ring_and_Field.inverse_unique"
- "REAL_RDISTRIB" > "Ring_and_Field.ring_eq_simps_1"
+ "REAL_RINV_UNIQ" > "Rings.inverse_unique"
+ "REAL_RDISTRIB" > "Rings.ring_eq_simps_1"
"REAL_POW_POW" > "Power.power_mult"
"REAL_POW_MONO_LT" > "HOL4Real.real.REAL_POW_MONO_LT"
"REAL_POW_LT2" > "HOL4Real.real.REAL_POW_LT2"
@@ -103,7 +103,7 @@
"REAL_POS_NZ" > "HOL4Real.real.REAL_POS_NZ"
"REAL_POS" > "RealDef.real_of_nat_ge_zero"
"REAL_POASQ" > "HOL4Real.real.REAL_POASQ"
- "REAL_OVER1" > "Ring_and_Field.divide_1"
+ "REAL_OVER1" > "Rings.divide_1"
"REAL_OF_NUM_SUC" > "RealDef.real_of_nat_Suc"
"REAL_OF_NUM_POW" > "RealPow.realpow_real_of_nat"
"REAL_OF_NUM_MUL" > "RealDef.real_of_nat_mult"
@@ -113,172 +113,172 @@
"REAL_NZ_IMP_LT" > "HOL4Real.real.REAL_NZ_IMP_LT"
"REAL_NOT_LT" > "HOL4Compat.real_lte"
"REAL_NOT_LE" > "Orderings.linorder_not_le"
- "REAL_NEG_SUB" > "OrderedGroup.minus_diff_eq"
- "REAL_NEG_RMUL" > "Ring_and_Field.mult_minus_right"
- "REAL_NEG_NEG" > "OrderedGroup.minus_minus"
- "REAL_NEG_MUL2" > "Ring_and_Field.minus_mult_minus"
+ "REAL_NEG_SUB" > "Groups.minus_diff_eq"
+ "REAL_NEG_RMUL" > "Rings.mult_minus_right"
+ "REAL_NEG_NEG" > "Groups.minus_minus"
+ "REAL_NEG_MUL2" > "Rings.minus_mult_minus"
"REAL_NEG_MINUS1" > "HOL4Real.real.REAL_NEG_MINUS1"
- "REAL_NEG_LT0" > "OrderedGroup.neg_less_0_iff_less"
- "REAL_NEG_LMUL" > "Ring_and_Field.mult_minus_left"
- "REAL_NEG_LE0" > "OrderedGroup.neg_le_0_iff_le"
- "REAL_NEG_INV" > "Ring_and_Field.nonzero_inverse_minus_eq"
- "REAL_NEG_GT0" > "OrderedGroup.neg_0_less_iff_less"
- "REAL_NEG_GE0" > "OrderedGroup.neg_0_le_iff_le"
- "REAL_NEG_EQ0" > "OrderedGroup.neg_equal_0_iff_equal"
+ "REAL_NEG_LT0" > "Groups.neg_less_0_iff_less"
+ "REAL_NEG_LMUL" > "Rings.mult_minus_left"
+ "REAL_NEG_LE0" > "Groups.neg_le_0_iff_le"
+ "REAL_NEG_INV" > "Rings.nonzero_inverse_minus_eq"
+ "REAL_NEG_GT0" > "Groups.neg_0_less_iff_less"
+ "REAL_NEG_GE0" > "Groups.neg_0_le_iff_le"
+ "REAL_NEG_EQ0" > "Groups.neg_equal_0_iff_equal"
"REAL_NEG_EQ" > "HOL4Real.real.REAL_NEG_EQ"
- "REAL_NEG_ADD" > "OrderedGroup.minus_add_distrib"
- "REAL_NEG_0" > "OrderedGroup.minus_zero"
- "REAL_NEGNEG" > "OrderedGroup.minus_minus"
+ "REAL_NEG_ADD" > "Groups.minus_add_distrib"
+ "REAL_NEG_0" > "Groups.minus_zero"
+ "REAL_NEGNEG" > "Groups.minus_minus"
"REAL_MUL_SYM" > "Int.zmult_ac_2"
- "REAL_MUL_RZERO" > "Ring_and_Field.mult_zero_right"
- "REAL_MUL_RNEG" > "Ring_and_Field.mult_minus_right"
- "REAL_MUL_RINV" > "Ring_and_Field.right_inverse"
+ "REAL_MUL_RZERO" > "Rings.mult_zero_right"
+ "REAL_MUL_RNEG" > "Rings.mult_minus_right"
+ "REAL_MUL_RINV" > "Rings.right_inverse"
"REAL_MUL_RID" > "Finite_Set.AC_mult.f_e.ident"
- "REAL_MUL_LZERO" > "Ring_and_Field.mult_zero_left"
- "REAL_MUL_LNEG" > "Ring_and_Field.mult_minus_left"
+ "REAL_MUL_LZERO" > "Rings.mult_zero_left"
+ "REAL_MUL_LNEG" > "Rings.mult_minus_left"
"REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV"
"REAL_MUL_LID" > "Finite_Set.AC_mult.f_e.left_ident"
"REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC"
"REAL_MUL" > "RealDef.real_of_nat_mult"
"REAL_MIDDLE2" > "HOL4Real.real.REAL_MIDDLE2"
"REAL_MIDDLE1" > "HOL4Real.real.REAL_MIDDLE1"
- "REAL_MEAN" > "Ring_and_Field.dense"
+ "REAL_MEAN" > "Rings.dense"
"REAL_LT_TRANS" > "Set.basic_trans_rules_21"
"REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
- "REAL_LT_SUB_RADD" > "OrderedGroup.compare_rls_6"
- "REAL_LT_SUB_LADD" > "OrderedGroup.compare_rls_7"
- "REAL_LT_RMUL_IMP" > "Ring_and_Field.mult_strict_right_mono"
+ "REAL_LT_SUB_RADD" > "Groups.compare_rls_6"
+ "REAL_LT_SUB_LADD" > "Groups.compare_rls_7"
+ "REAL_LT_RMUL_IMP" > "Rings.mult_strict_right_mono"
"REAL_LT_RMUL_0" > "HOL4Real.real.REAL_LT_RMUL_0"
"REAL_LT_RMUL" > "RealDef.real_mult_less_iff1"
"REAL_LT_REFL" > "Orderings.order_less_irrefl"
- "REAL_LT_RDIV_EQ" > "Ring_and_Field.pos_less_divide_eq"
+ "REAL_LT_RDIV_EQ" > "Rings.pos_less_divide_eq"
"REAL_LT_RDIV_0" > "HOL4Real.real.REAL_LT_RDIV_0"
"REAL_LT_RDIV" > "HOL4Real.real.REAL_LT_RDIV"
- "REAL_LT_RADD" > "OrderedGroup.add_less_cancel_right"
+ "REAL_LT_RADD" > "Groups.add_less_cancel_right"
"REAL_LT_NZ" > "HOL4Real.real.REAL_LT_NZ"
"REAL_LT_NEGTOTAL" > "HOL4Real.real.REAL_LT_NEGTOTAL"
- "REAL_LT_NEG" > "OrderedGroup.neg_less_iff_less"
+ "REAL_LT_NEG" > "Groups.neg_less_iff_less"
"REAL_LT_MULTIPLE" > "HOL4Real.real.REAL_LT_MULTIPLE"
- "REAL_LT_MUL2" > "Ring_and_Field.mult_strict_mono'"
- "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos"
- "REAL_LT_LMUL_IMP" > "Ring_and_Field.linordered_comm_semiring_strict_class.mult_strict_mono"
+ "REAL_LT_MUL2" > "Rings.mult_strict_mono'"
+ "REAL_LT_MUL" > "Rings.mult_pos_pos"
+ "REAL_LT_LMUL_IMP" > "Rings.linordered_comm_semiring_strict_class.mult_strict_mono"
"REAL_LT_LMUL_0" > "HOL4Real.real.REAL_LT_LMUL_0"
"REAL_LT_LMUL" > "HOL4Real.real.REAL_LT_LMUL"
"REAL_LT_LE" > "Orderings.order_class.order_less_le"
- "REAL_LT_LDIV_EQ" > "Ring_and_Field.pos_divide_less_eq"
- "REAL_LT_LADD" > "OrderedGroup.add_less_cancel_left"
- "REAL_LT_INV_EQ" > "Ring_and_Field.inverse_positive_iff_positive"
- "REAL_LT_INV" > "Ring_and_Field.less_imp_inverse_less"
+ "REAL_LT_LDIV_EQ" > "Rings.pos_divide_less_eq"
+ "REAL_LT_LADD" > "Groups.add_less_cancel_left"
+ "REAL_LT_INV_EQ" > "Rings.inverse_positive_iff_positive"
+ "REAL_LT_INV" > "Rings.less_imp_inverse_less"
"REAL_LT_IMP_NE" > "Orderings.less_imp_neq"
"REAL_LT_IMP_LE" > "Orderings.order_less_imp_le"
- "REAL_LT_IADD" > "OrderedGroup.add_strict_left_mono"
+ "REAL_LT_IADD" > "Groups.add_strict_left_mono"
"REAL_LT_HALF2" > "HOL4Real.real.REAL_LT_HALF2"
"REAL_LT_HALF1" > "NatSimprocs.half_gt_zero_iff"
"REAL_LT_GT" > "Orderings.order_less_not_sym"
"REAL_LT_FRACTION_0" > "HOL4Real.real.REAL_LT_FRACTION_0"
"REAL_LT_FRACTION" > "HOL4Real.real.REAL_LT_FRACTION"
- "REAL_LT_DIV" > "Ring_and_Field.divide_pos_pos"
+ "REAL_LT_DIV" > "Rings.divide_pos_pos"
"REAL_LT_ANTISYM" > "HOL4Real.real.REAL_LT_ANTISYM"
- "REAL_LT_ADD_SUB" > "OrderedGroup.compare_rls_7"
+ "REAL_LT_ADD_SUB" > "Groups.compare_rls_7"
"REAL_LT_ADDR" > "HOL4Real.real.REAL_LT_ADDR"
"REAL_LT_ADDNEG2" > "HOL4Real.real.REAL_LT_ADDNEG2"
"REAL_LT_ADDNEG" > "HOL4Real.real.REAL_LT_ADDNEG"
"REAL_LT_ADDL" > "HOL4Real.real.REAL_LT_ADDL"
- "REAL_LT_ADD2" > "OrderedGroup.add_strict_mono"
+ "REAL_LT_ADD2" > "Groups.add_strict_mono"
"REAL_LT_ADD1" > "HOL4Real.real.REAL_LT_ADD1"
- "REAL_LT_ADD" > "OrderedGroup.add_pos_pos"
+ "REAL_LT_ADD" > "Groups.add_pos_pos"
"REAL_LT_1" > "HOL4Real.real.REAL_LT_1"
- "REAL_LT_01" > "Ring_and_Field.zero_less_one"
+ "REAL_LT_01" > "Rings.zero_less_one"
"REAL_LTE_TRANS" > "Set.basic_trans_rules_24"
"REAL_LTE_TOTAL" > "HOL4Real.real.REAL_LTE_TOTAL"
"REAL_LTE_ANTSYM" > "HOL4Real.real.REAL_LTE_ANTSYM"
- "REAL_LTE_ADD2" > "OrderedGroup.add_less_le_mono"
- "REAL_LTE_ADD" > "OrderedGroup.add_pos_nonneg"
+ "REAL_LTE_ADD2" > "Groups.add_less_le_mono"
+ "REAL_LTE_ADD" > "Groups.add_pos_nonneg"
"REAL_LT1_POW2" > "HOL4Real.real.REAL_LT1_POW2"
"REAL_LT" > "RealDef.real_of_nat_less_iff"
"REAL_LNEG_UNIQ" > "HOL4Real.real.REAL_LNEG_UNIQ"
"REAL_LINV_UNIQ" > "HOL4Real.real.REAL_LINV_UNIQ"
"REAL_LE_TRANS" > "Set.basic_trans_rules_25"
"REAL_LE_TOTAL" > "Orderings.linorder_class.linorder_linear"
- "REAL_LE_SUB_RADD" > "OrderedGroup.compare_rls_8"
- "REAL_LE_SUB_LADD" > "OrderedGroup.compare_rls_9"
- "REAL_LE_SQUARE" > "Ring_and_Field.zero_le_square"
- "REAL_LE_RNEG" > "OrderedGroup.le_eq_neg"
- "REAL_LE_RMUL_IMP" > "Ring_and_Field.mult_right_mono"
+ "REAL_LE_SUB_RADD" > "Groups.compare_rls_8"
+ "REAL_LE_SUB_LADD" > "Groups.compare_rls_9"
+ "REAL_LE_SQUARE" > "Rings.zero_le_square"
+ "REAL_LE_RNEG" > "Groups.le_eq_neg"
+ "REAL_LE_RMUL_IMP" > "Rings.mult_right_mono"
"REAL_LE_RMUL" > "RealDef.real_mult_le_cancel_iff1"
"REAL_LE_REFL" > "Finite_Set.max.f_below.below_refl"
- "REAL_LE_RDIV_EQ" > "Ring_and_Field.pos_le_divide_eq"
- "REAL_LE_RDIV" > "Ring_and_Field.mult_imp_le_div_pos"
- "REAL_LE_RADD" > "OrderedGroup.add_le_cancel_right"
+ "REAL_LE_RDIV_EQ" > "Rings.pos_le_divide_eq"
+ "REAL_LE_RDIV" > "Rings.mult_imp_le_div_pos"
+ "REAL_LE_RADD" > "Groups.add_le_cancel_right"
"REAL_LE_POW2" > "Nat_Numeral.zero_compare_simps_12"
"REAL_LE_NEGTOTAL" > "HOL4Real.real.REAL_LE_NEGTOTAL"
- "REAL_LE_NEGR" > "OrderedGroup.le_minus_self_iff"
- "REAL_LE_NEGL" > "OrderedGroup.minus_le_self_iff"
- "REAL_LE_NEG2" > "OrderedGroup.neg_le_iff_le"
- "REAL_LE_NEG" > "OrderedGroup.neg_le_iff_le"
+ "REAL_LE_NEGR" > "Groups.le_minus_self_iff"
+ "REAL_LE_NEGL" > "Groups.minus_le_self_iff"
+ "REAL_LE_NEG2" > "Groups.neg_le_iff_le"
+ "REAL_LE_NEG" > "Groups.neg_le_iff_le"
"REAL_LE_MUL2" > "HOL4Real.real.REAL_LE_MUL2"
- "REAL_LE_MUL" > "Ring_and_Field.mult_nonneg_nonneg"
+ "REAL_LE_MUL" > "Rings.mult_nonneg_nonneg"
"REAL_LE_LT" > "Orderings.order_le_less"
"REAL_LE_LNEG" > "RealDef.real_0_le_add_iff"
- "REAL_LE_LMUL_IMP" > "Ring_and_Field.mult_mono"
+ "REAL_LE_LMUL_IMP" > "Rings.mult_mono"
"REAL_LE_LMUL" > "RealDef.real_mult_le_cancel_iff2"
- "REAL_LE_LDIV_EQ" > "Ring_and_Field.pos_divide_le_eq"
- "REAL_LE_LDIV" > "Ring_and_Field.mult_imp_div_pos_le"
- "REAL_LE_LADD_IMP" > "OrderedGroup.add_left_mono"
- "REAL_LE_LADD" > "OrderedGroup.add_le_cancel_left"
- "REAL_LE_INV_EQ" > "Ring_and_Field.inverse_nonnegative_iff_nonnegative"
+ "REAL_LE_LDIV_EQ" > "Rings.pos_divide_le_eq"
+ "REAL_LE_LDIV" > "Rings.mult_imp_div_pos_le"
+ "REAL_LE_LADD_IMP" > "Groups.add_left_mono"
+ "REAL_LE_LADD" > "Groups.add_le_cancel_left"
+ "REAL_LE_INV_EQ" > "Rings.inverse_nonnegative_iff_nonnegative"
"REAL_LE_INV" > "HOL4Real.real.REAL_LE_INV"
- "REAL_LE_DOUBLE" > "OrderedGroup.zero_le_double_add_iff_zero_le_single_add"
+ "REAL_LE_DOUBLE" > "Groups.zero_le_double_add_iff_zero_le_single_add"
"REAL_LE_DIV" > "HOL4Real.real.REAL_LE_DIV"
"REAL_LE_ANTISYM" > "Orderings.order_eq_iff"
"REAL_LE_ADDR" > "HOL4Real.real.REAL_LE_ADDR"
"REAL_LE_ADDL" > "HOL4Real.real.REAL_LE_ADDL"
- "REAL_LE_ADD2" > "OrderedGroup.add_mono"
- "REAL_LE_ADD" > "OrderedGroup.add_nonneg_nonneg"
- "REAL_LE_01" > "Ring_and_Field.zero_le_one"
+ "REAL_LE_ADD2" > "Groups.add_mono"
+ "REAL_LE_ADD" > "Groups.add_nonneg_nonneg"
+ "REAL_LE_01" > "Rings.zero_le_one"
"REAL_LET_TRANS" > "Set.basic_trans_rules_23"
"REAL_LET_TOTAL" > "Orderings.linorder_le_less_linear"
"REAL_LET_ANTISYM" > "HOL4Real.real.REAL_LET_ANTISYM"
- "REAL_LET_ADD2" > "OrderedGroup.add_le_less_mono"
- "REAL_LET_ADD" > "OrderedGroup.add_nonneg_pos"
+ "REAL_LET_ADD2" > "Groups.add_le_less_mono"
+ "REAL_LET_ADD" > "Groups.add_nonneg_pos"
"REAL_LE1_POW2" > "HOL4Real.real.REAL_LE1_POW2"
"REAL_LE" > "RealDef.real_of_nat_le_iff"
- "REAL_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2"
- "REAL_INV_POS" > "Ring_and_Field.positive_imp_inverse_positive"
- "REAL_INV_NZ" > "Ring_and_Field.nonzero_imp_inverse_nonzero"
+ "REAL_LDISTRIB" > "Rings.ring_eq_simps_2"
+ "REAL_INV_POS" > "Rings.positive_imp_inverse_positive"
+ "REAL_INV_NZ" > "Rings.nonzero_imp_inverse_nonzero"
"REAL_INV_MUL" > "HOL4Real.real.REAL_INV_MUL"
"REAL_INV_LT1" > "RealDef.real_inverse_gt_one"
- "REAL_INV_INV" > "Ring_and_Field.inverse_inverse_eq"
- "REAL_INV_EQ_0" > "Ring_and_Field.inverse_nonzero_iff_nonzero"
- "REAL_INV_1OVER" > "Ring_and_Field.inverse_eq_divide"
- "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero"
- "REAL_INVINV" > "Ring_and_Field.nonzero_inverse_inverse_eq"
- "REAL_INV1" > "Ring_and_Field.inverse_1"
+ "REAL_INV_INV" > "Rings.inverse_inverse_eq"
+ "REAL_INV_EQ_0" > "Rings.inverse_nonzero_iff_nonzero"
+ "REAL_INV_1OVER" > "Rings.inverse_eq_divide"
+ "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero"
+ "REAL_INVINV" > "Rings.nonzero_inverse_inverse_eq"
+ "REAL_INV1" > "Rings.inverse_1"
"REAL_INJ" > "RealDef.real_of_nat_inject"
"REAL_HALF_DOUBLE" > "RComplete.real_sum_of_halves"
"REAL_FACT_NZ" > "HOL4Real.real.REAL_FACT_NZ"
- "REAL_EQ_SUB_RADD" > "Ring_and_Field.ring_eq_simps_15"
- "REAL_EQ_SUB_LADD" > "Ring_and_Field.ring_eq_simps_16"
- "REAL_EQ_RMUL_IMP" > "Ring_and_Field.field_mult_cancel_right_lemma"
- "REAL_EQ_RMUL" > "Ring_and_Field.field_mult_cancel_right"
+ "REAL_EQ_SUB_RADD" > "Rings.ring_eq_simps_15"
+ "REAL_EQ_SUB_LADD" > "Rings.ring_eq_simps_16"
+ "REAL_EQ_RMUL_IMP" > "Rings.field_mult_cancel_right_lemma"
+ "REAL_EQ_RMUL" > "Rings.field_mult_cancel_right"
"REAL_EQ_RDIV_EQ" > "HOL4Real.real.REAL_EQ_RDIV_EQ"
- "REAL_EQ_RADD" > "OrderedGroup.add_right_cancel"
- "REAL_EQ_NEG" > "OrderedGroup.neg_equal_iff_equal"
- "REAL_EQ_MUL_LCANCEL" > "Ring_and_Field.field_mult_cancel_left"
+ "REAL_EQ_RADD" > "Groups.add_right_cancel"
+ "REAL_EQ_NEG" > "Groups.neg_equal_iff_equal"
+ "REAL_EQ_MUL_LCANCEL" > "Rings.field_mult_cancel_left"
"REAL_EQ_LMUL_IMP" > "HOL4Real.real.REAL_EQ_LMUL_IMP"
"REAL_EQ_LMUL2" > "RealDef.real_mult_left_cancel"
- "REAL_EQ_LMUL" > "Ring_and_Field.field_mult_cancel_left"
+ "REAL_EQ_LMUL" > "Rings.field_mult_cancel_left"
"REAL_EQ_LDIV_EQ" > "HOL4Real.real.REAL_EQ_LDIV_EQ"
- "REAL_EQ_LADD" > "OrderedGroup.add_left_cancel"
+ "REAL_EQ_LADD" > "Groups.add_left_cancel"
"REAL_EQ_IMP_LE" > "Orderings.order_eq_refl"
- "REAL_ENTIRE" > "Ring_and_Field.field_mult_eq_0_iff"
+ "REAL_ENTIRE" > "Rings.field_mult_eq_0_iff"
"REAL_DOWN2" > "RealDef.real_lbound_gt_zero"
"REAL_DOWN" > "HOL4Real.real.REAL_DOWN"
"REAL_DOUBLE" > "Int.mult_2"
"REAL_DIV_RMUL" > "HOL4Real.real.REAL_DIV_RMUL"
- "REAL_DIV_REFL" > "Ring_and_Field.divide_self"
+ "REAL_DIV_REFL" > "Rings.divide_self"
"REAL_DIV_MUL2" > "HOL4Real.real.REAL_DIV_MUL2"
- "REAL_DIV_LZERO" > "Ring_and_Field.divide_zero_left"
+ "REAL_DIV_LZERO" > "Rings.divide_zero_left"
"REAL_DIV_LMUL" > "HOL4Real.real.REAL_DIV_LMUL"
"REAL_DIFFSQ" > "HOL4Real.real.REAL_DIFFSQ"
"REAL_ARCH_LEAST" > "HOL4Real.real.REAL_ARCH_LEAST"
@@ -286,20 +286,20 @@
"REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
"REAL_ADD_SUB2" > "HOL4Real.real.REAL_ADD_SUB2"
"REAL_ADD_SUB" > "HOL4Real.real.REAL_ADD_SUB"
- "REAL_ADD_RINV" > "OrderedGroup.right_minus"
+ "REAL_ADD_RINV" > "Groups.right_minus"
"REAL_ADD_RID_UNIQ" > "HOL4Real.real.REAL_ADD_RID_UNIQ"
"REAL_ADD_RID" > "Finite_Set.AC_add.f_e.ident"
- "REAL_ADD_RDISTRIB" > "Ring_and_Field.ring_eq_simps_1"
+ "REAL_ADD_RDISTRIB" > "Rings.ring_eq_simps_1"
"REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
"REAL_ADD_LID_UNIQ" > "HOL4Real.real.REAL_ADD_LID_UNIQ"
"REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident"
- "REAL_ADD_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2"
+ "REAL_ADD_LDISTRIB" > "Rings.ring_eq_simps_2"
"REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC"
"REAL_ADD2_SUB2" > "HOL4Real.real.REAL_ADD2_SUB2"
"REAL_ADD" > "RealDef.real_of_nat_add"
- "REAL_ABS_TRIANGLE" > "OrderedGroup.abs_triangle_ineq"
- "REAL_ABS_POS" > "OrderedGroup.abs_ge_zero"
- "REAL_ABS_MUL" > "Ring_and_Field.abs_mult"
+ "REAL_ABS_TRIANGLE" > "Groups.abs_triangle_ineq"
+ "REAL_ABS_POS" > "Groups.abs_ge_zero"
+ "REAL_ABS_MUL" > "Rings.abs_mult"
"REAL_ABS_0" > "Int.bin_arith_simps_28"
"REAL_10" > "HOL4Compat.REAL_10"
"REAL_1" > "HOL4Real.real.REAL_1"
@@ -326,25 +326,25 @@
"POW_2" > "Nat_Numeral.power2_eq_square"
"POW_1" > "Power.power_one_right"
"POW_0" > "Power.power_0_Suc"
- "ABS_ZERO" > "OrderedGroup.abs_eq_0"
- "ABS_TRIANGLE" > "OrderedGroup.abs_triangle_ineq"
+ "ABS_ZERO" > "Groups.abs_eq_0"
+ "ABS_TRIANGLE" > "Groups.abs_triangle_ineq"
"ABS_SUM" > "HOL4Real.real.ABS_SUM"
- "ABS_SUB_ABS" > "OrderedGroup.abs_triangle_ineq3"
- "ABS_SUB" > "OrderedGroup.abs_minus_commute"
+ "ABS_SUB_ABS" > "Groups.abs_triangle_ineq3"
+ "ABS_SUB" > "Groups.abs_minus_commute"
"ABS_STILLNZ" > "HOL4Real.real.ABS_STILLNZ"
"ABS_SIGN2" > "HOL4Real.real.ABS_SIGN2"
"ABS_SIGN" > "HOL4Real.real.ABS_SIGN"
"ABS_REFL" > "HOL4Real.real.ABS_REFL"
"ABS_POW2" > "Nat_Numeral.abs_power2"
- "ABS_POS" > "OrderedGroup.abs_ge_zero"
- "ABS_NZ" > "OrderedGroup.zero_less_abs_iff"
- "ABS_NEG" > "OrderedGroup.abs_minus_cancel"
+ "ABS_POS" > "Groups.abs_ge_zero"
+ "ABS_NZ" > "Groups.zero_less_abs_iff"
+ "ABS_NEG" > "Groups.abs_minus_cancel"
"ABS_N" > "RealDef.abs_real_of_nat_cancel"
- "ABS_MUL" > "Ring_and_Field.abs_mult"
+ "ABS_MUL" > "Rings.abs_mult"
"ABS_LT_MUL2" > "HOL4Real.real.ABS_LT_MUL2"
- "ABS_LE" > "OrderedGroup.abs_ge_self"
- "ABS_INV" > "Ring_and_Field.nonzero_abs_inverse"
- "ABS_DIV" > "Ring_and_Field.nonzero_abs_divide"
+ "ABS_LE" > "Groups.abs_ge_self"
+ "ABS_INV" > "Rings.nonzero_abs_inverse"
+ "ABS_DIV" > "Rings.nonzero_abs_divide"
"ABS_CIRCLE" > "HOL4Real.real.ABS_CIRCLE"
"ABS_CASES" > "HOL4Real.real.ABS_CASES"
"ABS_BOUNDS" > "RealDef.abs_le_interval_iff"
@@ -352,7 +352,7 @@
"ABS_BETWEEN2" > "HOL4Real.real.ABS_BETWEEN2"
"ABS_BETWEEN1" > "HOL4Real.real.ABS_BETWEEN1"
"ABS_BETWEEN" > "HOL4Real.real.ABS_BETWEEN"
- "ABS_ABS" > "OrderedGroup.abs_idempotent"
+ "ABS_ABS" > "Groups.abs_idempotent"
"ABS_1" > "Int.bin_arith_simps_29"
"ABS_0" > "Int.bin_arith_simps_28"
--- a/src/HOL/Import/HOL/realax.imp Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Import/HOL/realax.imp Mon Feb 08 17:13:45 2010 +0100
@@ -98,10 +98,10 @@
"REAL_LT_TRANS" > "Set.basic_trans_rules_21"
"REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
"REAL_LT_REFL" > "Orderings.order_less_irrefl"
- "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos"
- "REAL_LT_IADD" > "OrderedGroup.add_strict_left_mono"
- "REAL_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2"
- "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero"
+ "REAL_LT_MUL" > "Rings.mult_pos_pos"
+ "REAL_LT_IADD" > "Groups.add_strict_left_mono"
+ "REAL_LDISTRIB" > "Rings.ring_eq_simps_2"
+ "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero"
"REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
"REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
"REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident"
--- a/src/HOL/Int.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Int.thy Mon Feb 08 17:13:45 2010 +0100
@@ -307,7 +307,7 @@
by (cases z, simp add: algebra_simps of_int minus)
lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"
-by (simp add: OrderedGroup.diff_minus diff_minus)
+by (simp add: diff_minus Groups.diff_minus)
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
apply (cases w, cases z)
@@ -519,7 +519,7 @@
text{*This version is proved for all ordered rings, not just integers!
It is proved here because attribute @{text arith_split} is not available
- in theory @{text Ring_and_Field}.
+ in theory @{text Rings}.
But is it really better than just rewriting with @{text abs_if}?*}
lemma abs_split [arith_split,noatp]:
"P(abs(a::'a::linordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
@@ -2317,9 +2317,9 @@
lemmas zadd_assoc = add_assoc [of "z1::int" "z2" "z3", standard]
lemmas zadd_left_commute = add_left_commute [of "x::int" "y" "z", standard]
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
-lemmas zmult_ac = OrderedGroup.mult_ac
-lemmas zadd_0 = OrderedGroup.add_0_left [of "z::int", standard]
-lemmas zadd_0_right = OrderedGroup.add_0_left [of "z::int", standard]
+lemmas zmult_ac = mult_ac
+lemmas zadd_0 = add_0_left [of "z::int", standard]
+lemmas zadd_0_right = add_0_right [of "z::int", standard]
lemmas zadd_zminus_inverse2 = left_minus [of "z::int", standard]
lemmas zmult_zminus = mult_minus_left [of "z::int" "w", standard]
lemmas zmult_commute = mult_commute [of "z::int" "w", standard]
--- a/src/HOL/IsaMakefile Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/IsaMakefile Mon Feb 08 17:13:45 2010 +0100
@@ -145,15 +145,16 @@
Complete_Lattice.thy \
Datatype.thy \
Extraction.thy \
+ Fields.thy \
Finite_Set.thy \
Fun.thy \
FunDef.thy \
+ Groups.thy \
Inductive.thy \
Lattices.thy \
Nat.thy \
Nitpick.thy \
Option.thy \
- OrderedGroup.thy \
Orderings.thy \
Plain.thy \
Power.thy \
@@ -162,7 +163,7 @@
Record.thy \
Refute.thy \
Relation.thy \
- Ring_and_Field.thy \
+ Rings.thy \
SAT.thy \
Set.thy \
Sum_Type.thy \
--- a/src/HOL/Library/Poly_Deriv.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Library/Poly_Deriv.thy Mon Feb 08 17:13:45 2010 +0100
@@ -139,7 +139,7 @@
lemma dvd_add_cancel1:
fixes a b c :: "'a::comm_ring_1"
shows "a dvd b + c \<Longrightarrow> a dvd b \<Longrightarrow> a dvd c"
- by (drule (1) Ring_and_Field.dvd_diff, simp)
+ by (drule (1) Rings.dvd_diff, simp)
lemma lemma_order_pderiv [rule_format]:
"\<forall>p q a. 0 < n &
--- a/src/HOL/Metis_Examples/BigO.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Metis_Examples/BigO.thy Mon Feb 08 17:13:45 2010 +0100
@@ -368,7 +368,7 @@
f : O(g)"
apply (auto simp add: bigo_def)
(*Version 1: one-shot proof*)
- apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult)
+ apply (metis abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) abs_mult)
done
lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==>
@@ -595,8 +595,8 @@
using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]]
prefer 2
apply (metis mult_assoc mult_left_commute
- OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute
- Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos)
+ abs_of_pos mult_left_commute
+ abs_mult mult_pos_pos)
apply (erule ssubst)
apply (subst abs_mult)
(*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has
@@ -613,32 +613,32 @@
\<le> (c\<Colon>'b\<Colon>linordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar> *
((ca\<Colon>'b\<Colon>linordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar>)"
have 4: "\<bar>c\<Colon>'b\<Colon>linordered_idom\<bar> = c"
- by (metis OrderedGroup.abs_of_pos 0)
+ by (metis abs_of_pos 0)
have 5: "\<And>X1\<Colon>'b\<Colon>linordered_idom. (c\<Colon>'b\<Colon>linordered_idom) * \<bar>X1\<bar> = \<bar>c * X1\<bar>"
- by (metis Ring_and_Field.abs_mult 4)
+ by (metis abs_mult 4)
have 6: "(0\<Colon>'b\<Colon>linordered_idom) = (1\<Colon>'b\<Colon>linordered_idom) \<or>
(0\<Colon>'b\<Colon>linordered_idom) < (1\<Colon>'b\<Colon>linordered_idom)"
- by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_linordered_idom)
+ by (metis abs_not_less_zero abs_one linorder_neqE_linordered_idom)
have 7: "(0\<Colon>'b\<Colon>linordered_idom) < (1\<Colon>'b\<Colon>linordered_idom)"
- by (metis 6 Ring_and_Field.one_neq_zero)
+ by (metis 6 one_neq_zero)
have 8: "\<bar>1\<Colon>'b\<Colon>linordered_idom\<bar> = (1\<Colon>'b\<Colon>linordered_idom)"
- by (metis OrderedGroup.abs_of_pos 7)
+ by (metis abs_of_pos 7)
have 9: "\<And>X1\<Colon>'b\<Colon>linordered_idom. (0\<Colon>'b\<Colon>linordered_idom) \<le> (c\<Colon>'b\<Colon>linordered_idom) * \<bar>X1\<bar>"
- by (metis OrderedGroup.abs_ge_zero 5)
+ by (metis abs_ge_zero 5)
have 10: "\<And>X1\<Colon>'b\<Colon>linordered_idom. X1 * (1\<Colon>'b\<Colon>linordered_idom) = X1"
- by (metis Ring_and_Field.mult_cancel_right2 mult_commute)
+ by (metis mult_cancel_right2 mult_commute)
have 11: "\<And>X1\<Colon>'b\<Colon>linordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar> * \<bar>1\<Colon>'b\<Colon>linordered_idom\<bar>"
- by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10)
+ by (metis abs_mult abs_idempotent 10)
have 12: "\<And>X1\<Colon>'b\<Colon>linordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar>"
by (metis 11 8 10)
have 13: "\<And>X1\<Colon>'b\<Colon>linordered_idom. (0\<Colon>'b\<Colon>linordered_idom) \<le> \<bar>X1\<bar>"
- by (metis OrderedGroup.abs_ge_zero 12)
+ by (metis abs_ge_zero 12)
have 14: "\<not> (0\<Colon>'b\<Colon>linordered_idom)
\<le> (c\<Colon>'b\<Colon>linordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) (x\<Colon>'a)\<bar> \<or>
\<not> (0\<Colon>'b\<Colon>linordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar> \<or>
\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>linordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar> \<or>
\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar> \<le> c * \<bar>f x\<bar>"
- by (metis 3 Ring_and_Field.mult_mono)
+ by (metis 3 mult_mono)
have 15: "\<not> (0\<Colon>'b\<Colon>linordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) (x\<Colon>'a)\<bar> \<or>
\<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>linordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar> \<or>
\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>linordered_idom) x\<bar>
@@ -1078,7 +1078,7 @@
apply (rule_tac x = c in exI)
apply safe
apply (case_tac "x = 0")
-apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le)
+apply (metis abs_ge_zero abs_zero order_less_le split_mult_pos_le)
apply (subgoal_tac "x = Suc (x - 1)")
apply metis
apply simp
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Feb 08 17:13:45 2010 +0100
@@ -836,7 +836,7 @@
lemma dot_rneg: "(x::'a::ring ^ 'n) \<bullet> (-y) = -(x \<bullet> y)" by vector
lemma dot_lzero[simp]: "0 \<bullet> x = (0::'a::{comm_monoid_add, mult_zero})" by vector
lemma dot_rzero[simp]: "x \<bullet> 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector
-lemma dot_pos_le[simp]: "(0::'a\<Colon>linlinordered_ring_strict) <= x \<bullet> x"
+lemma dot_pos_le[simp]: "(0::'a\<Colon>linordered_ring_strict) <= x \<bullet> x"
by (simp add: dot_def setsum_nonneg)
lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\<forall>x \<in> F. f x \<ge> (0 ::'a::ordered_ab_group_add)" shows "setsum f F = 0 \<longleftrightarrow> (ALL x:F. f x = 0)"
@@ -852,10 +852,10 @@
show ?case by (simp add: h)
qed
-lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{linlinordered_ring_strict,ring_no_zero_divisors} ^ 'n) = 0"
+lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{linordered_ring_strict,ring_no_zero_divisors} ^ 'n) = 0"
by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq)
-lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{linlinordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
+lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{linordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
by (auto simp add: le_less)
subsection{* The collapse of the general concepts to dimension one. *}
--- a/src/HOL/NSA/HyperDef.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/NSA/HyperDef.thy Mon Feb 08 17:13:45 2010 +0100
@@ -394,7 +394,7 @@
by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two)
(*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract
- result proved in Ring_and_Field*)
+ result proved in Rings or Fields*)
lemma hrabs_hrealpow_two [simp]:
"abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)"
by (simp add: abs_mult)
@@ -501,12 +501,12 @@
by transfer (rule power_mult_distrib)
lemma hyperpow_two_le [simp]:
- "(0::'a::{monoid_mult,linlinordered_ring_strict} star) \<le> r pow (1 + 1)"
+ "(0::'a::{monoid_mult,linordered_ring_strict} star) \<le> r pow (1 + 1)"
by (auto simp add: hyperpow_two zero_le_mult_iff)
lemma hrabs_hyperpow_two [simp]:
"abs(x pow (1 + 1)) =
- (x::'a::{monoid_mult,linlinordered_ring_strict} star) pow (1 + 1)"
+ (x::'a::{monoid_mult,linordered_ring_strict} star) pow (1 + 1)"
by (simp only: abs_of_nonneg hyperpow_two_le)
lemma hyperpow_two_hrabs [simp]:
@@ -516,7 +516,7 @@
text{*The precondition could be weakened to @{term "0\<le>x"}*}
lemma hypreal_mult_less_mono:
"[| u<v; x<y; (0::hypreal) < v; 0 < x |] ==> u*x < v* y"
- by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le)
+ by (simp add: mult_strict_mono order_less_imp_le)
lemma hyperpow_two_gt_one:
"\<And>r::'a::{linordered_semidom} star. 1 < r \<Longrightarrow> 1 < r pow (1 + 1)"
--- a/src/HOL/NSA/NSComplex.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/NSA/NSComplex.thy Mon Feb 08 17:13:45 2010 +0100
@@ -1,5 +1,4 @@
(* Title: NSComplex.thy
- ID: $Id$
Author: Jacques D. Fleuriot
Copyright: 2001 University of Edinburgh
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
@@ -161,7 +160,7 @@
lemma hcomplex_add_minus_eq_minus:
"x + y = (0::hcomplex) ==> x = -y"
-apply (drule OrderedGroup.minus_unique)
+apply (drule minus_unique)
apply (simp add: minus_equation_iff [of x y])
done
@@ -196,7 +195,7 @@
lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)"
(* TODO: delete *)
-by (rule OrderedGroup.diff_eq_eq)
+by (rule diff_eq_eq)
subsection{*Embedding Properties for @{term hcomplex_of_hypreal} Map*}
--- a/src/HOL/NSA/StarDef.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/NSA/StarDef.thy Mon Feb 08 17:13:45 2010 +0100
@@ -530,7 +530,7 @@
end
-instance star :: (Ring_and_Field.dvd) Ring_and_Field.dvd ..
+instance star :: (Rings.dvd) Rings.dvd ..
instantiation star :: (Divides.div) Divides.div
begin
@@ -912,7 +912,7 @@
instance star :: (ordered_cancel_semiring) ordered_cancel_semiring ..
-instance star :: (linlinordered_semiring_strict) linlinordered_semiring_strict
+instance star :: (linordered_semiring_strict) linordered_semiring_strict
apply (intro_classes)
apply (transfer, erule (1) mult_strict_left_mono)
apply (transfer, erule (1) mult_strict_right_mono)
@@ -936,7 +936,7 @@
instance star :: (sgn_if) sgn_if
by (intro_classes, transfer, rule sgn_if)
-instance star :: (linlinordered_ring_strict) linlinordered_ring_strict ..
+instance star :: (linordered_ring_strict) linordered_ring_strict ..
instance star :: (ordered_comm_ring) ordered_comm_ring ..
instance star :: (linordered_semidom) linordered_semidom
--- a/src/HOL/Nat.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Nat.thy Mon Feb 08 17:13:45 2010 +0100
@@ -8,7 +8,7 @@
header {* Natural numbers *}
theory Nat
-imports Inductive Product_Type Ring_and_Field
+imports Inductive Product_Type Fields
uses
"~~/src/Tools/rat.ML"
"~~/src/Provers/Arith/cancel_sums.ML"
@@ -176,6 +176,8 @@
end
+hide (open) fact add_0_right
+
instantiation nat :: comm_semiring_1_cancel
begin
@@ -730,7 +732,7 @@
apply (induct n)
apply (simp_all add: order_le_less)
apply (blast elim!: less_SucE
- intro!: add_0_right [symmetric] add_Suc_right [symmetric])
+ intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric])
done
text {* strict, in 1st argument; proof is by induction on @{text "k > 0"} *}
--- a/src/HOL/Nat_Numeral.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Nat_Numeral.thy Mon Feb 08 17:13:45 2010 +0100
@@ -64,7 +64,7 @@
lemma power_even_eq:
"a ^ (2*n) = (a ^ n) ^ 2"
- by (subst OrderedGroup.mult_commute) (simp add: power_mult)
+ by (subst mult_commute) (simp add: power_mult)
lemma power_odd_eq:
"a ^ Suc (2*n) = a * (a ^ n) ^ 2"
@@ -113,7 +113,7 @@
end
-context linlinordered_ring_strict
+context linordered_ring_strict
begin
lemma sum_squares_ge_zero:
--- a/src/HOL/Old_Number_Theory/WilsonBij.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Old_Number_Theory/WilsonBij.thy Mon Feb 08 17:13:45 2010 +0100
@@ -74,9 +74,9 @@
lemma aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
-- {* same as @{text WilsonRuss} *}
apply (unfold zcong_def)
- apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
+ apply (simp add: diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
- apply (simp add: mult_commute)
+ apply (simp add: algebra_simps)
apply (subst dvd_minus_iff)
apply (subst zdvd_reduce)
apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
--- a/src/HOL/Old_Number_Theory/WilsonRuss.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy Mon Feb 08 17:13:45 2010 +0100
@@ -82,9 +82,9 @@
lemma inv_not_p_minus_1_aux:
"[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
apply (unfold zcong_def)
- apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
+ apply (simp add: diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
- apply (simp add: mult_commute)
+ apply (simp add: algebra_simps)
apply (subst dvd_minus_iff)
apply (subst zdvd_reduce)
apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
--- a/src/HOL/OrderedGroup.thy Mon Feb 08 15:25:00 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1159 +0,0 @@
-(* Title: HOL/OrderedGroup.thy
- Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad
-*)
-
-header {* Ordered Groups *}
-
-theory OrderedGroup
-imports Lattices
-uses "~~/src/Provers/Arith/abel_cancel.ML"
-begin
-
-text {*
- The theory of partially ordered groups is taken from the books:
- \begin{itemize}
- \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
- \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
- \end{itemize}
- Most of the used notions can also be looked up in
- \begin{itemize}
- \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
- \item \emph{Algebra I} by van der Waerden, Springer.
- \end{itemize}
-*}
-
-ML {*
-structure Algebra_Simps = Named_Thms(
- val name = "algebra_simps"
- val description = "algebra simplification rules"
-)
-*}
-
-setup Algebra_Simps.setup
-
-text{* The rewrites accumulated in @{text algebra_simps} deal with the
-classical algebraic structures of groups, rings and family. They simplify
-terms by multiplying everything out (in case of a ring) and bringing sums and
-products into a canonical form (by ordered rewriting). As a result it decides
-group and ring equalities but also helps with inequalities.
-
-Of course it also works for fields, but it knows nothing about multiplicative
-inverses or division. This is catered for by @{text field_simps}. *}
-
-subsection {* Semigroups and Monoids *}
-
-class semigroup_add = plus +
- assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
-
-sublocale semigroup_add < plus!: semigroup plus proof
-qed (fact add_assoc)
-
-class ab_semigroup_add = semigroup_add +
- assumes add_commute [algebra_simps]: "a + b = b + a"
-
-sublocale ab_semigroup_add < plus!: abel_semigroup plus proof
-qed (fact add_commute)
-
-context ab_semigroup_add
-begin
-
-lemmas add_left_commute [algebra_simps] = plus.left_commute
-
-theorems add_ac = add_assoc add_commute add_left_commute
-
-end
-
-theorems add_ac = add_assoc add_commute add_left_commute
-
-class semigroup_mult = times +
- assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)"
-
-sublocale semigroup_mult < times!: semigroup times proof
-qed (fact mult_assoc)
-
-class ab_semigroup_mult = semigroup_mult +
- assumes mult_commute [algebra_simps]: "a * b = b * a"
-
-sublocale ab_semigroup_mult < times!: abel_semigroup times proof
-qed (fact mult_commute)
-
-context ab_semigroup_mult
-begin
-
-lemmas mult_left_commute [algebra_simps] = times.left_commute
-
-theorems mult_ac = mult_assoc mult_commute mult_left_commute
-
-end
-
-theorems mult_ac = mult_assoc mult_commute mult_left_commute
-
-class ab_semigroup_idem_mult = ab_semigroup_mult +
- assumes mult_idem: "x * x = x"
-
-sublocale ab_semigroup_idem_mult < times!: semilattice times proof
-qed (fact mult_idem)
-
-context ab_semigroup_idem_mult
-begin
-
-lemmas mult_left_idem = times.left_idem
-
-end
-
-class monoid_add = zero + semigroup_add +
- assumes add_0_left [simp]: "0 + a = a"
- and add_0_right [simp]: "a + 0 = a"
-
-lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"
-by (rule eq_commute)
-
-class comm_monoid_add = zero + ab_semigroup_add +
- assumes add_0: "0 + a = a"
-begin
-
-subclass monoid_add
- proof qed (insert add_0, simp_all add: add_commute)
-
-end
-
-class monoid_mult = one + semigroup_mult +
- assumes mult_1_left [simp]: "1 * a = a"
- assumes mult_1_right [simp]: "a * 1 = a"
-
-lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"
-by (rule eq_commute)
-
-class comm_monoid_mult = one + ab_semigroup_mult +
- assumes mult_1: "1 * a = a"
-begin
-
-subclass monoid_mult
- proof qed (insert mult_1, simp_all add: mult_commute)
-
-end
-
-class cancel_semigroup_add = semigroup_add +
- assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
- assumes add_right_imp_eq: "b + a = c + a \<Longrightarrow> b = c"
-begin
-
-lemma add_left_cancel [simp]:
- "a + b = a + c \<longleftrightarrow> b = c"
-by (blast dest: add_left_imp_eq)
-
-lemma add_right_cancel [simp]:
- "b + a = c + a \<longleftrightarrow> b = c"
-by (blast dest: add_right_imp_eq)
-
-end
-
-class cancel_ab_semigroup_add = ab_semigroup_add +
- assumes add_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
-begin
-
-subclass cancel_semigroup_add
-proof
- fix a b c :: 'a
- assume "a + b = a + c"
- then show "b = c" by (rule add_imp_eq)
-next
- fix a b c :: 'a
- assume "b + a = c + a"
- then have "a + b = a + c" by (simp only: add_commute)
- then show "b = c" by (rule add_imp_eq)
-qed
-
-end
-
-class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add
-
-
-subsection {* Groups *}
-
-class group_add = minus + uminus + monoid_add +
- assumes left_minus [simp]: "- a + a = 0"
- assumes diff_minus: "a - b = a + (- b)"
-begin
-
-lemma minus_unique:
- assumes "a + b = 0" shows "- a = b"
-proof -
- have "- a = - a + (a + b)" using assms by simp
- also have "\<dots> = b" by (simp add: add_assoc [symmetric])
- finally show ?thesis .
-qed
-
-lemmas equals_zero_I = minus_unique (* legacy name *)
-
-lemma minus_zero [simp]: "- 0 = 0"
-proof -
- have "0 + 0 = 0" by (rule add_0_right)
- thus "- 0 = 0" by (rule minus_unique)
-qed
-
-lemma minus_minus [simp]: "- (- a) = a"
-proof -
- have "- a + a = 0" by (rule left_minus)
- thus "- (- a) = a" by (rule minus_unique)
-qed
-
-lemma right_minus [simp]: "a + - a = 0"
-proof -
- have "a + - a = - (- a) + - a" by simp
- also have "\<dots> = 0" by (rule left_minus)
- finally show ?thesis .
-qed
-
-lemma minus_add_cancel: "- a + (a + b) = b"
-by (simp add: add_assoc [symmetric])
-
-lemma add_minus_cancel: "a + (- a + b) = b"
-by (simp add: add_assoc [symmetric])
-
-lemma minus_add: "- (a + b) = - b + - a"
-proof -
- have "(a + b) + (- b + - a) = 0"
- by (simp add: add_assoc add_minus_cancel)
- thus "- (a + b) = - b + - a"
- by (rule minus_unique)
-qed
-
-lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"
-proof
- assume "a - b = 0"
- have "a = (a - b) + b" by (simp add:diff_minus add_assoc)
- also have "\<dots> = b" using `a - b = 0` by simp
- finally show "a = b" .
-next
- assume "a = b" thus "a - b = 0" by (simp add: diff_minus)
-qed
-
-lemma diff_self [simp]: "a - a = 0"
-by (simp add: diff_minus)
-
-lemma diff_0 [simp]: "0 - a = - a"
-by (simp add: diff_minus)
-
-lemma diff_0_right [simp]: "a - 0 = a"
-by (simp add: diff_minus)
-
-lemma diff_minus_eq_add [simp]: "a - - b = a + b"
-by (simp add: diff_minus)
-
-lemma neg_equal_iff_equal [simp]:
- "- a = - b \<longleftrightarrow> a = b"
-proof
- assume "- a = - b"
- hence "- (- a) = - (- b)" by simp
- thus "a = b" by simp
-next
- assume "a = b"
- thus "- a = - b" by simp
-qed
-
-lemma neg_equal_0_iff_equal [simp]:
- "- a = 0 \<longleftrightarrow> a = 0"
-by (subst neg_equal_iff_equal [symmetric], simp)
-
-lemma neg_0_equal_iff_equal [simp]:
- "0 = - a \<longleftrightarrow> 0 = a"
-by (subst neg_equal_iff_equal [symmetric], simp)
-
-text{*The next two equations can make the simplifier loop!*}
-
-lemma equation_minus_iff:
- "a = - b \<longleftrightarrow> b = - a"
-proof -
- have "- (- a) = - b \<longleftrightarrow> - a = b" by (rule neg_equal_iff_equal)
- thus ?thesis by (simp add: eq_commute)
-qed
-
-lemma minus_equation_iff:
- "- a = b \<longleftrightarrow> - b = a"
-proof -
- have "- a = - (- b) \<longleftrightarrow> a = -b" by (rule neg_equal_iff_equal)
- thus ?thesis by (simp add: eq_commute)
-qed
-
-lemma diff_add_cancel: "a - b + b = a"
-by (simp add: diff_minus add_assoc)
-
-lemma add_diff_cancel: "a + b - b = a"
-by (simp add: diff_minus add_assoc)
-
-declare diff_minus[symmetric, algebra_simps]
-
-lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
-proof
- assume "a = - b" then show "a + b = 0" by simp
-next
- assume "a + b = 0"
- moreover have "a + (b + - b) = (a + b) + - b"
- by (simp only: add_assoc)
- ultimately show "a = - b" by simp
-qed
-
-end
-
-class ab_group_add = minus + uminus + comm_monoid_add +
- assumes ab_left_minus: "- a + a = 0"
- assumes ab_diff_minus: "a - b = a + (- b)"
-begin
-
-subclass group_add
- proof qed (simp_all add: ab_left_minus ab_diff_minus)
-
-subclass cancel_comm_monoid_add
-proof
- fix a b c :: 'a
- assume "a + b = a + c"
- then have "- a + a + b = - a + a + c"
- unfolding add_assoc by simp
- then show "b = c" by simp
-qed
-
-lemma uminus_add_conv_diff[algebra_simps]:
- "- a + b = b - a"
-by (simp add:diff_minus add_commute)
-
-lemma minus_add_distrib [simp]:
- "- (a + b) = - a + - b"
-by (rule minus_unique) (simp add: add_ac)
-
-lemma minus_diff_eq [simp]:
- "- (a - b) = b - a"
-by (simp add: diff_minus add_commute)
-
-lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c"
-by (simp add: diff_minus add_ac)
-
-lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b"
-by (simp add: diff_minus add_ac)
-
-lemma diff_eq_eq[algebra_simps]: "a - b = c \<longleftrightarrow> a = c + b"
-by (auto simp add: diff_minus add_assoc)
-
-lemma eq_diff_eq[algebra_simps]: "a = c - b \<longleftrightarrow> a + b = c"
-by (auto simp add: diff_minus add_assoc)
-
-lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)"
-by (simp add: diff_minus add_ac)
-
-lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b"
-by (simp add: diff_minus add_ac)
-
-lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
-by (simp add: algebra_simps)
-
-lemma diff_eq_0_iff_eq [simp, noatp]: "a - b = 0 \<longleftrightarrow> a = b"
-by (simp add: algebra_simps)
-
-end
-
-subsection {* (Partially) Ordered Groups *}
-
-class ordered_ab_semigroup_add = order + ab_semigroup_add +
- assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
-begin
-
-lemma add_right_mono:
- "a \<le> b \<Longrightarrow> a + c \<le> b + c"
-by (simp add: add_commute [of _ c] add_left_mono)
-
-text {* non-strict, in both arguments *}
-lemma add_mono:
- "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d"
- apply (erule add_right_mono [THEN order_trans])
- apply (simp add: add_commute add_left_mono)
- done
-
-end
-
-class ordered_cancel_ab_semigroup_add =
- ordered_ab_semigroup_add + cancel_ab_semigroup_add
-begin
-
-lemma add_strict_left_mono:
- "a < b \<Longrightarrow> c + a < c + b"
-by (auto simp add: less_le add_left_mono)
-
-lemma add_strict_right_mono:
- "a < b \<Longrightarrow> a + c < b + c"
-by (simp add: add_commute [of _ c] add_strict_left_mono)
-
-text{*Strict monotonicity in both arguments*}
-lemma add_strict_mono:
- "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
-apply (erule add_strict_right_mono [THEN less_trans])
-apply (erule add_strict_left_mono)
-done
-
-lemma add_less_le_mono:
- "a < b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c < b + d"
-apply (erule add_strict_right_mono [THEN less_le_trans])
-apply (erule add_left_mono)
-done
-
-lemma add_le_less_mono:
- "a \<le> b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
-apply (erule add_right_mono [THEN le_less_trans])
-apply (erule add_strict_left_mono)
-done
-
-end
-
-class ordered_ab_semigroup_add_imp_le =
- ordered_cancel_ab_semigroup_add +
- assumes add_le_imp_le_left: "c + a \<le> c + b \<Longrightarrow> a \<le> b"
-begin
-
-lemma add_less_imp_less_left:
- assumes less: "c + a < c + b" shows "a < b"
-proof -
- from less have le: "c + a <= c + b" by (simp add: order_le_less)
- have "a <= b"
- apply (insert le)
- apply (drule add_le_imp_le_left)
- by (insert le, drule add_le_imp_le_left, assumption)
- moreover have "a \<noteq> b"
- proof (rule ccontr)
- assume "~(a \<noteq> b)"
- then have "a = b" by simp
- then have "c + a = c + b" by simp
- with less show "False"by simp
- qed
- ultimately show "a < b" by (simp add: order_le_less)
-qed
-
-lemma add_less_imp_less_right:
- "a + c < b + c \<Longrightarrow> a < b"
-apply (rule add_less_imp_less_left [of c])
-apply (simp add: add_commute)
-done
-
-lemma add_less_cancel_left [simp]:
- "c + a < c + b \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_left add_strict_left_mono)
-
-lemma add_less_cancel_right [simp]:
- "a + c < b + c \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_right add_strict_right_mono)
-
-lemma add_le_cancel_left [simp]:
- "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
-by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)
-
-lemma add_le_cancel_right [simp]:
- "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
-by (simp add: add_commute [of a c] add_commute [of b c])
-
-lemma add_le_imp_le_right:
- "a + c \<le> b + c \<Longrightarrow> a \<le> b"
-by simp
-
-lemma max_add_distrib_left:
- "max x y + z = max (x + z) (y + z)"
- unfolding max_def by auto
-
-lemma min_add_distrib_left:
- "min x y + z = min (x + z) (y + z)"
- unfolding min_def by auto
-
-end
-
-subsection {* Support for reasoning about signs *}
-
-class ordered_comm_monoid_add =
- ordered_cancel_ab_semigroup_add + comm_monoid_add
-begin
-
-lemma add_pos_nonneg:
- assumes "0 < a" and "0 \<le> b" shows "0 < a + b"
-proof -
- have "0 + 0 < a + b"
- using assms by (rule add_less_le_mono)
- then show ?thesis by simp
-qed
-
-lemma add_pos_pos:
- assumes "0 < a" and "0 < b" shows "0 < a + b"
-by (rule add_pos_nonneg) (insert assms, auto)
-
-lemma add_nonneg_pos:
- assumes "0 \<le> a" and "0 < b" shows "0 < a + b"
-proof -
- have "0 + 0 < a + b"
- using assms by (rule add_le_less_mono)
- then show ?thesis by simp
-qed
-
-lemma add_nonneg_nonneg:
- assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"
-proof -
- have "0 + 0 \<le> a + b"
- using assms by (rule add_mono)
- then show ?thesis by simp
-qed
-
-lemma add_neg_nonpos:
- assumes "a < 0" and "b \<le> 0" shows "a + b < 0"
-proof -
- have "a + b < 0 + 0"
- using assms by (rule add_less_le_mono)
- then show ?thesis by simp
-qed
-
-lemma add_neg_neg:
- assumes "a < 0" and "b < 0" shows "a + b < 0"
-by (rule add_neg_nonpos) (insert assms, auto)
-
-lemma add_nonpos_neg:
- assumes "a \<le> 0" and "b < 0" shows "a + b < 0"
-proof -
- have "a + b < 0 + 0"
- using assms by (rule add_le_less_mono)
- then show ?thesis by simp
-qed
-
-lemma add_nonpos_nonpos:
- assumes "a \<le> 0" and "b \<le> 0" shows "a + b \<le> 0"
-proof -
- have "a + b \<le> 0 + 0"
- using assms by (rule add_mono)
- then show ?thesis by simp
-qed
-
-lemmas add_sign_intros =
- add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg
- add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos
-
-lemma add_nonneg_eq_0_iff:
- assumes x: "0 \<le> x" and y: "0 \<le> y"
- shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
-proof (intro iffI conjI)
- have "x = x + 0" by simp
- also have "x + 0 \<le> x + y" using y by (rule add_left_mono)
- also assume "x + y = 0"
- also have "0 \<le> x" using x .
- finally show "x = 0" .
-next
- have "y = 0 + y" by simp
- also have "0 + y \<le> x + y" using x by (rule add_right_mono)
- also assume "x + y = 0"
- also have "0 \<le> y" using y .
- finally show "y = 0" .
-next
- assume "x = 0 \<and> y = 0"
- then show "x + y = 0" by simp
-qed
-
-end
-
-class ordered_ab_group_add =
- ab_group_add + ordered_ab_semigroup_add
-begin
-
-subclass ordered_cancel_ab_semigroup_add ..
-
-subclass ordered_ab_semigroup_add_imp_le
-proof
- fix a b c :: 'a
- assume "c + a \<le> c + b"
- hence "(-c) + (c + a) \<le> (-c) + (c + b)" by (rule add_left_mono)
- hence "((-c) + c) + a \<le> ((-c) + c) + b" by (simp only: add_assoc)
- thus "a \<le> b" by simp
-qed
-
-subclass ordered_comm_monoid_add ..
-
-lemma max_diff_distrib_left:
- shows "max x y - z = max (x - z) (y - z)"
-by (simp add: diff_minus, rule max_add_distrib_left)
-
-lemma min_diff_distrib_left:
- shows "min x y - z = min (x - z) (y - z)"
-by (simp add: diff_minus, rule min_add_distrib_left)
-
-lemma le_imp_neg_le:
- assumes "a \<le> b" shows "-b \<le> -a"
-proof -
- have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono)
- hence "0 \<le> -a+b" by simp
- hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)
- thus ?thesis by (simp add: add_assoc)
-qed
-
-lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
-proof
- assume "- b \<le> - a"
- hence "- (- a) \<le> - (- b)" by (rule le_imp_neg_le)
- thus "a\<le>b" by simp
-next
- assume "a\<le>b"
- thus "-b \<le> -a" by (rule le_imp_neg_le)
-qed
-
-lemma neg_le_0_iff_le [simp]: "- a \<le> 0 \<longleftrightarrow> 0 \<le> a"
-by (subst neg_le_iff_le [symmetric], simp)
-
-lemma neg_0_le_iff_le [simp]: "0 \<le> - a \<longleftrightarrow> a \<le> 0"
-by (subst neg_le_iff_le [symmetric], simp)
-
-lemma neg_less_iff_less [simp]: "- b < - a \<longleftrightarrow> a < b"
-by (force simp add: less_le)
-
-lemma neg_less_0_iff_less [simp]: "- a < 0 \<longleftrightarrow> 0 < a"
-by (subst neg_less_iff_less [symmetric], simp)
-
-lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0"
-by (subst neg_less_iff_less [symmetric], simp)
-
-text{*The next several equations can make the simplifier loop!*}
-
-lemma less_minus_iff: "a < - b \<longleftrightarrow> b < - a"
-proof -
- have "(- (-a) < - b) = (b < - a)" by (rule neg_less_iff_less)
- thus ?thesis by simp
-qed
-
-lemma minus_less_iff: "- a < b \<longleftrightarrow> - b < a"
-proof -
- have "(- a < - (-b)) = (- b < a)" by (rule neg_less_iff_less)
- thus ?thesis by simp
-qed
-
-lemma le_minus_iff: "a \<le> - b \<longleftrightarrow> b \<le> - a"
-proof -
- have mm: "!! a (b::'a). (-(-a)) < -b \<Longrightarrow> -(-b) < -a" by (simp only: minus_less_iff)
- have "(- (- a) <= -b) = (b <= - a)"
- apply (auto simp only: le_less)
- apply (drule mm)
- apply (simp_all)
- apply (drule mm[simplified], assumption)
- done
- then show ?thesis by simp
-qed
-
-lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
-by (auto simp add: le_less minus_less_iff)
-
-lemma less_iff_diff_less_0: "a < b \<longleftrightarrow> a - b < 0"
-proof -
- have "(a < b) = (a + (- b) < b + (-b))"
- by (simp only: add_less_cancel_right)
- also have "... = (a - b < 0)" by (simp add: diff_minus)
- finally show ?thesis .
-qed
-
-lemma diff_less_eq[algebra_simps]: "a - b < c \<longleftrightarrow> a < c + b"
-apply (subst less_iff_diff_less_0 [of a])
-apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
-apply (simp add: diff_minus add_ac)
-done
-
-lemma less_diff_eq[algebra_simps]: "a < c - b \<longleftrightarrow> a + b < c"
-apply (subst less_iff_diff_less_0 [of "plus a b"])
-apply (subst less_iff_diff_less_0 [of a])
-apply (simp add: diff_minus add_ac)
-done
-
-lemma diff_le_eq[algebra_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
-by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
-
-lemma le_diff_eq[algebra_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
-by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
-
-lemma le_iff_diff_le_0: "a \<le> b \<longleftrightarrow> a - b \<le> 0"
-by (simp add: algebra_simps)
-
-text{*Legacy - use @{text algebra_simps} *}
-lemmas group_simps[noatp] = algebra_simps
-
-end
-
-text{*Legacy - use @{text algebra_simps} *}
-lemmas group_simps[noatp] = algebra_simps
-
-class linordered_ab_semigroup_add =
- linorder + ordered_ab_semigroup_add
-
-class linordered_cancel_ab_semigroup_add =
- linorder + ordered_cancel_ab_semigroup_add
-begin
-
-subclass linordered_ab_semigroup_add ..
-
-subclass ordered_ab_semigroup_add_imp_le
-proof
- fix a b c :: 'a
- assume le: "c + a <= c + b"
- show "a <= b"
- proof (rule ccontr)
- assume w: "~ a \<le> b"
- hence "b <= a" by (simp add: linorder_not_le)
- hence le2: "c + b <= c + a" by (rule add_left_mono)
- have "a = b"
- apply (insert le)
- apply (insert le2)
- apply (drule antisym, simp_all)
- done
- with w show False
- by (simp add: linorder_not_le [symmetric])
- qed
-qed
-
-end
-
-class linordered_ab_group_add = linorder + ordered_ab_group_add
-begin
-
-subclass linordered_cancel_ab_semigroup_add ..
-
-lemma neg_less_eq_nonneg [simp]:
- "- a \<le> a \<longleftrightarrow> 0 \<le> a"
-proof
- assume A: "- a \<le> a" show "0 \<le> a"
- proof (rule classical)
- assume "\<not> 0 \<le> a"
- then have "a < 0" by auto
- with A have "- a < 0" by (rule le_less_trans)
- then show ?thesis by auto
- qed
-next
- assume A: "0 \<le> a" show "- a \<le> a"
- proof (rule order_trans)
- show "- a \<le> 0" using A by (simp add: minus_le_iff)
- next
- show "0 \<le> a" using A .
- qed
-qed
-
-lemma neg_less_nonneg [simp]:
- "- a < a \<longleftrightarrow> 0 < a"
-proof
- assume A: "- a < a" show "0 < a"
- proof (rule classical)
- assume "\<not> 0 < a"
- then have "a \<le> 0" by auto
- with A have "- a < 0" by (rule less_le_trans)
- then show ?thesis by auto
- qed
-next
- assume A: "0 < a" show "- a < a"
- proof (rule less_trans)
- show "- a < 0" using A by (simp add: minus_le_iff)
- next
- show "0 < a" using A .
- qed
-qed
-
-lemma less_eq_neg_nonpos [simp]:
- "a \<le> - a \<longleftrightarrow> a \<le> 0"
-proof
- assume A: "a \<le> - a" show "a \<le> 0"
- proof (rule classical)
- assume "\<not> a \<le> 0"
- then have "0 < a" by auto
- then have "0 < - a" using A by (rule less_le_trans)
- then show ?thesis by auto
- qed
-next
- assume A: "a \<le> 0" show "a \<le> - a"
- proof (rule order_trans)
- show "0 \<le> - a" using A by (simp add: minus_le_iff)
- next
- show "a \<le> 0" using A .
- qed
-qed
-
-lemma equal_neg_zero [simp]:
- "a = - a \<longleftrightarrow> a = 0"
-proof
- assume "a = 0" then show "a = - a" by simp
-next
- assume A: "a = - a" show "a = 0"
- proof (cases "0 \<le> a")
- case True with A have "0 \<le> - a" by auto
- with le_minus_iff have "a \<le> 0" by simp
- with True show ?thesis by (auto intro: order_trans)
- next
- case False then have B: "a \<le> 0" by auto
- with A have "- a \<le> 0" by auto
- with B show ?thesis by (auto intro: order_trans)
- qed
-qed
-
-lemma neg_equal_zero [simp]:
- "- a = a \<longleftrightarrow> a = 0"
- by (auto dest: sym)
-
-lemma double_zero [simp]:
- "a + a = 0 \<longleftrightarrow> a = 0"
-proof
- assume assm: "a + a = 0"
- then have a: "- a = a" by (rule minus_unique)
- then show "a = 0" by (simp add: neg_equal_zero)
-qed simp
-
-lemma double_zero_sym [simp]:
- "0 = a + a \<longleftrightarrow> a = 0"
- by (rule, drule sym) simp_all
-
-lemma zero_less_double_add_iff_zero_less_single_add [simp]:
- "0 < a + a \<longleftrightarrow> 0 < a"
-proof
- assume "0 < a + a"
- then have "0 - a < a" by (simp only: diff_less_eq)
- then have "- a < a" by simp
- then show "0 < a" by (simp add: neg_less_nonneg)
-next
- assume "0 < a"
- with this have "0 + 0 < a + a"
- by (rule add_strict_mono)
- then show "0 < a + a" by simp
-qed
-
-lemma zero_le_double_add_iff_zero_le_single_add [simp]:
- "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
- by (auto simp add: le_less)
-
-lemma double_add_less_zero_iff_single_add_less_zero [simp]:
- "a + a < 0 \<longleftrightarrow> a < 0"
-proof -
- have "\<not> a + a < 0 \<longleftrightarrow> \<not> a < 0"
- by (simp add: not_less)
- then show ?thesis by simp
-qed
-
-lemma double_add_le_zero_iff_single_add_le_zero [simp]:
- "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
-proof -
- have "\<not> a + a \<le> 0 \<longleftrightarrow> \<not> a \<le> 0"
- by (simp add: not_le)
- then show ?thesis by simp
-qed
-
-lemma le_minus_self_iff:
- "a \<le> - a \<longleftrightarrow> a \<le> 0"
-proof -
- from add_le_cancel_left [of "- a" "a + a" 0]
- have "a \<le> - a \<longleftrightarrow> a + a \<le> 0"
- by (simp add: add_assoc [symmetric])
- thus ?thesis by simp
-qed
-
-lemma minus_le_self_iff:
- "- a \<le> a \<longleftrightarrow> 0 \<le> a"
-proof -
- from add_le_cancel_left [of "- a" 0 "a + a"]
- have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"
- by (simp add: add_assoc [symmetric])
- thus ?thesis by simp
-qed
-
-lemma minus_max_eq_min:
- "- max x y = min (-x) (-y)"
- by (auto simp add: max_def min_def)
-
-lemma minus_min_eq_max:
- "- min x y = max (-x) (-y)"
- by (auto simp add: max_def min_def)
-
-end
-
--- {* FIXME localize the following *}
-
-lemma add_increasing:
- fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
- shows "[|0\<le>a; b\<le>c|] ==> b \<le> a + c"
-by (insert add_mono [of 0 a b c], simp)
-
-lemma add_increasing2:
- fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
- shows "[|0\<le>c; b\<le>a|] ==> b \<le> a + c"
-by (simp add:add_increasing add_commute[of a])
-
-lemma add_strict_increasing:
- fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
- shows "[|0<a; b\<le>c|] ==> b < a + c"
-by (insert add_less_le_mono [of 0 a b c], simp)
-
-lemma add_strict_increasing2:
- fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
- shows "[|0\<le>a; b<c|] ==> b < a + c"
-by (insert add_le_less_mono [of 0 a b c], simp)
-
-
-class ordered_ab_group_add_abs = ordered_ab_group_add + abs +
- assumes abs_ge_zero [simp]: "\<bar>a\<bar> \<ge> 0"
- and abs_ge_self: "a \<le> \<bar>a\<bar>"
- and abs_leI: "a \<le> b \<Longrightarrow> - a \<le> b \<Longrightarrow> \<bar>a\<bar> \<le> b"
- and abs_minus_cancel [simp]: "\<bar>-a\<bar> = \<bar>a\<bar>"
- and abs_triangle_ineq: "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
-begin
-
-lemma abs_minus_le_zero: "- \<bar>a\<bar> \<le> 0"
- unfolding neg_le_0_iff_le by simp
-
-lemma abs_of_nonneg [simp]:
- assumes nonneg: "0 \<le> a" shows "\<bar>a\<bar> = a"
-proof (rule antisym)
- from nonneg le_imp_neg_le have "- a \<le> 0" by simp
- from this nonneg have "- a \<le> a" by (rule order_trans)
- then show "\<bar>a\<bar> \<le> a" by (auto intro: abs_leI)
-qed (rule abs_ge_self)
-
-lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"
-by (rule antisym)
- (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"])
-
-lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"
-proof -
- have "\<bar>a\<bar> = 0 \<Longrightarrow> a = 0"
- proof (rule antisym)
- assume zero: "\<bar>a\<bar> = 0"
- with abs_ge_self show "a \<le> 0" by auto
- from zero have "\<bar>-a\<bar> = 0" by simp
- with abs_ge_self [of "uminus a"] have "- a \<le> 0" by auto
- with neg_le_0_iff_le show "0 \<le> a" by auto
- qed
- then show ?thesis by auto
-qed
-
-lemma abs_zero [simp]: "\<bar>0\<bar> = 0"
-by simp
-
-lemma abs_0_eq [simp, noatp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
-proof -
- have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac)
- thus ?thesis by simp
-qed
-
-lemma abs_le_zero_iff [simp]: "\<bar>a\<bar> \<le> 0 \<longleftrightarrow> a = 0"
-proof
- assume "\<bar>a\<bar> \<le> 0"
- then have "\<bar>a\<bar> = 0" by (rule antisym) simp
- thus "a = 0" by simp
-next
- assume "a = 0"
- thus "\<bar>a\<bar> \<le> 0" by simp
-qed
-
-lemma zero_less_abs_iff [simp]: "0 < \<bar>a\<bar> \<longleftrightarrow> a \<noteq> 0"
-by (simp add: less_le)
-
-lemma abs_not_less_zero [simp]: "\<not> \<bar>a\<bar> < 0"
-proof -
- have a: "\<And>x y. x \<le> y \<Longrightarrow> \<not> y < x" by auto
- show ?thesis by (simp add: a)
-qed
-
-lemma abs_ge_minus_self: "- a \<le> \<bar>a\<bar>"
-proof -
- have "- a \<le> \<bar>-a\<bar>" by (rule abs_ge_self)
- then show ?thesis by simp
-qed
-
-lemma abs_minus_commute:
- "\<bar>a - b\<bar> = \<bar>b - a\<bar>"
-proof -
- have "\<bar>a - b\<bar> = \<bar>- (a - b)\<bar>" by (simp only: abs_minus_cancel)
- also have "... = \<bar>b - a\<bar>" by simp
- finally show ?thesis .
-qed
-
-lemma abs_of_pos: "0 < a \<Longrightarrow> \<bar>a\<bar> = a"
-by (rule abs_of_nonneg, rule less_imp_le)
-
-lemma abs_of_nonpos [simp]:
- assumes "a \<le> 0" shows "\<bar>a\<bar> = - a"
-proof -
- let ?b = "- a"
- have "- ?b \<le> 0 \<Longrightarrow> \<bar>- ?b\<bar> = - (- ?b)"
- unfolding abs_minus_cancel [of "?b"]
- unfolding neg_le_0_iff_le [of "?b"]
- unfolding minus_minus by (erule abs_of_nonneg)
- then show ?thesis using assms by auto
-qed
-
-lemma abs_of_neg: "a < 0 \<Longrightarrow> \<bar>a\<bar> = - a"
-by (rule abs_of_nonpos, rule less_imp_le)
-
-lemma abs_le_D1: "\<bar>a\<bar> \<le> b \<Longrightarrow> a \<le> b"
-by (insert abs_ge_self, blast intro: order_trans)
-
-lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b"
-by (insert abs_le_D1 [of "uminus a"], simp)
-
-lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b"
-by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
-
-lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
- apply (simp add: algebra_simps)
- apply (subgoal_tac "abs a = abs (plus b (minus a b))")
- apply (erule ssubst)
- apply (rule abs_triangle_ineq)
- apply (rule arg_cong[of _ _ abs])
- apply (simp add: algebra_simps)
-done
-
-lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>"
- apply (subst abs_le_iff)
- apply auto
- apply (rule abs_triangle_ineq2)
- apply (subst abs_minus_commute)
- apply (rule abs_triangle_ineq2)
-done
-
-lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
-proof -
- have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl)
- also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq)
- finally show ?thesis by simp
-qed
-
-lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
-proof -
- have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
- also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
- finally show ?thesis .
-qed
-
-lemma abs_add_abs [simp]:
- "\<bar>\<bar>a\<bar> + \<bar>b\<bar>\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" (is "?L = ?R")
-proof (rule antisym)
- show "?L \<ge> ?R" by(rule abs_ge_self)
-next
- have "?L \<le> \<bar>\<bar>a\<bar>\<bar> + \<bar>\<bar>b\<bar>\<bar>" by(rule abs_triangle_ineq)
- also have "\<dots> = ?R" by simp
- finally show "?L \<le> ?R" .
-qed
-
-end
-
-text {* Needed for abelian cancellation simprocs: *}
-
-lemma add_cancel_21: "((x::'a::ab_group_add) + (y + z) = y + u) = (x + z = u)"
-apply (subst add_left_commute)
-apply (subst add_left_cancel)
-apply simp
-done
-
-lemma add_cancel_end: "(x + (y + z) = y) = (x = - (z::'a::ab_group_add))"
-apply (subst add_cancel_21[of _ _ _ 0, simplified])
-apply (simp add: add_right_cancel[symmetric, of "x" "-z" "z", simplified])
-done
-
-lemma less_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \<Longrightarrow> (x < y) = (x' < y')"
-by (simp add: less_iff_diff_less_0[of x y] less_iff_diff_less_0[of x' y'])
-
-lemma le_eqI: "(x::'a::ordered_ab_group_add) - y = x' - y' \<Longrightarrow> (y <= x) = (y' <= x')"
-apply (simp add: le_iff_diff_le_0[of y x] le_iff_diff_le_0[of y' x'])
-apply (simp add: neg_le_iff_le[symmetric, of "y-x" 0] neg_le_iff_le[symmetric, of "y'-x'" 0])
-done
-
-lemma eq_eqI: "(x::'a::ab_group_add) - y = x' - y' \<Longrightarrow> (x = y) = (x' = y')"
-by (simp only: eq_iff_diff_eq_0[of x y] eq_iff_diff_eq_0[of x' y'])
-
-lemma diff_def: "(x::'a::ab_group_add) - y == x + (-y)"
-by (simp add: diff_minus)
-
-lemma le_add_right_mono:
- assumes
- "a <= b + (c::'a::ordered_ab_group_add)"
- "c <= d"
- shows "a <= b + d"
- apply (rule_tac order_trans[where y = "b+c"])
- apply (simp_all add: prems)
- done
-
-
-subsection {* Tools setup *}
-
-lemma add_mono_thms_linordered_semiring [noatp]:
- fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
- shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
- and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
- and "i \<le> j \<and> k = l \<Longrightarrow> i + k \<le> j + l"
- and "i = j \<and> k = l \<Longrightarrow> i + k = j + l"
-by (rule add_mono, clarify+)+
-
-lemma add_mono_thms_linordered_field [noatp]:
- fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"
- shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"
- and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"
- and "i < j \<and> k \<le> l \<Longrightarrow> i + k < j + l"
- and "i \<le> j \<and> k < l \<Longrightarrow> i + k < j + l"
- and "i < j \<and> k < l \<Longrightarrow> i + k < j + l"
-by (auto intro: add_strict_right_mono add_strict_left_mono
- add_less_le_mono add_le_less_mono add_strict_mono)
-
-text{*Simplification of @{term "x-y < 0"}, etc.*}
-lemmas diff_less_0_iff_less [simp, noatp] = less_iff_diff_less_0 [symmetric]
-lemmas diff_le_0_iff_le [simp, noatp] = le_iff_diff_le_0 [symmetric]
-
-ML {*
-structure ab_group_add_cancel = Abel_Cancel
-(
-
-(* term order for abelian groups *)
-
-fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a')
- [@{const_name Algebras.zero}, @{const_name Algebras.plus},
- @{const_name Algebras.uminus}, @{const_name Algebras.minus}]
- | agrp_ord _ = ~1;
-
-fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
-
-local
- val ac1 = mk_meta_eq @{thm add_assoc};
- val ac2 = mk_meta_eq @{thm add_commute};
- val ac3 = mk_meta_eq @{thm add_left_commute};
- fun solve_add_ac thy _ (_ $ (Const (@{const_name Algebras.plus},_) $ _ $ _) $ _) =
- SOME ac1
- | solve_add_ac thy _ (_ $ x $ (Const (@{const_name Algebras.plus},_) $ y $ z)) =
- if termless_agrp (y, x) then SOME ac3 else NONE
- | solve_add_ac thy _ (_ $ x $ y) =
- if termless_agrp (y, x) then SOME ac2 else NONE
- | solve_add_ac thy _ _ = NONE
-in
- val add_ac_proc = Simplifier.simproc @{theory}
- "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac;
-end;
-
-val eq_reflection = @{thm eq_reflection};
-
-val T = @{typ "'a::ab_group_add"};
-
-val cancel_ss = HOL_basic_ss settermless termless_agrp
- addsimprocs [add_ac_proc] addsimps
- [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_def},
- @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero},
- @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel},
- @{thm minus_add_cancel}];
-
-val sum_pats = [@{cterm "x + y::'a::ab_group_add"}, @{cterm "x - y::'a::ab_group_add"}];
-
-val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}];
-
-val dest_eqI =
- fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of;
-
-);
-*}
-
-ML {*
- Addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv];
-*}
-
-code_modulename SML
- OrderedGroup Arith
-
-code_modulename OCaml
- OrderedGroup Arith
-
-code_modulename Haskell
- OrderedGroup Arith
-
-end
--- a/src/HOL/PReal.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/PReal.thy Mon Feb 08 17:13:45 2010 +0100
@@ -12,7 +12,7 @@
imports Rational
begin
-text{*Could be generalized and moved to @{text Ring_and_Field}*}
+text{*Could be generalized and moved to @{text Groups}*}
lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
by (rule_tac x="b-a" in exI, simp)
--- a/src/HOL/Parity.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Parity.thy Mon Feb 08 17:13:45 2010 +0100
@@ -218,7 +218,7 @@
done
lemma zero_le_even_power: "even n ==>
- 0 <= (x::'a::{linlinordered_ring_strict,monoid_mult}) ^ n"
+ 0 <= (x::'a::{linordered_ring_strict,monoid_mult}) ^ n"
apply (simp add: even_nat_equiv_def2)
apply (erule exE)
apply (erule ssubst)
--- a/src/HOL/Presburger.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Presburger.thy Mon Feb 08 17:13:45 2010 +0100
@@ -30,8 +30,8 @@
"\<exists>(z ::'a::{linorder}).\<forall>x<z.(x \<le> t) = True"
"\<exists>(z ::'a::{linorder}).\<forall>x<z.(x > t) = False"
"\<exists>(z ::'a::{linorder}).\<forall>x<z.(x \<ge> t) = False"
- "\<exists>z.\<forall>(x::'a::{linorder,plus,Ring_and_Field.dvd})<z. (d dvd x + s) = (d dvd x + s)"
- "\<exists>z.\<forall>(x::'a::{linorder,plus,Ring_and_Field.dvd})<z. (\<not> d dvd x + s) = (\<not> d dvd x + s)"
+ "\<exists>z.\<forall>(x::'a::{linorder,plus,Rings.dvd})<z. (d dvd x + s) = (d dvd x + s)"
+ "\<exists>z.\<forall>(x::'a::{linorder,plus,Rings.dvd})<z. (\<not> d dvd x + s) = (\<not> d dvd x + s)"
"\<exists>z.\<forall>x<z. F = F"
by ((erule exE, erule exE,rule_tac x="min z za" in exI,simp)+, (rule_tac x="t" in exI,fastsimp)+) simp_all
@@ -46,8 +46,8 @@
"\<exists>(z ::'a::{linorder}).\<forall>x>z.(x \<le> t) = False"
"\<exists>(z ::'a::{linorder}).\<forall>x>z.(x > t) = True"
"\<exists>(z ::'a::{linorder}).\<forall>x>z.(x \<ge> t) = True"
- "\<exists>z.\<forall>(x::'a::{linorder,plus,Ring_and_Field.dvd})>z. (d dvd x + s) = (d dvd x + s)"
- "\<exists>z.\<forall>(x::'a::{linorder,plus,Ring_and_Field.dvd})>z. (\<not> d dvd x + s) = (\<not> d dvd x + s)"
+ "\<exists>z.\<forall>(x::'a::{linorder,plus,Rings.dvd})>z. (d dvd x + s) = (d dvd x + s)"
+ "\<exists>z.\<forall>(x::'a::{linorder,plus,Rings.dvd})>z. (\<not> d dvd x + s) = (\<not> d dvd x + s)"
"\<exists>z.\<forall>x>z. F = F"
by ((erule exE, erule exE,rule_tac x="max z za" in exI,simp)+,(rule_tac x="t" in exI,fastsimp)+) simp_all
@@ -56,8 +56,8 @@
\<Longrightarrow> \<forall>x k. (P x \<and> Q x) = (P (x - k*D) \<and> Q (x - k*D))"
"\<lbrakk>\<forall>x k. P x = P (x - k*D); \<forall>x k. Q x = Q (x - k*D)\<rbrakk>
\<Longrightarrow> \<forall>x k. (P x \<or> Q x) = (P (x - k*D) \<or> Q (x - k*D))"
- "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (d dvd x + t) = (d dvd (x - k*D) + t)"
- "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (\<not>d dvd x + t) = (\<not>d dvd (x - k*D) + t)"
+ "(d::'a::{comm_ring,Rings.dvd}) dvd D \<Longrightarrow> \<forall>x k. (d dvd x + t) = (d dvd (x - k*D) + t)"
+ "(d::'a::{comm_ring,Rings.dvd}) dvd D \<Longrightarrow> \<forall>x k. (\<not>d dvd x + t) = (\<not>d dvd (x - k*D) + t)"
"\<forall>x k. F = F"
apply (auto elim!: dvdE simp add: algebra_simps)
unfolding mult_assoc [symmetric] left_distrib [symmetric] left_diff_distrib [symmetric]
@@ -243,7 +243,7 @@
{fix x
have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast
also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)" using minus[THEN spec, of "x - i * d"]
- by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric])
+ by (simp add: algebra_simps)
ultimately have "P x \<longrightarrow> P(x - (i + 1) * d)" by blast}
thus ?case ..
qed
@@ -360,7 +360,7 @@
apply(fastsimp)
done
-theorem unity_coeff_ex: "(\<exists>(x::'a::{semiring_0,Ring_and_Field.dvd}). P (l * x)) \<equiv> (\<exists>x. l dvd (x + 0) \<and> P x)"
+theorem unity_coeff_ex: "(\<exists>(x::'a::{semiring_0,Rings.dvd}). P (l * x)) \<equiv> (\<exists>x. l dvd (x + 0) \<and> P x)"
apply (rule eq_reflection [symmetric])
apply (rule iffI)
defer
--- a/src/HOL/Probability/Borel.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Probability/Borel.thy Mon Feb 08 17:13:45 2010 +0100
@@ -355,7 +355,7 @@
borel_measurable_add_borel_measurable f g)
have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) =
(\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))"
- by (simp add: Ring_and_Field.minus_divide_right)
+ by (simp add: minus_divide_right)
also have "... \<in> borel_measurable M"
by (fast intro: affine_borel_measurable borel_measurable_square
borel_measurable_add_borel_measurable
--- a/src/HOL/RealDef.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/RealDef.thy Mon Feb 08 17:13:45 2010 +0100
@@ -521,7 +521,7 @@
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
by (force elim: order_less_asym
- simp add: Ring_and_Field.mult_less_cancel_right)
+ simp add: mult_less_cancel_right)
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
apply (simp add: mult_le_cancel_right)
@@ -1015,7 +1015,7 @@
done
-text{*Similar results are proved in @{text Ring_and_Field}*}
+text{*Similar results are proved in @{text Fields}*}
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
by auto
@@ -1030,7 +1030,7 @@
(* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *)
lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
-by (force simp add: OrderedGroup.abs_le_iff)
+by (force simp add: abs_le_iff)
lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
by (simp add: abs_if)
--- a/src/HOL/Ring_and_Field.thy Mon Feb 08 15:25:00 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2226 +0,0 @@
-(* Title: HOL/Ring_and_Field.thy
- Author: Gertrud Bauer
- Author: Steven Obua
- Author: Tobias Nipkow
- Author: Lawrence C Paulson
- Author: Markus Wenzel
- Author: Jeremy Avigad
-*)
-
-header {* (Ordered) Rings and Fields *}
-
-theory Ring_and_Field
-imports OrderedGroup
-begin
-
-text {*
- The theory of partially ordered rings is taken from the books:
- \begin{itemize}
- \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
- \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
- \end{itemize}
- Most of the used notions can also be looked up in
- \begin{itemize}
- \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
- \item \emph{Algebra I} by van der Waerden, Springer.
- \end{itemize}
-*}
-
-class semiring = ab_semigroup_add + semigroup_mult +
- assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
- assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
-begin
-
-text{*For the @{text combine_numerals} simproc*}
-lemma combine_common_factor:
- "a * e + (b * e + c) = (a + b) * e + c"
-by (simp add: left_distrib add_ac)
-
-end
-
-class mult_zero = times + zero +
- assumes mult_zero_left [simp]: "0 * a = 0"
- assumes mult_zero_right [simp]: "a * 0 = 0"
-
-class semiring_0 = semiring + comm_monoid_add + mult_zero
-
-class semiring_0_cancel = semiring + cancel_comm_monoid_add
-begin
-
-subclass semiring_0
-proof
- fix a :: 'a
- have "0 * a + 0 * a = 0 * a + 0" by (simp add: left_distrib [symmetric])
- thus "0 * a = 0" by (simp only: add_left_cancel)
-next
- fix a :: 'a
- have "a * 0 + a * 0 = a * 0 + 0" by (simp add: right_distrib [symmetric])
- thus "a * 0 = 0" by (simp only: add_left_cancel)
-qed
-
-end
-
-class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
- assumes distrib: "(a + b) * c = a * c + b * c"
-begin
-
-subclass semiring
-proof
- fix a b c :: 'a
- show "(a + b) * c = a * c + b * c" by (simp add: distrib)
- have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
- also have "... = b * a + c * a" by (simp only: distrib)
- also have "... = a * b + a * c" by (simp add: mult_ac)
- finally show "a * (b + c) = a * b + a * c" by blast
-qed
-
-end
-
-class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
-begin
-
-subclass semiring_0 ..
-
-end
-
-class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add
-begin
-
-subclass semiring_0_cancel ..
-
-subclass comm_semiring_0 ..
-
-end
-
-class zero_neq_one = zero + one +
- assumes zero_neq_one [simp]: "0 \<noteq> 1"
-begin
-
-lemma one_neq_zero [simp]: "1 \<noteq> 0"
-by (rule not_sym) (rule zero_neq_one)
-
-end
-
-class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
-
-text {* Abstract divisibility *}
-
-class dvd = times
-begin
-
-definition dvd :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "dvd" 50) where
- [code del]: "b dvd a \<longleftrightarrow> (\<exists>k. a = b * k)"
-
-lemma dvdI [intro?]: "a = b * k \<Longrightarrow> b dvd a"
- unfolding dvd_def ..
-
-lemma dvdE [elim?]: "b dvd a \<Longrightarrow> (\<And>k. a = b * k \<Longrightarrow> P) \<Longrightarrow> P"
- unfolding dvd_def by blast
-
-end
-
-class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult + dvd
- (*previously almost_semiring*)
-begin
-
-subclass semiring_1 ..
-
-lemma dvd_refl[simp]: "a dvd a"
-proof
- show "a = a * 1" by simp
-qed
-
-lemma dvd_trans:
- assumes "a dvd b" and "b dvd c"
- shows "a dvd c"
-proof -
- from assms obtain v where "b = a * v" by (auto elim!: dvdE)
- moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE)
- ultimately have "c = a * (v * w)" by (simp add: mult_assoc)
- then show ?thesis ..
-qed
-
-lemma dvd_0_left_iff [noatp, simp]: "0 dvd a \<longleftrightarrow> a = 0"
-by (auto intro: dvd_refl elim!: dvdE)
-
-lemma dvd_0_right [iff]: "a dvd 0"
-proof
- show "0 = a * 0" by simp
-qed
-
-lemma one_dvd [simp]: "1 dvd a"
-by (auto intro!: dvdI)
-
-lemma dvd_mult[simp]: "a dvd c \<Longrightarrow> a dvd (b * c)"
-by (auto intro!: mult_left_commute dvdI elim!: dvdE)
-
-lemma dvd_mult2[simp]: "a dvd b \<Longrightarrow> a dvd (b * c)"
- apply (subst mult_commute)
- apply (erule dvd_mult)
- done
-
-lemma dvd_triv_right [simp]: "a dvd b * a"
-by (rule dvd_mult) (rule dvd_refl)
-
-lemma dvd_triv_left [simp]: "a dvd a * b"
-by (rule dvd_mult2) (rule dvd_refl)
-
-lemma mult_dvd_mono:
- assumes "a dvd b"
- and "c dvd d"
- shows "a * c dvd b * d"
-proof -
- from `a dvd b` obtain b' where "b = a * b'" ..
- moreover from `c dvd d` obtain d' where "d = c * d'" ..
- ultimately have "b * d = (a * c) * (b' * d')" by (simp add: mult_ac)
- then show ?thesis ..
-qed
-
-lemma dvd_mult_left: "a * b dvd c \<Longrightarrow> a dvd c"
-by (simp add: dvd_def mult_assoc, blast)
-
-lemma dvd_mult_right: "a * b dvd c \<Longrightarrow> b dvd c"
- unfolding mult_ac [of a] by (rule dvd_mult_left)
-
-lemma dvd_0_left: "0 dvd a \<Longrightarrow> a = 0"
-by simp
-
-lemma dvd_add[simp]:
- assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)"
-proof -
- from `a dvd b` obtain b' where "b = a * b'" ..
- moreover from `a dvd c` obtain c' where "c = a * c'" ..
- ultimately have "b + c = a * (b' + c')" by (simp add: right_distrib)
- then show ?thesis ..
-qed
-
-end
-
-
-class no_zero_divisors = zero + times +
- assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
-
-class semiring_1_cancel = semiring + cancel_comm_monoid_add
- + zero_neq_one + monoid_mult
-begin
-
-subclass semiring_0_cancel ..
-
-subclass semiring_1 ..
-
-end
-
-class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add
- + zero_neq_one + comm_monoid_mult
-begin
-
-subclass semiring_1_cancel ..
-subclass comm_semiring_0_cancel ..
-subclass comm_semiring_1 ..
-
-end
-
-class ring = semiring + ab_group_add
-begin
-
-subclass semiring_0_cancel ..
-
-text {* Distribution rules *}
-
-lemma minus_mult_left: "- (a * b) = - a * b"
-by (rule minus_unique) (simp add: left_distrib [symmetric])
-
-lemma minus_mult_right: "- (a * b) = a * - b"
-by (rule minus_unique) (simp add: right_distrib [symmetric])
-
-text{*Extract signs from products*}
-lemmas mult_minus_left [simp, noatp] = minus_mult_left [symmetric]
-lemmas mult_minus_right [simp,noatp] = minus_mult_right [symmetric]
-
-lemma minus_mult_minus [simp]: "- a * - b = a * b"
-by simp
-
-lemma minus_mult_commute: "- a * b = a * - b"
-by simp
-
-lemma right_diff_distrib[algebra_simps]: "a * (b - c) = a * b - a * c"
-by (simp add: right_distrib diff_minus)
-
-lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c"
-by (simp add: left_distrib diff_minus)
-
-lemmas ring_distribs[noatp] =
- right_distrib left_distrib left_diff_distrib right_diff_distrib
-
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[noatp] = algebra_simps
-
-lemma eq_add_iff1:
- "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
-by (simp add: algebra_simps)
-
-lemma eq_add_iff2:
- "a * e + c = b * e + d \<longleftrightarrow> c = (b - a) * e + d"
-by (simp add: algebra_simps)
-
-end
-
-lemmas ring_distribs[noatp] =
- right_distrib left_distrib left_diff_distrib right_diff_distrib
-
-class comm_ring = comm_semiring + ab_group_add
-begin
-
-subclass ring ..
-subclass comm_semiring_0_cancel ..
-
-end
-
-class ring_1 = ring + zero_neq_one + monoid_mult
-begin
-
-subclass semiring_1_cancel ..
-
-end
-
-class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
- (*previously ring*)
-begin
-
-subclass ring_1 ..
-subclass comm_semiring_1_cancel ..
-
-lemma dvd_minus_iff [simp]: "x dvd - y \<longleftrightarrow> x dvd y"
-proof
- assume "x dvd - y"
- then have "x dvd - 1 * - y" by (rule dvd_mult)
- then show "x dvd y" by simp
-next
- assume "x dvd y"
- then have "x dvd - 1 * y" by (rule dvd_mult)
- then show "x dvd - y" by simp
-qed
-
-lemma minus_dvd_iff [simp]: "- x dvd y \<longleftrightarrow> x dvd y"
-proof
- assume "- x dvd y"
- then obtain k where "y = - x * k" ..
- then have "y = x * - k" by simp
- then show "x dvd y" ..
-next
- assume "x dvd y"
- then obtain k where "y = x * k" ..
- then have "y = - x * - k" by simp
- then show "- x dvd y" ..
-qed
-
-lemma dvd_diff[simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
-by (simp add: diff_minus dvd_minus_iff)
-
-end
-
-class ring_no_zero_divisors = ring + no_zero_divisors
-begin
-
-lemma mult_eq_0_iff [simp]:
- shows "a * b = 0 \<longleftrightarrow> (a = 0 \<or> b = 0)"
-proof (cases "a = 0 \<or> b = 0")
- case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
- then show ?thesis using no_zero_divisors by simp
-next
- case True then show ?thesis by auto
-qed
-
-text{*Cancellation of equalities with a common factor*}
-lemma mult_cancel_right [simp, noatp]:
- "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
-proof -
- have "(a * c = b * c) = ((a - b) * c = 0)"
- by (simp add: algebra_simps right_minus_eq)
- thus ?thesis by (simp add: disj_commute right_minus_eq)
-qed
-
-lemma mult_cancel_left [simp, noatp]:
- "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
-proof -
- have "(c * a = c * b) = (c * (a - b) = 0)"
- by (simp add: algebra_simps right_minus_eq)
- thus ?thesis by (simp add: right_minus_eq)
-qed
-
-end
-
-class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
-begin
-
-lemma mult_cancel_right1 [simp]:
- "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
-by (insert mult_cancel_right [of 1 c b], force)
-
-lemma mult_cancel_right2 [simp]:
- "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
-by (insert mult_cancel_right [of a c 1], simp)
-
-lemma mult_cancel_left1 [simp]:
- "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
-by (insert mult_cancel_left [of c 1 b], force)
-
-lemma mult_cancel_left2 [simp]:
- "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
-by (insert mult_cancel_left [of c a 1], simp)
-
-end
-
-class idom = comm_ring_1 + no_zero_divisors
-begin
-
-subclass ring_1_no_zero_divisors ..
-
-lemma square_eq_iff: "a * a = b * b \<longleftrightarrow> (a = b \<or> a = - b)"
-proof
- assume "a * a = b * b"
- then have "(a - b) * (a + b) = 0"
- by (simp add: algebra_simps)
- then show "a = b \<or> a = - b"
- by (simp add: right_minus_eq eq_neg_iff_add_eq_0)
-next
- assume "a = b \<or> a = - b"
- then show "a * a = b * b" by auto
-qed
-
-lemma dvd_mult_cancel_right [simp]:
- "a * c dvd b * c \<longleftrightarrow> c = 0 \<or> a dvd b"
-proof -
- have "a * c dvd b * c \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
- unfolding dvd_def by (simp add: mult_ac)
- also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
- unfolding dvd_def by simp
- finally show ?thesis .
-qed
-
-lemma dvd_mult_cancel_left [simp]:
- "c * a dvd c * b \<longleftrightarrow> c = 0 \<or> a dvd b"
-proof -
- have "c * a dvd c * b \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
- unfolding dvd_def by (simp add: mult_ac)
- also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
- unfolding dvd_def by simp
- finally show ?thesis .
-qed
-
-end
-
-class division_ring = ring_1 + inverse +
- assumes left_inverse [simp]: "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
- assumes right_inverse [simp]: "a \<noteq> 0 \<Longrightarrow> a * inverse a = 1"
-begin
-
-subclass ring_1_no_zero_divisors
-proof
- fix a b :: 'a
- assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
- show "a * b \<noteq> 0"
- proof
- assume ab: "a * b = 0"
- hence "0 = inverse a * (a * b) * inverse b" by simp
- also have "\<dots> = (inverse a * a) * (b * inverse b)"
- by (simp only: mult_assoc)
- also have "\<dots> = 1" using a b by simp
- finally show False by simp
- qed
-qed
-
-lemma nonzero_imp_inverse_nonzero:
- "a \<noteq> 0 \<Longrightarrow> inverse a \<noteq> 0"
-proof
- assume ianz: "inverse a = 0"
- assume "a \<noteq> 0"
- hence "1 = a * inverse a" by simp
- also have "... = 0" by (simp add: ianz)
- finally have "1 = 0" .
- thus False by (simp add: eq_commute)
-qed
-
-lemma inverse_zero_imp_zero:
- "inverse a = 0 \<Longrightarrow> a = 0"
-apply (rule classical)
-apply (drule nonzero_imp_inverse_nonzero)
-apply auto
-done
-
-lemma inverse_unique:
- assumes ab: "a * b = 1"
- shows "inverse a = b"
-proof -
- have "a \<noteq> 0" using ab by (cases "a = 0") simp_all
- moreover have "inverse a * (a * b) = inverse a" by (simp add: ab)
- ultimately show ?thesis by (simp add: mult_assoc [symmetric])
-qed
-
-lemma nonzero_inverse_minus_eq:
- "a \<noteq> 0 \<Longrightarrow> inverse (- a) = - inverse a"
-by (rule inverse_unique) simp
-
-lemma nonzero_inverse_inverse_eq:
- "a \<noteq> 0 \<Longrightarrow> inverse (inverse a) = a"
-by (rule inverse_unique) simp
-
-lemma nonzero_inverse_eq_imp_eq:
- assumes "inverse a = inverse b" and "a \<noteq> 0" and "b \<noteq> 0"
- shows "a = b"
-proof -
- from `inverse a = inverse b`
- have "inverse (inverse a) = inverse (inverse b)" by (rule arg_cong)
- with `a \<noteq> 0` and `b \<noteq> 0` show "a = b"
- by (simp add: nonzero_inverse_inverse_eq)
-qed
-
-lemma inverse_1 [simp]: "inverse 1 = 1"
-by (rule inverse_unique) simp
-
-lemma nonzero_inverse_mult_distrib:
- assumes "a \<noteq> 0" and "b \<noteq> 0"
- shows "inverse (a * b) = inverse b * inverse a"
-proof -
- have "a * (b * inverse b) * inverse a = 1" using assms by simp
- hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult_assoc)
- thus ?thesis by (rule inverse_unique)
-qed
-
-lemma division_ring_inverse_add:
- "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a + inverse b = inverse a * (a + b) * inverse b"
-by (simp add: algebra_simps)
-
-lemma division_ring_inverse_diff:
- "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a - inverse b = inverse a * (b - a) * inverse b"
-by (simp add: algebra_simps)
-
-end
-
-class field = comm_ring_1 + inverse +
- assumes field_inverse: "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
- assumes divide_inverse: "a / b = a * inverse b"
-begin
-
-subclass division_ring
-proof
- fix a :: 'a
- assume "a \<noteq> 0"
- thus "inverse a * a = 1" by (rule field_inverse)
- thus "a * inverse a = 1" by (simp only: mult_commute)
-qed
-
-subclass idom ..
-
-lemma right_inverse_eq: "b \<noteq> 0 \<Longrightarrow> a / b = 1 \<longleftrightarrow> a = b"
-proof
- assume neq: "b \<noteq> 0"
- {
- hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
- also assume "a / b = 1"
- finally show "a = b" by simp
- next
- assume "a = b"
- with neq show "a / b = 1" by (simp add: divide_inverse)
- }
-qed
-
-lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
-by (simp add: divide_inverse)
-
-lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
-by (simp add: divide_inverse)
-
-lemma divide_zero_left [simp]: "0 / a = 0"
-by (simp add: divide_inverse)
-
-lemma inverse_eq_divide: "inverse a = 1 / a"
-by (simp add: divide_inverse)
-
-lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
-by (simp add: divide_inverse algebra_simps)
-
-text{*There is no slick version using division by zero.*}
-lemma inverse_add:
- "[| a \<noteq> 0; b \<noteq> 0 |]
- ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
-by (simp add: division_ring_inverse_add mult_ac)
-
-lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]:
-assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
-proof -
- have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
- by (simp add: divide_inverse nonzero_inverse_mult_distrib)
- also have "... = a * inverse b * (inverse c * c)"
- by (simp only: mult_ac)
- also have "... = a * inverse b" by simp
- finally show ?thesis by (simp add: divide_inverse)
-qed
-
-lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]:
- "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
-by (simp add: mult_commute [of _ c])
-
-lemma divide_1 [simp]: "a / 1 = a"
-by (simp add: divide_inverse)
-
-lemma times_divide_eq_right: "a * (b / c) = (a * b) / c"
-by (simp add: divide_inverse mult_assoc)
-
-lemma times_divide_eq_left: "(b / c) * a = (b * a) / c"
-by (simp add: divide_inverse mult_ac)
-
-text {* These are later declared as simp rules. *}
-lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left
-
-lemma add_frac_eq:
- assumes "y \<noteq> 0" and "z \<noteq> 0"
- shows "x / y + w / z = (x * z + w * y) / (y * z)"
-proof -
- have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)"
- using assms by simp
- also have "\<dots> = (x * z + y * w) / (y * z)"
- by (simp only: add_divide_distrib)
- finally show ?thesis
- by (simp only: mult_commute)
-qed
-
-text{*Special Cancellation Simprules for Division*}
-
-lemma nonzero_mult_divide_cancel_right [simp, noatp]:
- "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
-using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
-
-lemma nonzero_mult_divide_cancel_left [simp, noatp]:
- "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
-using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
-
-lemma nonzero_divide_mult_cancel_right [simp, noatp]:
- "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
-using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
-
-lemma nonzero_divide_mult_cancel_left [simp, noatp]:
- "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
-using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
-
-lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]:
- "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
-using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
-
-lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]:
- "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
-using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
-
-lemma minus_divide_left: "- (a / b) = (-a) / b"
-by (simp add: divide_inverse)
-
-lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma divide_minus_left [simp, noatp]: "(-a) / b = - (a / b)"
-by (simp add: divide_inverse)
-
-lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
-by (simp add: diff_minus add_divide_distrib)
-
-lemma add_divide_eq_iff:
- "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
-by (simp add: add_divide_distrib)
-
-lemma divide_add_eq_iff:
- "z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
-by (simp add: add_divide_distrib)
-
-lemma diff_divide_eq_iff:
- "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
-by (simp add: diff_divide_distrib)
-
-lemma divide_diff_eq_iff:
- "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
-by (simp add: diff_divide_distrib)
-
-lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
-proof -
- assume [simp]: "c \<noteq> 0"
- have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
- also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
-proof -
- assume [simp]: "c \<noteq> 0"
- have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
- also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
-by simp
-
-lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
-by (erule subst, simp)
-
-lemmas field_eq_simps[noatp] = algebra_simps
- (* pull / out*)
- add_divide_eq_iff divide_add_eq_iff
- diff_divide_eq_iff divide_diff_eq_iff
- (* multiply eqn *)
- nonzero_eq_divide_eq nonzero_divide_eq_eq
-(* is added later:
- times_divide_eq_left times_divide_eq_right
-*)
-
-text{*An example:*}
-lemma "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f\<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
-apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
- apply(simp add:field_eq_simps)
-apply(simp)
-done
-
-lemma diff_frac_eq:
- "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y - w / z = (x * z - w * y) / (y * z)"
-by (simp add: field_eq_simps times_divide_eq)
-
-lemma frac_eq_eq:
- "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (x / y = w / z) = (x * z = w * y)"
-by (simp add: field_eq_simps times_divide_eq)
-
-end
-
-class division_by_zero = zero + inverse +
- assumes inverse_zero [simp]: "inverse 0 = 0"
-
-lemma divide_zero [simp]:
- "a / 0 = (0::'a::{field,division_by_zero})"
-by (simp add: divide_inverse)
-
-lemma divide_self_if [simp]:
- "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
-by simp
-
-class mult_mono = times + zero + ord +
- assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
- assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
-
-class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add
-begin
-
-lemma mult_mono:
- "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> c
- \<Longrightarrow> a * c \<le> b * d"
-apply (erule mult_right_mono [THEN order_trans], assumption)
-apply (erule mult_left_mono, assumption)
-done
-
-lemma mult_mono':
- "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c
- \<Longrightarrow> a * c \<le> b * d"
-apply (rule mult_mono)
-apply (fast intro: order_trans)+
-done
-
-end
-
-class ordered_cancel_semiring = mult_mono + ordered_ab_semigroup_add
- + semiring + cancel_comm_monoid_add
-begin
-
-subclass semiring_0_cancel ..
-subclass ordered_semiring ..
-
-lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
-using mult_left_mono [of zero b a] by simp
-
-lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
-using mult_left_mono [of b zero a] by simp
-
-lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
-using mult_right_mono [of a zero b] by simp
-
-text {* Legacy - use @{text mult_nonpos_nonneg} *}
-lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
-by (drule mult_right_mono [of b zero], auto)
-
-lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> 0"
-by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
-
-end
-
-class linordered_semiring = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + mult_mono
-begin
-
-subclass ordered_cancel_semiring ..
-
-subclass ordered_comm_monoid_add ..
-
-lemma mult_left_less_imp_less:
- "c * a < c * b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
-by (force simp add: mult_left_mono not_le [symmetric])
-
-lemma mult_right_less_imp_less:
- "a * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
-by (force simp add: mult_right_mono not_le [symmetric])
-
-end
-
-class linlinordered_semiring_1 = linordered_semiring + semiring_1
-
-class linlinordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add +
- assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
- assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
-begin
-
-subclass semiring_0_cancel ..
-
-subclass linordered_semiring
-proof
- fix a b c :: 'a
- assume A: "a \<le> b" "0 \<le> c"
- from A show "c * a \<le> c * b"
- unfolding le_less
- using mult_strict_left_mono by (cases "c = 0") auto
- from A show "a * c \<le> b * c"
- unfolding le_less
- using mult_strict_right_mono by (cases "c = 0") auto
-qed
-
-lemma mult_left_le_imp_le:
- "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
-by (force simp add: mult_strict_left_mono _not_less [symmetric])
-
-lemma mult_right_le_imp_le:
- "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
-by (force simp add: mult_strict_right_mono not_less [symmetric])
-
-lemma mult_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
-using mult_strict_left_mono [of zero b a] by simp
-
-lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
-using mult_strict_left_mono [of b zero a] by simp
-
-lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
-using mult_strict_right_mono [of a zero b] by simp
-
-text {* Legacy - use @{text mult_neg_pos} *}
-lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
-by (drule mult_strict_right_mono [of b zero], auto)
-
-lemma zero_less_mult_pos:
- "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0")
- apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg [of a b])
- apply (auto dest: less_not_sym)
-done
-
-lemma zero_less_mult_pos2:
- "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
-apply (cases "b\<le>0")
- apply (auto simp add: le_less not_less)
-apply (drule_tac mult_pos_neg2 [of a b])
- apply (auto dest: less_not_sym)
-done
-
-text{*Strict monotonicity in both arguments*}
-lemma mult_strict_mono:
- assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
- shows "a * c < b * d"
- using assms apply (cases "c=0")
- apply (simp add: mult_pos_pos)
- apply (erule mult_strict_right_mono [THEN less_trans])
- apply (force simp add: le_less)
- apply (erule mult_strict_left_mono, assumption)
- done
-
-text{*This weaker variant has more natural premises*}
-lemma mult_strict_mono':
- assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
- shows "a * c < b * d"
-by (rule mult_strict_mono) (insert assms, auto)
-
-lemma mult_less_le_imp_less:
- assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
- shows "a * c < b * d"
- using assms apply (subgoal_tac "a * c < b * c")
- apply (erule less_le_trans)
- apply (erule mult_left_mono)
- apply simp
- apply (erule mult_strict_right_mono)
- apply assumption
- done
-
-lemma mult_le_less_imp_less:
- assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
- shows "a * c < b * d"
- using assms apply (subgoal_tac "a * c \<le> b * c")
- apply (erule le_less_trans)
- apply (erule mult_strict_left_mono)
- apply simp
- apply (erule mult_right_mono)
- apply simp
- done
-
-lemma mult_less_imp_less_left:
- assumes less: "c * a < c * b" and nonneg: "0 \<le> c"
- shows "a < b"
-proof (rule ccontr)
- assume "\<not> a < b"
- hence "b \<le> a" by (simp add: linorder_not_less)
- hence "c * b \<le> c * a" using nonneg by (rule mult_left_mono)
- with this and less show False by (simp add: not_less [symmetric])
-qed
-
-lemma mult_less_imp_less_right:
- assumes less: "a * c < b * c" and nonneg: "0 \<le> c"
- shows "a < b"
-proof (rule ccontr)
- assume "\<not> a < b"
- hence "b \<le> a" by (simp add: linorder_not_less)
- hence "b * c \<le> a * c" using nonneg by (rule mult_right_mono)
- with this and less show False by (simp add: not_less [symmetric])
-qed
-
-end
-
-class linlinlinordered_semiring_1_strict = linlinordered_semiring_strict + semiring_1
-
-class mult_mono1 = times + zero + ord +
- assumes mult_mono1: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
-
-class ordered_comm_semiring = comm_semiring_0
- + ordered_ab_semigroup_add + mult_mono1
-begin
-
-subclass ordered_semiring
-proof
- fix a b c :: 'a
- assume "a \<le> b" "0 \<le> c"
- thus "c * a \<le> c * b" by (rule mult_mono1)
- thus "a * c \<le> b * c" by (simp only: mult_commute)
-qed
-
-end
-
-class ordered_cancel_comm_semiring = comm_semiring_0_cancel
- + ordered_ab_semigroup_add + mult_mono1
-begin
-
-subclass ordered_comm_semiring ..
-subclass ordered_cancel_semiring ..
-
-end
-
-class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add +
- assumes mult_strict_left_mono_comm: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
-begin
-
-subclass linlinordered_semiring_strict
-proof
- fix a b c :: 'a
- assume "a < b" "0 < c"
- thus "c * a < c * b" by (rule mult_strict_left_mono_comm)
- thus "a * c < b * c" by (simp only: mult_commute)
-qed
-
-subclass ordered_cancel_comm_semiring
-proof
- fix a b c :: 'a
- assume "a \<le> b" "0 \<le> c"
- thus "c * a \<le> c * b"
- unfolding le_less
- using mult_strict_left_mono by (cases "c = 0") auto
-qed
-
-end
-
-class ordered_ring = ring + ordered_cancel_semiring
-begin
-
-subclass ordered_ab_group_add ..
-
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[noatp] = algebra_simps
-
-lemma less_add_iff1:
- "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
-by (simp add: algebra_simps)
-
-lemma less_add_iff2:
- "a * e + c < b * e + d \<longleftrightarrow> c < (b - a) * e + d"
-by (simp add: algebra_simps)
-
-lemma le_add_iff1:
- "a * e + c \<le> b * e + d \<longleftrightarrow> (a - b) * e + c \<le> d"
-by (simp add: algebra_simps)
-
-lemma le_add_iff2:
- "a * e + c \<le> b * e + d \<longleftrightarrow> c \<le> (b - a) * e + d"
-by (simp add: algebra_simps)
-
-lemma mult_left_mono_neg:
- "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
- apply (drule mult_left_mono [of _ _ "uminus c"])
- apply (simp_all add: minus_mult_left [symmetric])
- done
-
-lemma mult_right_mono_neg:
- "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
- apply (drule mult_right_mono [of _ _ "uminus c"])
- apply (simp_all add: minus_mult_right [symmetric])
- done
-
-lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
-using mult_right_mono_neg [of a zero b] by simp
-
-lemma split_mult_pos_le:
- "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
-by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos)
-
-end
-
-class abs_if = minus + uminus + ord + zero + abs +
- assumes abs_if: "\<bar>a\<bar> = (if a < 0 then - a else a)"
-
-class sgn_if = minus + uminus + zero + one + ord + sgn +
- assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
-
-lemma (in sgn_if) sgn0[simp]: "sgn 0 = 0"
-by(simp add:sgn_if)
-
-class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if
-begin
-
-subclass ordered_ring ..
-
-subclass ordered_ab_group_add_abs
-proof
- fix a b
- show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
- by (auto simp add: abs_if not_less neg_less_eq_nonneg less_eq_neg_nonpos)
- (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric]
- neg_less_eq_nonneg less_eq_neg_nonpos, auto intro: add_nonneg_nonneg,
- auto intro!: less_imp_le add_neg_neg)
-qed (auto simp add: abs_if less_eq_neg_nonpos neg_equal_zero)
-
-end
-
-(* The "strict" suffix can be seen as describing the combination of linordered_ring and no_zero_divisors.
- Basically, linordered_ring + no_zero_divisors = linlinordered_ring_strict.
- *)
-class linlinordered_ring_strict = ring + linlinordered_semiring_strict
- + ordered_ab_group_add + abs_if
-begin
-
-subclass linordered_ring ..
-
-lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
-using mult_strict_left_mono [of b a "- c"] by simp
-
-lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
-using mult_strict_right_mono [of b a "- c"] by simp
-
-lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
-using mult_strict_right_mono_neg [of a zero b] by simp
-
-subclass ring_no_zero_divisors
-proof
- fix a b
- assume "a \<noteq> 0" then have A: "a < 0 \<or> 0 < a" by (simp add: neq_iff)
- assume "b \<noteq> 0" then have B: "b < 0 \<or> 0 < b" by (simp add: neq_iff)
- have "a * b < 0 \<or> 0 < a * b"
- proof (cases "a < 0")
- case True note A' = this
- show ?thesis proof (cases "b < 0")
- case True with A'
- show ?thesis by (auto dest: mult_neg_neg)
- next
- case False with B have "0 < b" by auto
- with A' show ?thesis by (auto dest: mult_strict_right_mono)
- qed
- next
- case False with A have A': "0 < a" by auto
- show ?thesis proof (cases "b < 0")
- case True with A'
- show ?thesis by (auto dest: mult_strict_right_mono_neg)
- next
- case False with B have "0 < b" by auto
- with A' show ?thesis by (auto dest: mult_pos_pos)
- qed
- qed
- then show "a * b \<noteq> 0" by (simp add: neq_iff)
-qed
-
-lemma zero_less_mult_iff:
- "0 < a * b \<longleftrightarrow> 0 < a \<and> 0 < b \<or> a < 0 \<and> b < 0"
- apply (auto simp add: mult_pos_pos mult_neg_neg)
- apply (simp_all add: not_less le_less)
- apply (erule disjE) apply assumption defer
- apply (erule disjE) defer apply (drule sym) apply simp
- apply (erule disjE) defer apply (drule sym) apply simp
- apply (erule disjE) apply assumption apply (drule sym) apply simp
- apply (drule sym) apply simp
- apply (blast dest: zero_less_mult_pos)
- apply (blast dest: zero_less_mult_pos2)
- done
-
-lemma zero_le_mult_iff:
- "0 \<le> a * b \<longleftrightarrow> 0 \<le> a \<and> 0 \<le> b \<or> a \<le> 0 \<and> b \<le> 0"
-by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff)
-
-lemma mult_less_0_iff:
- "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
- apply (insert zero_less_mult_iff [of "-a" b])
- apply (force simp add: minus_mult_left[symmetric])
- done
-
-lemma mult_le_0_iff:
- "a * b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
- apply (insert zero_le_mult_iff [of "-a" b])
- apply (force simp add: minus_mult_left[symmetric])
- done
-
-lemma zero_le_square [simp]: "0 \<le> a * a"
-by (simp add: zero_le_mult_iff linear)
-
-lemma not_square_less_zero [simp]: "\<not> (a * a < 0)"
-by (simp add: not_less)
-
-text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
- also with the relations @{text "\<le>"} and equality.*}
-
-text{*These ``disjunction'' versions produce two cases when the comparison is
- an assumption, but effectively four when the comparison is a goal.*}
-
-lemma mult_less_cancel_right_disj:
- "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
- apply (cases "c = 0")
- apply (auto simp add: neq_iff mult_strict_right_mono
- mult_strict_right_mono_neg)
- apply (auto simp add: not_less
- not_le [symmetric, of "a*c"]
- not_le [symmetric, of a])
- apply (erule_tac [!] notE)
- apply (auto simp add: less_imp_le mult_right_mono
- mult_right_mono_neg)
- done
-
-lemma mult_less_cancel_left_disj:
- "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
- apply (cases "c = 0")
- apply (auto simp add: neq_iff mult_strict_left_mono
- mult_strict_left_mono_neg)
- apply (auto simp add: not_less
- not_le [symmetric, of "c*a"]
- not_le [symmetric, of a])
- apply (erule_tac [!] notE)
- apply (auto simp add: less_imp_le mult_left_mono
- mult_left_mono_neg)
- done
-
-text{*The ``conjunction of implication'' lemmas produce two cases when the
-comparison is a goal, but give four when the comparison is an assumption.*}
-
-lemma mult_less_cancel_right:
- "a * c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
- using mult_less_cancel_right_disj [of a c b] by auto
-
-lemma mult_less_cancel_left:
- "c * a < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
- using mult_less_cancel_left_disj [of c a b] by auto
-
-lemma mult_le_cancel_right:
- "a * c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
-by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
-
-lemma mult_le_cancel_left:
- "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
-by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
-
-lemma mult_le_cancel_left_pos:
- "0 < c \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> a \<le> b"
-by (auto simp: mult_le_cancel_left)
-
-lemma mult_le_cancel_left_neg:
- "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
-by (auto simp: mult_le_cancel_left)
-
-lemma mult_less_cancel_left_pos:
- "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
-by (auto simp: mult_less_cancel_left)
-
-lemma mult_less_cancel_left_neg:
- "c < 0 \<Longrightarrow> c * a < c * b \<longleftrightarrow> b < a"
-by (auto simp: mult_less_cancel_left)
-
-end
-
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[noatp] = algebra_simps
-
-lemmas mult_sign_intros =
- mult_nonneg_nonneg mult_nonneg_nonpos
- mult_nonpos_nonneg mult_nonpos_nonpos
- mult_pos_pos mult_pos_neg
- mult_neg_pos mult_neg_neg
-
-class ordered_comm_ring = comm_ring + ordered_comm_semiring
-begin
-
-subclass ordered_ring ..
-subclass ordered_cancel_comm_semiring ..
-
-end
-
-class linordered_semidom = comm_semiring_1_cancel + linordered_comm_semiring_strict +
- (*previously linordered_semiring*)
- assumes zero_less_one [simp]: "0 < 1"
-begin
-
-lemma pos_add_strict:
- shows "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
- using add_strict_mono [of zero a b c] by simp
-
-lemma zero_le_one [simp]: "0 \<le> 1"
-by (rule zero_less_one [THEN less_imp_le])
-
-lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
-by (simp add: not_le)
-
-lemma not_one_less_zero [simp]: "\<not> 1 < 0"
-by (simp add: not_less)
-
-lemma less_1_mult:
- assumes "1 < m" and "1 < n"
- shows "1 < m * n"
- using assms mult_strict_mono [of 1 m 1 n]
- by (simp add: less_trans [OF zero_less_one])
-
-end
-
-class linordered_idom = comm_ring_1 +
- linordered_comm_semiring_strict + ordered_ab_group_add +
- abs_if + sgn_if
- (*previously linordered_ring*)
-begin
-
-subclass linlinordered_ring_strict ..
-subclass ordered_comm_ring ..
-subclass idom ..
-
-subclass linordered_semidom
-proof
- have "0 \<le> 1 * 1" by (rule zero_le_square)
- thus "0 < 1" by (simp add: le_less)
-qed
-
-lemma linorder_neqE_linordered_idom:
- assumes "x \<noteq> y" obtains "x < y" | "y < x"
- using assms by (rule neqE)
-
-text {* These cancellation simprules also produce two cases when the comparison is a goal. *}
-
-lemma mult_le_cancel_right1:
- "c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
-by (insert mult_le_cancel_right [of 1 c b], simp)
-
-lemma mult_le_cancel_right2:
- "a * c \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
-by (insert mult_le_cancel_right [of a c 1], simp)
-
-lemma mult_le_cancel_left1:
- "c \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
-by (insert mult_le_cancel_left [of c 1 b], simp)
-
-lemma mult_le_cancel_left2:
- "c * a \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
-by (insert mult_le_cancel_left [of c a 1], simp)
-
-lemma mult_less_cancel_right1:
- "c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
-by (insert mult_less_cancel_right [of 1 c b], simp)
-
-lemma mult_less_cancel_right2:
- "a * c < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
-by (insert mult_less_cancel_right [of a c 1], simp)
-
-lemma mult_less_cancel_left1:
- "c < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
-by (insert mult_less_cancel_left [of c 1 b], simp)
-
-lemma mult_less_cancel_left2:
- "c * a < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
-by (insert mult_less_cancel_left [of c a 1], simp)
-
-lemma sgn_sgn [simp]:
- "sgn (sgn a) = sgn a"
-unfolding sgn_if by simp
-
-lemma sgn_0_0:
- "sgn a = 0 \<longleftrightarrow> a = 0"
-unfolding sgn_if by simp
-
-lemma sgn_1_pos:
- "sgn a = 1 \<longleftrightarrow> a > 0"
-unfolding sgn_if by (simp add: neg_equal_zero)
-
-lemma sgn_1_neg:
- "sgn a = - 1 \<longleftrightarrow> a < 0"
-unfolding sgn_if by (auto simp add: equal_neg_zero)
-
-lemma sgn_pos [simp]:
- "0 < a \<Longrightarrow> sgn a = 1"
-unfolding sgn_1_pos .
-
-lemma sgn_neg [simp]:
- "a < 0 \<Longrightarrow> sgn a = - 1"
-unfolding sgn_1_neg .
-
-lemma sgn_times:
- "sgn (a * b) = sgn a * sgn b"
-by (auto simp add: sgn_if zero_less_mult_iff)
-
-lemma abs_sgn: "abs k = k * sgn k"
-unfolding sgn_if abs_if by auto
-
-lemma sgn_greater [simp]:
- "0 < sgn a \<longleftrightarrow> 0 < a"
- unfolding sgn_if by auto
-
-lemma sgn_less [simp]:
- "sgn a < 0 \<longleftrightarrow> a < 0"
- unfolding sgn_if by auto
-
-lemma abs_dvd_iff [simp]: "(abs m) dvd k \<longleftrightarrow> m dvd k"
- by (simp add: abs_if)
-
-lemma dvd_abs_iff [simp]: "m dvd (abs k) \<longleftrightarrow> m dvd k"
- by (simp add: abs_if)
-
-lemma dvd_if_abs_eq:
- "abs l = abs (k) \<Longrightarrow> l dvd k"
-by(subst abs_dvd_iff[symmetric]) simp
-
-end
-
-class linordered_field = field + linordered_idom
-
-text {* Simprules for comparisons where common factors can be cancelled. *}
-
-lemmas mult_compare_simps[noatp] =
- mult_le_cancel_right mult_le_cancel_left
- mult_le_cancel_right1 mult_le_cancel_right2
- mult_le_cancel_left1 mult_le_cancel_left2
- mult_less_cancel_right mult_less_cancel_left
- mult_less_cancel_right1 mult_less_cancel_right2
- mult_less_cancel_left1 mult_less_cancel_left2
- mult_cancel_right mult_cancel_left
- mult_cancel_right1 mult_cancel_right2
- mult_cancel_left1 mult_cancel_left2
-
--- {* FIXME continue localization here *}
-
-lemma inverse_nonzero_iff_nonzero [simp]:
- "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
-by (force dest: inverse_zero_imp_zero)
-
-lemma inverse_minus_eq [simp]:
- "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
-proof cases
- assume "a=0" thus ?thesis by (simp add: inverse_zero)
-next
- assume "a\<noteq>0"
- thus ?thesis by (simp add: nonzero_inverse_minus_eq)
-qed
-
-lemma inverse_eq_imp_eq:
- "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
-apply (cases "a=0 | b=0")
- apply (force dest!: inverse_zero_imp_zero
- simp add: eq_commute [of "0::'a"])
-apply (force dest!: nonzero_inverse_eq_imp_eq)
-done
-
-lemma inverse_eq_iff_eq [simp]:
- "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))"
-by (force dest!: inverse_eq_imp_eq)
-
-lemma inverse_inverse_eq [simp]:
- "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
- proof cases
- assume "a=0" thus ?thesis by simp
- next
- assume "a\<noteq>0"
- thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
- qed
-
-text{*This version builds in division by zero while also re-orienting
- the right-hand side.*}
-lemma inverse_mult_distrib [simp]:
- "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
- proof cases
- assume "a \<noteq> 0 & b \<noteq> 0"
- thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute)
- next
- assume "~ (a \<noteq> 0 & b \<noteq> 0)"
- thus ?thesis by force
- qed
-
-lemma inverse_divide [simp]:
- "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
-by (simp add: divide_inverse mult_commute)
-
-
-subsection {* Calculations with fractions *}
-
-text{* There is a whole bunch of simp-rules just for class @{text
-field} but none for class @{text field} and @{text nonzero_divides}
-because the latter are covered by a simproc. *}
-
-lemma mult_divide_mult_cancel_left:
- "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
-apply (cases "b = 0")
-apply (simp_all add: nonzero_mult_divide_mult_cancel_left)
-done
-
-lemma mult_divide_mult_cancel_right:
- "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
-apply (cases "b = 0")
-apply (simp_all add: nonzero_mult_divide_mult_cancel_right)
-done
-
-lemma divide_divide_eq_right [simp,noatp]:
- "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
-by (simp add: divide_inverse mult_ac)
-
-lemma divide_divide_eq_left [simp,noatp]:
- "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
-by (simp add: divide_inverse mult_assoc)
-
-
-subsubsection{*Special Cancellation Simprules for Division*}
-
-lemma mult_divide_mult_cancel_left_if[simp,noatp]:
-fixes c :: "'a :: {field,division_by_zero}"
-shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
-by (simp add: mult_divide_mult_cancel_left)
-
-
-subsection {* Division and Unary Minus *}
-
-lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
-by (simp add: divide_inverse)
-
-lemma divide_minus_right [simp, noatp]:
- "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
-by (simp add: divide_inverse)
-
-lemma minus_divide_divide:
- "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
-apply (cases "b=0", simp)
-apply (simp add: nonzero_minus_divide_divide)
-done
-
-lemma eq_divide_eq:
- "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
-by (simp add: nonzero_eq_divide_eq)
-
-lemma divide_eq_eq:
- "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
-by (force simp add: nonzero_divide_eq_eq)
-
-
-subsection {* Ordered Fields *}
-
-lemma positive_imp_inverse_positive:
-assumes a_gt_0: "0 < a" shows "0 < inverse (a::'a::linordered_field)"
-proof -
- have "0 < a * inverse a"
- by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one)
- thus "0 < inverse a"
- by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff)
-qed
-
-lemma negative_imp_inverse_negative:
- "a < 0 ==> inverse a < (0::'a::linordered_field)"
-by (insert positive_imp_inverse_positive [of "-a"],
- simp add: nonzero_inverse_minus_eq order_less_imp_not_eq)
-
-lemma inverse_le_imp_le:
-assumes invle: "inverse a \<le> inverse b" and apos: "0 < a"
-shows "b \<le> (a::'a::linordered_field)"
-proof (rule classical)
- assume "~ b \<le> a"
- hence "a < b" by (simp add: linorder_not_le)
- hence bpos: "0 < b" by (blast intro: apos order_less_trans)
- hence "a * inverse a \<le> a * inverse b"
- by (simp add: apos invle order_less_imp_le mult_left_mono)
- hence "(a * inverse a) * b \<le> (a * inverse b) * b"
- by (simp add: bpos order_less_imp_le mult_right_mono)
- thus "b \<le> a" by (simp add: mult_assoc apos bpos order_less_imp_not_eq2)
-qed
-
-lemma inverse_positive_imp_positive:
-assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
-shows "0 < (a::'a::linordered_field)"
-proof -
- have "0 < inverse (inverse a)"
- using inv_gt_0 by (rule positive_imp_inverse_positive)
- thus "0 < a"
- using nz by (simp add: nonzero_inverse_inverse_eq)
-qed
-
-lemma inverse_positive_iff_positive [simp]:
- "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
-done
-
-lemma inverse_negative_imp_negative:
-assumes inv_less_0: "inverse a < 0" and nz: "a \<noteq> 0"
-shows "a < (0::'a::linordered_field)"
-proof -
- have "inverse (inverse a) < 0"
- using inv_less_0 by (rule negative_imp_inverse_negative)
- thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq)
-qed
-
-lemma inverse_negative_iff_negative [simp]:
- "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
-done
-
-lemma inverse_nonnegative_iff_nonnegative [simp]:
- "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_less [symmetric])
-
-lemma inverse_nonpositive_iff_nonpositive [simp]:
- "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_less [symmetric])
-
-lemma linlinordered_field_no_lb: "\<forall> x. \<exists>y. y < (x::'a::linordered_field)"
-proof
- fix x::'a
- have m1: "- (1::'a) < 0" by simp
- from add_strict_right_mono[OF m1, where c=x]
- have "(- 1) + x < x" by simp
- thus "\<exists>y. y < x" by blast
-qed
-
-lemma linlinordered_field_no_ub: "\<forall> x. \<exists>y. y > (x::'a::linordered_field)"
-proof
- fix x::'a
- have m1: " (1::'a) > 0" by simp
- from add_strict_right_mono[OF m1, where c=x]
- have "1 + x > x" by simp
- thus "\<exists>y. y > x" by blast
-qed
-
-subsection{*Anti-Monotonicity of @{term inverse}*}
-
-lemma less_imp_inverse_less:
-assumes less: "a < b" and apos: "0 < a"
-shows "inverse b < inverse (a::'a::linordered_field)"
-proof (rule ccontr)
- assume "~ inverse b < inverse a"
- hence "inverse a \<le> inverse b" by (simp add: linorder_not_less)
- hence "~ (a < b)"
- by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
- thus False by (rule notE [OF _ less])
-qed
-
-lemma inverse_less_imp_less:
- "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)"
-apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"])
-apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq)
-done
-
-text{*Both premises are essential. Consider -1 and 1.*}
-lemma inverse_less_iff_less [simp,noatp]:
- "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
-by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less)
-
-lemma le_imp_inverse_le:
- "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
-by (force simp add: order_le_less less_imp_inverse_less)
-
-lemma inverse_le_iff_le [simp,noatp]:
- "[|0 < a; 0 < b|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
-by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le)
-
-
-text{*These results refer to both operands being negative. The opposite-sign
-case is trivial, since inverse preserves signs.*}
-lemma inverse_le_imp_le_neg:
- "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::linordered_field)"
-apply (rule classical)
-apply (subgoal_tac "a < 0")
- prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans)
-apply (insert inverse_le_imp_le [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
-done
-
-lemma less_imp_inverse_less_neg:
- "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)"
-apply (subgoal_tac "a < 0")
- prefer 2 apply (blast intro: order_less_trans)
-apply (insert less_imp_inverse_less [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
-done
-
-lemma inverse_less_imp_less_neg:
- "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)"
-apply (rule classical)
-apply (subgoal_tac "a < 0")
- prefer 2
- apply (force simp add: linorder_not_less intro: order_le_less_trans)
-apply (insert inverse_less_imp_less [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq)
-done
-
-lemma inverse_less_iff_less_neg [simp,noatp]:
- "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
-apply (insert inverse_less_iff_less [of "-b" "-a"])
-apply (simp del: inverse_less_iff_less
- add: order_less_imp_not_eq nonzero_inverse_minus_eq)
-done
-
-lemma le_imp_inverse_le_neg:
- "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
-by (force simp add: order_le_less less_imp_inverse_less_neg)
-
-lemma inverse_le_iff_le_neg [simp,noatp]:
- "[|a < 0; b < 0|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
-by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg)
-
-
-subsection{*Inverses and the Number One*}
-
-lemma one_less_inverse_iff:
- "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))"
-proof cases
- assume "0 < x"
- with inverse_less_iff_less [OF zero_less_one, of x]
- show ?thesis by simp
-next
- assume notless: "~ (0 < x)"
- have "~ (1 < inverse x)"
- proof
- assume "1 < inverse x"
- also with notless have "... \<le> 0" by (simp add: linorder_not_less)
- also have "... < 1" by (rule zero_less_one)
- finally show False by auto
- qed
- with notless show ?thesis by simp
-qed
-
-lemma inverse_eq_1_iff [simp]:
- "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
-by (insert inverse_eq_iff_eq [of x 1], simp)
-
-lemma one_le_inverse_iff:
- "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_by_zero}))"
-by (force simp add: order_le_less one_less_inverse_iff zero_less_one
- eq_commute [of 1])
-
-lemma inverse_less_1_iff:
- "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_le [symmetric] one_le_inverse_iff)
-
-lemma inverse_le_1_iff:
- "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_less [symmetric] one_less_inverse_iff)
-
-
-subsection{*Simplification of Inequalities Involving Literal Divisors*}
-
-lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
-proof -
- assume less: "0<c"
- hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
- by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
- also have "... = (a*c \<le> b)"
- by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
-proof -
- assume less: "c<0"
- hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
- by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
- also have "... = (b \<le> a*c)"
- by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma le_divide_eq:
- "(a \<le> b/c) =
- (if 0 < c then a*c \<le> b
- else if c < 0 then b \<le> a*c
- else a \<le> (0::'a::{linordered_field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff)
-done
-
-lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
-proof -
- assume less: "0<c"
- hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
- by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
- also have "... = (b \<le> a*c)"
- by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
-proof -
- assume less: "c<0"
- hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
- by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
- also have "... = (a*c \<le> b)"
- by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma divide_le_eq:
- "(b/c \<le> a) =
- (if 0 < c then b \<le> a*c
- else if c < 0 then a*c \<le> b
- else 0 \<le> (a::'a::{linordered_field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff)
-done
-
-lemma pos_less_divide_eq:
- "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)"
-proof -
- assume less: "0<c"
- hence "(a < b/c) = (a*c < (b/c)*c)"
- by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
- also have "... = (a*c < b)"
- by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma neg_less_divide_eq:
- "c < (0::'a::linordered_field) ==> (a < b/c) = (b < a*c)"
-proof -
- assume less: "c<0"
- hence "(a < b/c) = ((b/c)*c < a*c)"
- by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
- also have "... = (b < a*c)"
- by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma less_divide_eq:
- "(a < b/c) =
- (if 0 < c then a*c < b
- else if c < 0 then b < a*c
- else a < (0::'a::{linordered_field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff)
-done
-
-lemma pos_divide_less_eq:
- "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)"
-proof -
- assume less: "0<c"
- hence "(b/c < a) = ((b/c)*c < a*c)"
- by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
- also have "... = (b < a*c)"
- by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma neg_divide_less_eq:
- "c < (0::'a::linordered_field) ==> (b/c < a) = (a*c < b)"
-proof -
- assume less: "c<0"
- hence "(b/c < a) = (a*c < (b/c)*c)"
- by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
- also have "... = (a*c < b)"
- by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc)
- finally show ?thesis .
-qed
-
-lemma divide_less_eq:
- "(b/c < a) =
- (if 0 < c then b < a*c
- else if c < 0 then a*c < b
- else 0 < (a::'a::{linordered_field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff)
-done
-
-
-subsection{*Field simplification*}
-
-text{* Lemmas @{text field_simps} multiply with denominators in in(equations)
-if they can be proved to be non-zero (for equations) or positive/negative
-(for inequations). Can be too aggressive and is therefore separate from the
-more benign @{text algebra_simps}. *}
-
-lemmas field_simps[noatp] = field_eq_simps
- (* multiply ineqn *)
- pos_divide_less_eq neg_divide_less_eq
- pos_less_divide_eq neg_less_divide_eq
- pos_divide_le_eq neg_divide_le_eq
- pos_le_divide_eq neg_le_divide_eq
-
-text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
-of positivity/negativity needed for @{text field_simps}. Have not added @{text
-sign_simps} to @{text field_simps} because the former can lead to case
-explosions. *}
-
-lemmas sign_simps[noatp] = group_simps
- zero_less_mult_iff mult_less_0_iff
-
-(* Only works once linear arithmetic is installed:
-text{*An example:*}
-lemma fixes a b c d e f :: "'a::linordered_field"
-shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
- ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
- ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
-apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
- prefer 2 apply(simp add:sign_simps)
-apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
- prefer 2 apply(simp add:sign_simps)
-apply(simp add:field_simps)
-done
-*)
-
-
-subsection{*Division and Signs*}
-
-lemma zero_less_divide_iff:
- "((0::'a::{linordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
-by (simp add: divide_inverse zero_less_mult_iff)
-
-lemma divide_less_0_iff:
- "(a/b < (0::'a::{linordered_field,division_by_zero})) =
- (0 < a & b < 0 | a < 0 & 0 < b)"
-by (simp add: divide_inverse mult_less_0_iff)
-
-lemma zero_le_divide_iff:
- "((0::'a::{linordered_field,division_by_zero}) \<le> a/b) =
- (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
-by (simp add: divide_inverse zero_le_mult_iff)
-
-lemma divide_le_0_iff:
- "(a/b \<le> (0::'a::{linordered_field,division_by_zero})) =
- (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
-by (simp add: divide_inverse mult_le_0_iff)
-
-lemma divide_eq_0_iff [simp,noatp]:
- "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
-by (simp add: divide_inverse)
-
-lemma divide_pos_pos:
- "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y"
-by(simp add:field_simps)
-
-
-lemma divide_nonneg_pos:
- "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y"
-by(simp add:field_simps)
-
-lemma divide_neg_pos:
- "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0"
-by(simp add:field_simps)
-
-lemma divide_nonpos_pos:
- "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0"
-by(simp add:field_simps)
-
-lemma divide_pos_neg:
- "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0"
-by(simp add:field_simps)
-
-lemma divide_nonneg_neg:
- "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0"
-by(simp add:field_simps)
-
-lemma divide_neg_neg:
- "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y"
-by(simp add:field_simps)
-
-lemma divide_nonpos_neg:
- "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y"
-by(simp add:field_simps)
-
-
-subsection{*Cancellation Laws for Division*}
-
-lemma divide_cancel_right [simp,noatp]:
- "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (simp add: divide_inverse)
-done
-
-lemma divide_cancel_left [simp,noatp]:
- "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (simp add: divide_inverse)
-done
-
-
-subsection {* Division and the Number One *}
-
-text{*Simplify expressions equated with 1*}
-lemma divide_eq_1_iff [simp,noatp]:
- "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
-apply (cases "b=0", simp)
-apply (simp add: right_inverse_eq)
-done
-
-lemma one_eq_divide_iff [simp,noatp]:
- "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
-by (simp add: eq_commute [of 1])
-
-lemma zero_eq_1_divide_iff [simp,noatp]:
- "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)"
-apply (cases "a=0", simp)
-apply (auto simp add: nonzero_eq_divide_eq)
-done
-
-lemma one_divide_eq_0_iff [simp,noatp]:
- "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)"
-apply (cases "a=0", simp)
-apply (insert zero_neq_one [THEN not_sym])
-apply (auto simp add: nonzero_divide_eq_eq)
-done
-
-text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
-lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified]
-lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified]
-lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified]
-lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified]
-
-declare zero_less_divide_1_iff [simp,noatp]
-declare divide_less_0_1_iff [simp,noatp]
-declare zero_le_divide_1_iff [simp,noatp]
-declare divide_le_0_1_iff [simp,noatp]
-
-
-subsection {* Ordering Rules for Division *}
-
-lemma divide_strict_right_mono:
- "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)"
-by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono
- positive_imp_inverse_positive)
-
-lemma divide_right_mono:
- "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_by_zero})"
-by (force simp add: divide_strict_right_mono order_le_less)
-
-lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b
- ==> c <= 0 ==> b / c <= a / c"
-apply (drule divide_right_mono [of _ _ "- c"])
-apply auto
-done
-
-lemma divide_strict_right_mono_neg:
- "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)"
-apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
-apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric])
-done
-
-text{*The last premise ensures that @{term a} and @{term b}
- have the same sign*}
-lemma divide_strict_left_mono:
- "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono)
-
-lemma divide_left_mono:
- "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono)
-
-lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b
- ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
- apply (drule divide_left_mono [of _ _ "- c"])
- apply (auto simp add: mult_commute)
-done
-
-lemma divide_strict_left_mono_neg:
- "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg)
-
-
-text{*Simplify quotients that are compared with the value 1.*}
-
-lemma le_divide_eq_1 [noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
-by (auto simp add: le_divide_eq)
-
-lemma divide_le_eq_1 [noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
-by (auto simp add: divide_le_eq)
-
-lemma less_divide_eq_1 [noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
-by (auto simp add: less_divide_eq)
-
-lemma divide_less_eq_1 [noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
-by (auto simp add: divide_less_eq)
-
-
-subsection{*Conditional Simplification Rules: No Case Splits*}
-
-lemma le_divide_eq_1_pos [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
-by (auto simp add: le_divide_eq)
-
-lemma le_divide_eq_1_neg [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
-by (auto simp add: le_divide_eq)
-
-lemma divide_le_eq_1_pos [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
-by (auto simp add: divide_le_eq)
-
-lemma divide_le_eq_1_neg [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
-by (auto simp add: divide_le_eq)
-
-lemma less_divide_eq_1_pos [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
-by (auto simp add: less_divide_eq)
-
-lemma less_divide_eq_1_neg [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
-by (auto simp add: less_divide_eq)
-
-lemma divide_less_eq_1_pos [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
-by (auto simp add: divide_less_eq)
-
-lemma divide_less_eq_1_neg [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
-by (auto simp add: divide_less_eq)
-
-lemma eq_divide_eq_1 [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(1 = b/a) = ((a \<noteq> 0 & a = b))"
-by (auto simp add: eq_divide_eq)
-
-lemma divide_eq_eq_1 [simp,noatp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
- shows "(b/a = 1) = ((a \<noteq> 0 & a = b))"
-by (auto simp add: divide_eq_eq)
-
-
-subsection {* Reasoning about inequalities with division *}
-
-lemma mult_right_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
- ==> x * y <= x"
-by (auto simp add: mult_compare_simps)
-
-lemma mult_left_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
- ==> y * x <= x"
-by (auto simp add: mult_compare_simps)
-
-lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==>
- x / y <= z"
-by (subst pos_divide_le_eq, assumption+)
-
-lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==>
- z <= x / y"
-by(simp add:field_simps)
-
-lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==>
- x / y < z"
-by(simp add:field_simps)
-
-lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==>
- z < x / y"
-by(simp add:field_simps)
-
-lemma frac_le: "(0::'a::linordered_field) <= x ==>
- x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w"
- apply (rule mult_imp_div_pos_le)
- apply simp
- apply (subst times_divide_eq_left)
- apply (rule mult_imp_le_div_pos, assumption)
- apply (rule mult_mono)
- apply simp_all
-done
-
-lemma frac_less: "(0::'a::linordered_field) <= x ==>
- x < y ==> 0 < w ==> w <= z ==> x / z < y / w"
- apply (rule mult_imp_div_pos_less)
- apply simp
- apply (subst times_divide_eq_left)
- apply (rule mult_imp_less_div_pos, assumption)
- apply (erule mult_less_le_imp_less)
- apply simp_all
-done
-
-lemma frac_less2: "(0::'a::linordered_field) < x ==>
- x <= y ==> 0 < w ==> w < z ==> x / z < y / w"
- apply (rule mult_imp_div_pos_less)
- apply simp_all
- apply (subst times_divide_eq_left)
- apply (rule mult_imp_less_div_pos, assumption)
- apply (erule mult_le_less_imp_less)
- apply simp_all
-done
-
-text{*It's not obvious whether these should be simprules or not.
- Their effect is to gather terms into one big fraction, like
- a*b*c / x*y*z. The rationale for that is unclear, but many proofs
- seem to need them.*}
-
-declare times_divide_eq [simp]
-
-
-subsection {* Ordered Fields are Dense *}
-
-context linordered_semidom
-begin
-
-lemma less_add_one: "a < a + 1"
-proof -
- have "a + 0 < a + 1"
- by (blast intro: zero_less_one add_strict_left_mono)
- thus ?thesis by simp
-qed
-
-lemma zero_less_two: "0 < 1 + 1"
-by (blast intro: less_trans zero_less_one less_add_one)
-
-end
-
-lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)"
-by (simp add: field_simps zero_less_two)
-
-lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b"
-by (simp add: field_simps zero_less_two)
-
-instance linordered_field < dense_linorder
-proof
- fix x y :: 'a
- have "x < x + 1" by simp
- then show "\<exists>y. x < y" ..
- have "x - 1 < x" by simp
- then show "\<exists>y. y < x" ..
- show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
-qed
-
-
-subsection {* Absolute Value *}
-
-context linordered_idom
-begin
-
-lemma mult_sgn_abs: "sgn x * abs x = x"
- unfolding abs_if sgn_if by auto
-
-end
-
-lemma abs_one [simp]: "abs 1 = (1::'a::linordered_idom)"
-by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
-
-class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs +
- assumes abs_eq_mult:
- "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0) \<Longrightarrow> \<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
-
-context linordered_idom
-begin
-
-subclass ordered_ring_abs proof
-qed (auto simp add: abs_if not_less equal_neg_zero neg_equal_zero mult_less_0_iff)
-
-lemma abs_mult:
- "abs (a * b) = abs a * abs b"
- by (rule abs_eq_mult) auto
-
-lemma abs_mult_self:
- "abs a * abs a = a * a"
- by (simp add: abs_if)
-
-end
-
-lemma nonzero_abs_inverse:
- "a \<noteq> 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)"
-apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq
- negative_imp_inverse_negative)
-apply (blast intro: positive_imp_inverse_positive elim: order_less_asym)
-done
-
-lemma abs_inverse [simp]:
- "abs (inverse (a::'a::{linordered_field,division_by_zero})) =
- inverse (abs a)"
-apply (cases "a=0", simp)
-apply (simp add: nonzero_abs_inverse)
-done
-
-lemma nonzero_abs_divide:
- "b \<noteq> 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b"
-by (simp add: divide_inverse abs_mult nonzero_abs_inverse)
-
-lemma abs_divide [simp]:
- "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b"
-apply (cases "b=0", simp)
-apply (simp add: nonzero_abs_divide)
-done
-
-lemma abs_mult_less:
- "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::linordered_idom)"
-proof -
- assume ac: "abs a < c"
- hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
- assume "abs b < d"
- thus ?thesis by (simp add: ac cpos mult_strict_mono)
-qed
-
-lemmas eq_minus_self_iff[noatp] = equal_neg_zero
-
-lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::linordered_idom))"
- unfolding order_less_le less_eq_neg_nonpos equal_neg_zero ..
-
-lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::linordered_idom))"
-apply (simp add: order_less_le abs_le_iff)
-apply (auto simp add: abs_if neg_less_eq_nonneg less_eq_neg_nonpos)
-done
-
-lemma abs_mult_pos: "(0::'a::linordered_idom) <= x ==>
- (abs y) * x = abs (y * x)"
- apply (subst abs_mult)
- apply simp
-done
-
-lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==>
- abs x / y = abs (x / y)"
- apply (subst abs_divide)
- apply (simp add: order_less_imp_le)
-done
-
-code_modulename SML
- Ring_and_Field Arith
-
-code_modulename OCaml
- Ring_and_Field Arith
-
-code_modulename Haskell
- Ring_and_Field Arith
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Rings.thy Mon Feb 08 17:13:45 2010 +0100
@@ -0,0 +1,1212 @@
+(* Title: HOL/Rings.thy
+ Author: Gertrud Bauer
+ Author: Steven Obua
+ Author: Tobias Nipkow
+ Author: Lawrence C Paulson
+ Author: Markus Wenzel
+ Author: Jeremy Avigad
+*)
+
+header {* Rings *}
+
+theory Rings
+imports Groups
+begin
+
+text {*
+ The theory of partially ordered rings is taken from the books:
+ \begin{itemize}
+ \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
+ \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+ \end{itemize}
+ Most of the used notions can also be looked up in
+ \begin{itemize}
+ \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+ \item \emph{Algebra I} by van der Waerden, Springer.
+ \end{itemize}
+*}
+
+class semiring = ab_semigroup_add + semigroup_mult +
+ assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
+ assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
+begin
+
+text{*For the @{text combine_numerals} simproc*}
+lemma combine_common_factor:
+ "a * e + (b * e + c) = (a + b) * e + c"
+by (simp add: left_distrib add_ac)
+
+end
+
+class mult_zero = times + zero +
+ assumes mult_zero_left [simp]: "0 * a = 0"
+ assumes mult_zero_right [simp]: "a * 0 = 0"
+
+class semiring_0 = semiring + comm_monoid_add + mult_zero
+
+class semiring_0_cancel = semiring + cancel_comm_monoid_add
+begin
+
+subclass semiring_0
+proof
+ fix a :: 'a
+ have "0 * a + 0 * a = 0 * a + 0" by (simp add: left_distrib [symmetric])
+ thus "0 * a = 0" by (simp only: add_left_cancel)
+next
+ fix a :: 'a
+ have "a * 0 + a * 0 = a * 0 + 0" by (simp add: right_distrib [symmetric])
+ thus "a * 0 = 0" by (simp only: add_left_cancel)
+qed
+
+end
+
+class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
+ assumes distrib: "(a + b) * c = a * c + b * c"
+begin
+
+subclass semiring
+proof
+ fix a b c :: 'a
+ show "(a + b) * c = a * c + b * c" by (simp add: distrib)
+ have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
+ also have "... = b * a + c * a" by (simp only: distrib)
+ also have "... = a * b + a * c" by (simp add: mult_ac)
+ finally show "a * (b + c) = a * b + a * c" by blast
+qed
+
+end
+
+class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
+begin
+
+subclass semiring_0 ..
+
+end
+
+class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add
+begin
+
+subclass semiring_0_cancel ..
+
+subclass comm_semiring_0 ..
+
+end
+
+class zero_neq_one = zero + one +
+ assumes zero_neq_one [simp]: "0 \<noteq> 1"
+begin
+
+lemma one_neq_zero [simp]: "1 \<noteq> 0"
+by (rule not_sym) (rule zero_neq_one)
+
+end
+
+class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
+
+text {* Abstract divisibility *}
+
+class dvd = times
+begin
+
+definition dvd :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "dvd" 50) where
+ [code del]: "b dvd a \<longleftrightarrow> (\<exists>k. a = b * k)"
+
+lemma dvdI [intro?]: "a = b * k \<Longrightarrow> b dvd a"
+ unfolding dvd_def ..
+
+lemma dvdE [elim?]: "b dvd a \<Longrightarrow> (\<And>k. a = b * k \<Longrightarrow> P) \<Longrightarrow> P"
+ unfolding dvd_def by blast
+
+end
+
+class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult + dvd
+ (*previously almost_semiring*)
+begin
+
+subclass semiring_1 ..
+
+lemma dvd_refl[simp]: "a dvd a"
+proof
+ show "a = a * 1" by simp
+qed
+
+lemma dvd_trans:
+ assumes "a dvd b" and "b dvd c"
+ shows "a dvd c"
+proof -
+ from assms obtain v where "b = a * v" by (auto elim!: dvdE)
+ moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE)
+ ultimately have "c = a * (v * w)" by (simp add: mult_assoc)
+ then show ?thesis ..
+qed
+
+lemma dvd_0_left_iff [noatp, simp]: "0 dvd a \<longleftrightarrow> a = 0"
+by (auto intro: dvd_refl elim!: dvdE)
+
+lemma dvd_0_right [iff]: "a dvd 0"
+proof
+ show "0 = a * 0" by simp
+qed
+
+lemma one_dvd [simp]: "1 dvd a"
+by (auto intro!: dvdI)
+
+lemma dvd_mult[simp]: "a dvd c \<Longrightarrow> a dvd (b * c)"
+by (auto intro!: mult_left_commute dvdI elim!: dvdE)
+
+lemma dvd_mult2[simp]: "a dvd b \<Longrightarrow> a dvd (b * c)"
+ apply (subst mult_commute)
+ apply (erule dvd_mult)
+ done
+
+lemma dvd_triv_right [simp]: "a dvd b * a"
+by (rule dvd_mult) (rule dvd_refl)
+
+lemma dvd_triv_left [simp]: "a dvd a * b"
+by (rule dvd_mult2) (rule dvd_refl)
+
+lemma mult_dvd_mono:
+ assumes "a dvd b"
+ and "c dvd d"
+ shows "a * c dvd b * d"
+proof -
+ from `a dvd b` obtain b' where "b = a * b'" ..
+ moreover from `c dvd d` obtain d' where "d = c * d'" ..
+ ultimately have "b * d = (a * c) * (b' * d')" by (simp add: mult_ac)
+ then show ?thesis ..
+qed
+
+lemma dvd_mult_left: "a * b dvd c \<Longrightarrow> a dvd c"
+by (simp add: dvd_def mult_assoc, blast)
+
+lemma dvd_mult_right: "a * b dvd c \<Longrightarrow> b dvd c"
+ unfolding mult_ac [of a] by (rule dvd_mult_left)
+
+lemma dvd_0_left: "0 dvd a \<Longrightarrow> a = 0"
+by simp
+
+lemma dvd_add[simp]:
+ assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)"
+proof -
+ from `a dvd b` obtain b' where "b = a * b'" ..
+ moreover from `a dvd c` obtain c' where "c = a * c'" ..
+ ultimately have "b + c = a * (b' + c')" by (simp add: right_distrib)
+ then show ?thesis ..
+qed
+
+end
+
+
+class no_zero_divisors = zero + times +
+ assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
+
+class semiring_1_cancel = semiring + cancel_comm_monoid_add
+ + zero_neq_one + monoid_mult
+begin
+
+subclass semiring_0_cancel ..
+
+subclass semiring_1 ..
+
+end
+
+class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add
+ + zero_neq_one + comm_monoid_mult
+begin
+
+subclass semiring_1_cancel ..
+subclass comm_semiring_0_cancel ..
+subclass comm_semiring_1 ..
+
+end
+
+class ring = semiring + ab_group_add
+begin
+
+subclass semiring_0_cancel ..
+
+text {* Distribution rules *}
+
+lemma minus_mult_left: "- (a * b) = - a * b"
+by (rule minus_unique) (simp add: left_distrib [symmetric])
+
+lemma minus_mult_right: "- (a * b) = a * - b"
+by (rule minus_unique) (simp add: right_distrib [symmetric])
+
+text{*Extract signs from products*}
+lemmas mult_minus_left [simp, noatp] = minus_mult_left [symmetric]
+lemmas mult_minus_right [simp,noatp] = minus_mult_right [symmetric]
+
+lemma minus_mult_minus [simp]: "- a * - b = a * b"
+by simp
+
+lemma minus_mult_commute: "- a * b = a * - b"
+by simp
+
+lemma right_diff_distrib[algebra_simps]: "a * (b - c) = a * b - a * c"
+by (simp add: right_distrib diff_minus)
+
+lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c"
+by (simp add: left_distrib diff_minus)
+
+lemmas ring_distribs[noatp] =
+ right_distrib left_distrib left_diff_distrib right_diff_distrib
+
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps[noatp] = algebra_simps
+
+lemma eq_add_iff1:
+ "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
+by (simp add: algebra_simps)
+
+lemma eq_add_iff2:
+ "a * e + c = b * e + d \<longleftrightarrow> c = (b - a) * e + d"
+by (simp add: algebra_simps)
+
+end
+
+lemmas ring_distribs[noatp] =
+ right_distrib left_distrib left_diff_distrib right_diff_distrib
+
+class comm_ring = comm_semiring + ab_group_add
+begin
+
+subclass ring ..
+subclass comm_semiring_0_cancel ..
+
+end
+
+class ring_1 = ring + zero_neq_one + monoid_mult
+begin
+
+subclass semiring_1_cancel ..
+
+end
+
+class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
+ (*previously ring*)
+begin
+
+subclass ring_1 ..
+subclass comm_semiring_1_cancel ..
+
+lemma dvd_minus_iff [simp]: "x dvd - y \<longleftrightarrow> x dvd y"
+proof
+ assume "x dvd - y"
+ then have "x dvd - 1 * - y" by (rule dvd_mult)
+ then show "x dvd y" by simp
+next
+ assume "x dvd y"
+ then have "x dvd - 1 * y" by (rule dvd_mult)
+ then show "x dvd - y" by simp
+qed
+
+lemma minus_dvd_iff [simp]: "- x dvd y \<longleftrightarrow> x dvd y"
+proof
+ assume "- x dvd y"
+ then obtain k where "y = - x * k" ..
+ then have "y = x * - k" by simp
+ then show "x dvd y" ..
+next
+ assume "x dvd y"
+ then obtain k where "y = x * k" ..
+ then have "y = - x * - k" by simp
+ then show "- x dvd y" ..
+qed
+
+lemma dvd_diff[simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
+by (simp add: diff_minus dvd_minus_iff)
+
+end
+
+class ring_no_zero_divisors = ring + no_zero_divisors
+begin
+
+lemma mult_eq_0_iff [simp]:
+ shows "a * b = 0 \<longleftrightarrow> (a = 0 \<or> b = 0)"
+proof (cases "a = 0 \<or> b = 0")
+ case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
+ then show ?thesis using no_zero_divisors by simp
+next
+ case True then show ?thesis by auto
+qed
+
+text{*Cancellation of equalities with a common factor*}
+lemma mult_cancel_right [simp, noatp]:
+ "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
+proof -
+ have "(a * c = b * c) = ((a - b) * c = 0)"
+ by (simp add: algebra_simps right_minus_eq)
+ thus ?thesis by (simp add: disj_commute right_minus_eq)
+qed
+
+lemma mult_cancel_left [simp, noatp]:
+ "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
+proof -
+ have "(c * a = c * b) = (c * (a - b) = 0)"
+ by (simp add: algebra_simps right_minus_eq)
+ thus ?thesis by (simp add: right_minus_eq)
+qed
+
+end
+
+class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
+begin
+
+lemma mult_cancel_right1 [simp]:
+ "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
+by (insert mult_cancel_right [of 1 c b], force)
+
+lemma mult_cancel_right2 [simp]:
+ "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
+by (insert mult_cancel_right [of a c 1], simp)
+
+lemma mult_cancel_left1 [simp]:
+ "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
+by (insert mult_cancel_left [of c 1 b], force)
+
+lemma mult_cancel_left2 [simp]:
+ "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
+by (insert mult_cancel_left [of c a 1], simp)
+
+end
+
+class idom = comm_ring_1 + no_zero_divisors
+begin
+
+subclass ring_1_no_zero_divisors ..
+
+lemma square_eq_iff: "a * a = b * b \<longleftrightarrow> (a = b \<or> a = - b)"
+proof
+ assume "a * a = b * b"
+ then have "(a - b) * (a + b) = 0"
+ by (simp add: algebra_simps)
+ then show "a = b \<or> a = - b"
+ by (simp add: right_minus_eq eq_neg_iff_add_eq_0)
+next
+ assume "a = b \<or> a = - b"
+ then show "a * a = b * b" by auto
+qed
+
+lemma dvd_mult_cancel_right [simp]:
+ "a * c dvd b * c \<longleftrightarrow> c = 0 \<or> a dvd b"
+proof -
+ have "a * c dvd b * c \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
+ unfolding dvd_def by (simp add: mult_ac)
+ also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
+ unfolding dvd_def by simp
+ finally show ?thesis .
+qed
+
+lemma dvd_mult_cancel_left [simp]:
+ "c * a dvd c * b \<longleftrightarrow> c = 0 \<or> a dvd b"
+proof -
+ have "c * a dvd c * b \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
+ unfolding dvd_def by (simp add: mult_ac)
+ also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
+ unfolding dvd_def by simp
+ finally show ?thesis .
+qed
+
+end
+
+class division_ring = ring_1 + inverse +
+ assumes left_inverse [simp]: "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
+ assumes right_inverse [simp]: "a \<noteq> 0 \<Longrightarrow> a * inverse a = 1"
+begin
+
+subclass ring_1_no_zero_divisors
+proof
+ fix a b :: 'a
+ assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
+ show "a * b \<noteq> 0"
+ proof
+ assume ab: "a * b = 0"
+ hence "0 = inverse a * (a * b) * inverse b" by simp
+ also have "\<dots> = (inverse a * a) * (b * inverse b)"
+ by (simp only: mult_assoc)
+ also have "\<dots> = 1" using a b by simp
+ finally show False by simp
+ qed
+qed
+
+lemma nonzero_imp_inverse_nonzero:
+ "a \<noteq> 0 \<Longrightarrow> inverse a \<noteq> 0"
+proof
+ assume ianz: "inverse a = 0"
+ assume "a \<noteq> 0"
+ hence "1 = a * inverse a" by simp
+ also have "... = 0" by (simp add: ianz)
+ finally have "1 = 0" .
+ thus False by (simp add: eq_commute)
+qed
+
+lemma inverse_zero_imp_zero:
+ "inverse a = 0 \<Longrightarrow> a = 0"
+apply (rule classical)
+apply (drule nonzero_imp_inverse_nonzero)
+apply auto
+done
+
+lemma inverse_unique:
+ assumes ab: "a * b = 1"
+ shows "inverse a = b"
+proof -
+ have "a \<noteq> 0" using ab by (cases "a = 0") simp_all
+ moreover have "inverse a * (a * b) = inverse a" by (simp add: ab)
+ ultimately show ?thesis by (simp add: mult_assoc [symmetric])
+qed
+
+lemma nonzero_inverse_minus_eq:
+ "a \<noteq> 0 \<Longrightarrow> inverse (- a) = - inverse a"
+by (rule inverse_unique) simp
+
+lemma nonzero_inverse_inverse_eq:
+ "a \<noteq> 0 \<Longrightarrow> inverse (inverse a) = a"
+by (rule inverse_unique) simp
+
+lemma nonzero_inverse_eq_imp_eq:
+ assumes "inverse a = inverse b" and "a \<noteq> 0" and "b \<noteq> 0"
+ shows "a = b"
+proof -
+ from `inverse a = inverse b`
+ have "inverse (inverse a) = inverse (inverse b)" by (rule arg_cong)
+ with `a \<noteq> 0` and `b \<noteq> 0` show "a = b"
+ by (simp add: nonzero_inverse_inverse_eq)
+qed
+
+lemma inverse_1 [simp]: "inverse 1 = 1"
+by (rule inverse_unique) simp
+
+lemma nonzero_inverse_mult_distrib:
+ assumes "a \<noteq> 0" and "b \<noteq> 0"
+ shows "inverse (a * b) = inverse b * inverse a"
+proof -
+ have "a * (b * inverse b) * inverse a = 1" using assms by simp
+ hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult_assoc)
+ thus ?thesis by (rule inverse_unique)
+qed
+
+lemma division_ring_inverse_add:
+ "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a + inverse b = inverse a * (a + b) * inverse b"
+by (simp add: algebra_simps)
+
+lemma division_ring_inverse_diff:
+ "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a - inverse b = inverse a * (b - a) * inverse b"
+by (simp add: algebra_simps)
+
+end
+
+class mult_mono = times + zero + ord +
+ assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
+ assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
+
+class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add
+begin
+
+lemma mult_mono:
+ "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> c
+ \<Longrightarrow> a * c \<le> b * d"
+apply (erule mult_right_mono [THEN order_trans], assumption)
+apply (erule mult_left_mono, assumption)
+done
+
+lemma mult_mono':
+ "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c
+ \<Longrightarrow> a * c \<le> b * d"
+apply (rule mult_mono)
+apply (fast intro: order_trans)+
+done
+
+end
+
+class ordered_cancel_semiring = mult_mono + ordered_ab_semigroup_add
+ + semiring + cancel_comm_monoid_add
+begin
+
+subclass semiring_0_cancel ..
+subclass ordered_semiring ..
+
+lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
+using mult_left_mono [of zero b a] by simp
+
+lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
+using mult_left_mono [of b zero a] by simp
+
+lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
+using mult_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_nonpos_nonneg} *}
+lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
+by (drule mult_right_mono [of b zero], auto)
+
+lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> 0"
+by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
+
+end
+
+class linordered_semiring = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + mult_mono
+begin
+
+subclass ordered_cancel_semiring ..
+
+subclass ordered_comm_monoid_add ..
+
+lemma mult_left_less_imp_less:
+ "c * a < c * b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
+by (force simp add: mult_left_mono not_le [symmetric])
+
+lemma mult_right_less_imp_less:
+ "a * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
+by (force simp add: mult_right_mono not_le [symmetric])
+
+end
+
+class linordered_semiring_1 = linordered_semiring + semiring_1
+
+class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add +
+ assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
+ assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
+begin
+
+subclass semiring_0_cancel ..
+
+subclass linordered_semiring
+proof
+ fix a b c :: 'a
+ assume A: "a \<le> b" "0 \<le> c"
+ from A show "c * a \<le> c * b"
+ unfolding le_less
+ using mult_strict_left_mono by (cases "c = 0") auto
+ from A show "a * c \<le> b * c"
+ unfolding le_less
+ using mult_strict_right_mono by (cases "c = 0") auto
+qed
+
+lemma mult_left_le_imp_le:
+ "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
+by (force simp add: mult_strict_left_mono _not_less [symmetric])
+
+lemma mult_right_le_imp_le:
+ "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
+by (force simp add: mult_strict_right_mono not_less [symmetric])
+
+lemma mult_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
+using mult_strict_left_mono [of zero b a] by simp
+
+lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
+using mult_strict_left_mono [of b zero a] by simp
+
+lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
+using mult_strict_right_mono [of a zero b] by simp
+
+text {* Legacy - use @{text mult_neg_pos} *}
+lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
+by (drule mult_strict_right_mono [of b zero], auto)
+
+lemma zero_less_mult_pos:
+ "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
+apply (cases "b\<le>0")
+ apply (auto simp add: le_less not_less)
+apply (drule_tac mult_pos_neg [of a b])
+ apply (auto dest: less_not_sym)
+done
+
+lemma zero_less_mult_pos2:
+ "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
+apply (cases "b\<le>0")
+ apply (auto simp add: le_less not_less)
+apply (drule_tac mult_pos_neg2 [of a b])
+ apply (auto dest: less_not_sym)
+done
+
+text{*Strict monotonicity in both arguments*}
+lemma mult_strict_mono:
+ assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
+ shows "a * c < b * d"
+ using assms apply (cases "c=0")
+ apply (simp add: mult_pos_pos)
+ apply (erule mult_strict_right_mono [THEN less_trans])
+ apply (force simp add: le_less)
+ apply (erule mult_strict_left_mono, assumption)
+ done
+
+text{*This weaker variant has more natural premises*}
+lemma mult_strict_mono':
+ assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
+ shows "a * c < b * d"
+by (rule mult_strict_mono) (insert assms, auto)
+
+lemma mult_less_le_imp_less:
+ assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
+ shows "a * c < b * d"
+ using assms apply (subgoal_tac "a * c < b * c")
+ apply (erule less_le_trans)
+ apply (erule mult_left_mono)
+ apply simp
+ apply (erule mult_strict_right_mono)
+ apply assumption
+ done
+
+lemma mult_le_less_imp_less:
+ assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
+ shows "a * c < b * d"
+ using assms apply (subgoal_tac "a * c \<le> b * c")
+ apply (erule le_less_trans)
+ apply (erule mult_strict_left_mono)
+ apply simp
+ apply (erule mult_right_mono)
+ apply simp
+ done
+
+lemma mult_less_imp_less_left:
+ assumes less: "c * a < c * b" and nonneg: "0 \<le> c"
+ shows "a < b"
+proof (rule ccontr)
+ assume "\<not> a < b"
+ hence "b \<le> a" by (simp add: linorder_not_less)
+ hence "c * b \<le> c * a" using nonneg by (rule mult_left_mono)
+ with this and less show False by (simp add: not_less [symmetric])
+qed
+
+lemma mult_less_imp_less_right:
+ assumes less: "a * c < b * c" and nonneg: "0 \<le> c"
+ shows "a < b"
+proof (rule ccontr)
+ assume "\<not> a < b"
+ hence "b \<le> a" by (simp add: linorder_not_less)
+ hence "b * c \<le> a * c" using nonneg by (rule mult_right_mono)
+ with this and less show False by (simp add: not_less [symmetric])
+qed
+
+end
+
+class linlinordered_semiring_1_strict = linordered_semiring_strict + semiring_1
+
+class mult_mono1 = times + zero + ord +
+ assumes mult_mono1: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
+
+class ordered_comm_semiring = comm_semiring_0
+ + ordered_ab_semigroup_add + mult_mono1
+begin
+
+subclass ordered_semiring
+proof
+ fix a b c :: 'a
+ assume "a \<le> b" "0 \<le> c"
+ thus "c * a \<le> c * b" by (rule mult_mono1)
+ thus "a * c \<le> b * c" by (simp only: mult_commute)
+qed
+
+end
+
+class ordered_cancel_comm_semiring = comm_semiring_0_cancel
+ + ordered_ab_semigroup_add + mult_mono1
+begin
+
+subclass ordered_comm_semiring ..
+subclass ordered_cancel_semiring ..
+
+end
+
+class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add +
+ assumes mult_strict_left_mono_comm: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
+begin
+
+subclass linordered_semiring_strict
+proof
+ fix a b c :: 'a
+ assume "a < b" "0 < c"
+ thus "c * a < c * b" by (rule mult_strict_left_mono_comm)
+ thus "a * c < b * c" by (simp only: mult_commute)
+qed
+
+subclass ordered_cancel_comm_semiring
+proof
+ fix a b c :: 'a
+ assume "a \<le> b" "0 \<le> c"
+ thus "c * a \<le> c * b"
+ unfolding le_less
+ using mult_strict_left_mono by (cases "c = 0") auto
+qed
+
+end
+
+class ordered_ring = ring + ordered_cancel_semiring
+begin
+
+subclass ordered_ab_group_add ..
+
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps[noatp] = algebra_simps
+
+lemma less_add_iff1:
+ "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
+by (simp add: algebra_simps)
+
+lemma less_add_iff2:
+ "a * e + c < b * e + d \<longleftrightarrow> c < (b - a) * e + d"
+by (simp add: algebra_simps)
+
+lemma le_add_iff1:
+ "a * e + c \<le> b * e + d \<longleftrightarrow> (a - b) * e + c \<le> d"
+by (simp add: algebra_simps)
+
+lemma le_add_iff2:
+ "a * e + c \<le> b * e + d \<longleftrightarrow> c \<le> (b - a) * e + d"
+by (simp add: algebra_simps)
+
+lemma mult_left_mono_neg:
+ "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
+ apply (drule mult_left_mono [of _ _ "uminus c"])
+ apply (simp_all add: minus_mult_left [symmetric])
+ done
+
+lemma mult_right_mono_neg:
+ "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
+ apply (drule mult_right_mono [of _ _ "uminus c"])
+ apply (simp_all add: minus_mult_right [symmetric])
+ done
+
+lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
+using mult_right_mono_neg [of a zero b] by simp
+
+lemma split_mult_pos_le:
+ "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
+by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos)
+
+end
+
+class abs_if = minus + uminus + ord + zero + abs +
+ assumes abs_if: "\<bar>a\<bar> = (if a < 0 then - a else a)"
+
+class sgn_if = minus + uminus + zero + one + ord + sgn +
+ assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
+
+lemma (in sgn_if) sgn0[simp]: "sgn 0 = 0"
+by(simp add:sgn_if)
+
+class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if
+begin
+
+subclass ordered_ring ..
+
+subclass ordered_ab_group_add_abs
+proof
+ fix a b
+ show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
+ by (auto simp add: abs_if not_less neg_less_eq_nonneg less_eq_neg_nonpos)
+ (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric]
+ neg_less_eq_nonneg less_eq_neg_nonpos, auto intro: add_nonneg_nonneg,
+ auto intro!: less_imp_le add_neg_neg)
+qed (auto simp add: abs_if less_eq_neg_nonpos neg_equal_zero)
+
+end
+
+(* The "strict" suffix can be seen as describing the combination of linordered_ring and no_zero_divisors.
+ Basically, linordered_ring + no_zero_divisors = linordered_ring_strict.
+ *)
+class linordered_ring_strict = ring + linordered_semiring_strict
+ + ordered_ab_group_add + abs_if
+begin
+
+subclass linordered_ring ..
+
+lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
+using mult_strict_left_mono [of b a "- c"] by simp
+
+lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
+using mult_strict_right_mono [of b a "- c"] by simp
+
+lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
+using mult_strict_right_mono_neg [of a zero b] by simp
+
+subclass ring_no_zero_divisors
+proof
+ fix a b
+ assume "a \<noteq> 0" then have A: "a < 0 \<or> 0 < a" by (simp add: neq_iff)
+ assume "b \<noteq> 0" then have B: "b < 0 \<or> 0 < b" by (simp add: neq_iff)
+ have "a * b < 0 \<or> 0 < a * b"
+ proof (cases "a < 0")
+ case True note A' = this
+ show ?thesis proof (cases "b < 0")
+ case True with A'
+ show ?thesis by (auto dest: mult_neg_neg)
+ next
+ case False with B have "0 < b" by auto
+ with A' show ?thesis by (auto dest: mult_strict_right_mono)
+ qed
+ next
+ case False with A have A': "0 < a" by auto
+ show ?thesis proof (cases "b < 0")
+ case True with A'
+ show ?thesis by (auto dest: mult_strict_right_mono_neg)
+ next
+ case False with B have "0 < b" by auto
+ with A' show ?thesis by (auto dest: mult_pos_pos)
+ qed
+ qed
+ then show "a * b \<noteq> 0" by (simp add: neq_iff)
+qed
+
+lemma zero_less_mult_iff:
+ "0 < a * b \<longleftrightarrow> 0 < a \<and> 0 < b \<or> a < 0 \<and> b < 0"
+ apply (auto simp add: mult_pos_pos mult_neg_neg)
+ apply (simp_all add: not_less le_less)
+ apply (erule disjE) apply assumption defer
+ apply (erule disjE) defer apply (drule sym) apply simp
+ apply (erule disjE) defer apply (drule sym) apply simp
+ apply (erule disjE) apply assumption apply (drule sym) apply simp
+ apply (drule sym) apply simp
+ apply (blast dest: zero_less_mult_pos)
+ apply (blast dest: zero_less_mult_pos2)
+ done
+
+lemma zero_le_mult_iff:
+ "0 \<le> a * b \<longleftrightarrow> 0 \<le> a \<and> 0 \<le> b \<or> a \<le> 0 \<and> b \<le> 0"
+by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff)
+
+lemma mult_less_0_iff:
+ "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
+ apply (insert zero_less_mult_iff [of "-a" b])
+ apply (force simp add: minus_mult_left[symmetric])
+ done
+
+lemma mult_le_0_iff:
+ "a * b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
+ apply (insert zero_le_mult_iff [of "-a" b])
+ apply (force simp add: minus_mult_left[symmetric])
+ done
+
+lemma zero_le_square [simp]: "0 \<le> a * a"
+by (simp add: zero_le_mult_iff linear)
+
+lemma not_square_less_zero [simp]: "\<not> (a * a < 0)"
+by (simp add: not_less)
+
+text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
+ also with the relations @{text "\<le>"} and equality.*}
+
+text{*These ``disjunction'' versions produce two cases when the comparison is
+ an assumption, but effectively four when the comparison is a goal.*}
+
+lemma mult_less_cancel_right_disj:
+ "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
+ apply (cases "c = 0")
+ apply (auto simp add: neq_iff mult_strict_right_mono
+ mult_strict_right_mono_neg)
+ apply (auto simp add: not_less
+ not_le [symmetric, of "a*c"]
+ not_le [symmetric, of a])
+ apply (erule_tac [!] notE)
+ apply (auto simp add: less_imp_le mult_right_mono
+ mult_right_mono_neg)
+ done
+
+lemma mult_less_cancel_left_disj:
+ "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
+ apply (cases "c = 0")
+ apply (auto simp add: neq_iff mult_strict_left_mono
+ mult_strict_left_mono_neg)
+ apply (auto simp add: not_less
+ not_le [symmetric, of "c*a"]
+ not_le [symmetric, of a])
+ apply (erule_tac [!] notE)
+ apply (auto simp add: less_imp_le mult_left_mono
+ mult_left_mono_neg)
+ done
+
+text{*The ``conjunction of implication'' lemmas produce two cases when the
+comparison is a goal, but give four when the comparison is an assumption.*}
+
+lemma mult_less_cancel_right:
+ "a * c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
+ using mult_less_cancel_right_disj [of a c b] by auto
+
+lemma mult_less_cancel_left:
+ "c * a < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
+ using mult_less_cancel_left_disj [of c a b] by auto
+
+lemma mult_le_cancel_right:
+ "a * c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
+
+lemma mult_le_cancel_left:
+ "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
+
+lemma mult_le_cancel_left_pos:
+ "0 < c \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> a \<le> b"
+by (auto simp: mult_le_cancel_left)
+
+lemma mult_le_cancel_left_neg:
+ "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
+by (auto simp: mult_le_cancel_left)
+
+lemma mult_less_cancel_left_pos:
+ "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
+by (auto simp: mult_less_cancel_left)
+
+lemma mult_less_cancel_left_neg:
+ "c < 0 \<Longrightarrow> c * a < c * b \<longleftrightarrow> b < a"
+by (auto simp: mult_less_cancel_left)
+
+end
+
+text{*Legacy - use @{text algebra_simps} *}
+lemmas ring_simps[noatp] = algebra_simps
+
+lemmas mult_sign_intros =
+ mult_nonneg_nonneg mult_nonneg_nonpos
+ mult_nonpos_nonneg mult_nonpos_nonpos
+ mult_pos_pos mult_pos_neg
+ mult_neg_pos mult_neg_neg
+
+class ordered_comm_ring = comm_ring + ordered_comm_semiring
+begin
+
+subclass ordered_ring ..
+subclass ordered_cancel_comm_semiring ..
+
+end
+
+class linordered_semidom = comm_semiring_1_cancel + linordered_comm_semiring_strict +
+ (*previously linordered_semiring*)
+ assumes zero_less_one [simp]: "0 < 1"
+begin
+
+lemma pos_add_strict:
+ shows "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
+ using add_strict_mono [of zero a b c] by simp
+
+lemma zero_le_one [simp]: "0 \<le> 1"
+by (rule zero_less_one [THEN less_imp_le])
+
+lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
+by (simp add: not_le)
+
+lemma not_one_less_zero [simp]: "\<not> 1 < 0"
+by (simp add: not_less)
+
+lemma less_1_mult:
+ assumes "1 < m" and "1 < n"
+ shows "1 < m * n"
+ using assms mult_strict_mono [of 1 m 1 n]
+ by (simp add: less_trans [OF zero_less_one])
+
+end
+
+class linordered_idom = comm_ring_1 +
+ linordered_comm_semiring_strict + ordered_ab_group_add +
+ abs_if + sgn_if
+ (*previously linordered_ring*)
+begin
+
+subclass linordered_ring_strict ..
+subclass ordered_comm_ring ..
+subclass idom ..
+
+subclass linordered_semidom
+proof
+ have "0 \<le> 1 * 1" by (rule zero_le_square)
+ thus "0 < 1" by (simp add: le_less)
+qed
+
+lemma linorder_neqE_linordered_idom:
+ assumes "x \<noteq> y" obtains "x < y" | "y < x"
+ using assms by (rule neqE)
+
+text {* These cancellation simprules also produce two cases when the comparison is a goal. *}
+
+lemma mult_le_cancel_right1:
+ "c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
+by (insert mult_le_cancel_right [of 1 c b], simp)
+
+lemma mult_le_cancel_right2:
+ "a * c \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
+by (insert mult_le_cancel_right [of a c 1], simp)
+
+lemma mult_le_cancel_left1:
+ "c \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
+by (insert mult_le_cancel_left [of c 1 b], simp)
+
+lemma mult_le_cancel_left2:
+ "c * a \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
+by (insert mult_le_cancel_left [of c a 1], simp)
+
+lemma mult_less_cancel_right1:
+ "c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
+by (insert mult_less_cancel_right [of 1 c b], simp)
+
+lemma mult_less_cancel_right2:
+ "a * c < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
+by (insert mult_less_cancel_right [of a c 1], simp)
+
+lemma mult_less_cancel_left1:
+ "c < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
+by (insert mult_less_cancel_left [of c 1 b], simp)
+
+lemma mult_less_cancel_left2:
+ "c * a < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
+by (insert mult_less_cancel_left [of c a 1], simp)
+
+lemma sgn_sgn [simp]:
+ "sgn (sgn a) = sgn a"
+unfolding sgn_if by simp
+
+lemma sgn_0_0:
+ "sgn a = 0 \<longleftrightarrow> a = 0"
+unfolding sgn_if by simp
+
+lemma sgn_1_pos:
+ "sgn a = 1 \<longleftrightarrow> a > 0"
+unfolding sgn_if by (simp add: neg_equal_zero)
+
+lemma sgn_1_neg:
+ "sgn a = - 1 \<longleftrightarrow> a < 0"
+unfolding sgn_if by (auto simp add: equal_neg_zero)
+
+lemma sgn_pos [simp]:
+ "0 < a \<Longrightarrow> sgn a = 1"
+unfolding sgn_1_pos .
+
+lemma sgn_neg [simp]:
+ "a < 0 \<Longrightarrow> sgn a = - 1"
+unfolding sgn_1_neg .
+
+lemma sgn_times:
+ "sgn (a * b) = sgn a * sgn b"
+by (auto simp add: sgn_if zero_less_mult_iff)
+
+lemma abs_sgn: "abs k = k * sgn k"
+unfolding sgn_if abs_if by auto
+
+lemma sgn_greater [simp]:
+ "0 < sgn a \<longleftrightarrow> 0 < a"
+ unfolding sgn_if by auto
+
+lemma sgn_less [simp]:
+ "sgn a < 0 \<longleftrightarrow> a < 0"
+ unfolding sgn_if by auto
+
+lemma abs_dvd_iff [simp]: "(abs m) dvd k \<longleftrightarrow> m dvd k"
+ by (simp add: abs_if)
+
+lemma dvd_abs_iff [simp]: "m dvd (abs k) \<longleftrightarrow> m dvd k"
+ by (simp add: abs_if)
+
+lemma dvd_if_abs_eq:
+ "abs l = abs (k) \<Longrightarrow> l dvd k"
+by(subst abs_dvd_iff[symmetric]) simp
+
+end
+
+text {* Simprules for comparisons where common factors can be cancelled. *}
+
+lemmas mult_compare_simps[noatp] =
+ mult_le_cancel_right mult_le_cancel_left
+ mult_le_cancel_right1 mult_le_cancel_right2
+ mult_le_cancel_left1 mult_le_cancel_left2
+ mult_less_cancel_right mult_less_cancel_left
+ mult_less_cancel_right1 mult_less_cancel_right2
+ mult_less_cancel_left1 mult_less_cancel_left2
+ mult_cancel_right mult_cancel_left
+ mult_cancel_right1 mult_cancel_right2
+ mult_cancel_left1 mult_cancel_left2
+
+-- {* FIXME continue localization here *}
+
+subsection {* Reasoning about inequalities with division *}
+
+lemma mult_right_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
+ ==> x * y <= x"
+by (auto simp add: mult_compare_simps)
+
+lemma mult_left_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
+ ==> y * x <= x"
+by (auto simp add: mult_compare_simps)
+
+context linordered_semidom
+begin
+
+lemma less_add_one: "a < a + 1"
+proof -
+ have "a + 0 < a + 1"
+ by (blast intro: zero_less_one add_strict_left_mono)
+ thus ?thesis by simp
+qed
+
+lemma zero_less_two: "0 < 1 + 1"
+by (blast intro: less_trans zero_less_one less_add_one)
+
+end
+
+
+subsection {* Absolute Value *}
+
+context linordered_idom
+begin
+
+lemma mult_sgn_abs: "sgn x * abs x = x"
+ unfolding abs_if sgn_if by auto
+
+end
+
+lemma abs_one [simp]: "abs 1 = (1::'a::linordered_idom)"
+by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
+
+class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs +
+ assumes abs_eq_mult:
+ "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0) \<Longrightarrow> \<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
+
+context linordered_idom
+begin
+
+subclass ordered_ring_abs proof
+qed (auto simp add: abs_if not_less equal_neg_zero neg_equal_zero mult_less_0_iff)
+
+lemma abs_mult:
+ "abs (a * b) = abs a * abs b"
+ by (rule abs_eq_mult) auto
+
+lemma abs_mult_self:
+ "abs a * abs a = a * a"
+ by (simp add: abs_if)
+
+end
+
+lemma abs_mult_less:
+ "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::linordered_idom)"
+proof -
+ assume ac: "abs a < c"
+ hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
+ assume "abs b < d"
+ thus ?thesis by (simp add: ac cpos mult_strict_mono)
+qed
+
+lemmas eq_minus_self_iff[noatp] = equal_neg_zero
+
+lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::linordered_idom))"
+ unfolding order_less_le less_eq_neg_nonpos equal_neg_zero ..
+
+lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::linordered_idom))"
+apply (simp add: order_less_le abs_le_iff)
+apply (auto simp add: abs_if neg_less_eq_nonneg less_eq_neg_nonpos)
+done
+
+lemma abs_mult_pos: "(0::'a::linordered_idom) <= x ==>
+ (abs y) * x = abs (y * x)"
+ apply (subst abs_mult)
+ apply simp
+done
+
+code_modulename SML
+ Rings Arith
+
+code_modulename OCaml
+ Rings Arith
+
+code_modulename Haskell
+ Rings Arith
+
+end
--- a/src/HOL/SMT/Examples/SMT_Examples.certs Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/SMT/Examples/SMT_Examples.certs Mon Feb 08 17:13:45 2010 +0100
@@ -6677,15 +6677,25 @@
[mp #29 #65]: false
unsat
-rOkYPs+Q++z/M3OPR/88JQ 1272
+rOkYPs+Q++z/M3OPR/88JQ 1654
#2 := false
#55 := 0::int
#7 := 2::int
decl uf_1 :: int
#4 := uf_1
#8 := (mod uf_1 2::int)
-#58 := (>= #8 0::int)
-#61 := (not #58)
+#67 := (>= #8 0::int)
+#1 := true
+#156 := (iff #67 true)
+#158 := (iff #156 #67)
+#159 := [rewrite]: #158
+#28 := [true-axiom]: true
+#153 := (or false #67)
+#154 := [th-lemma]: #153
+#155 := [unit-resolution #154 #28]: #67
+#157 := [iff-true #155]: #156
+#160 := [mp #157 #159]: #67
+#70 := (not #67)
#5 := 1::int
#9 := (* 2::int #8)
#10 := (+ #9 1::int)
@@ -6693,24 +6703,35 @@
#6 := (+ uf_1 1::int)
#12 := (<= #6 #11)
#13 := (not #12)
-#66 := (iff #13 #61)
+#77 := (iff #13 #70)
#39 := (+ uf_1 #9)
#40 := (+ 1::int #39)
#30 := (+ 1::int uf_1)
#45 := (<= #30 #40)
#48 := (not #45)
-#64 := (iff #48 #61)
+#75 := (iff #48 #70)
+#58 := (* 1::int #8)
+#59 := (>= #58 0::int)
+#62 := (not #59)
+#71 := (iff #62 #70)
+#68 := (iff #59 #67)
+#65 := (= #58 #8)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#72 := [monotonicity #69]: #71
+#73 := (iff #48 #62)
#56 := (>= #9 0::int)
#51 := (not #56)
-#62 := (iff #51 #61)
-#59 := (iff #56 #58)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
+#63 := (iff #51 #62)
+#60 := (iff #56 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
#52 := (iff #48 #51)
#53 := (iff #45 #56)
#54 := [rewrite]: #53
#57 := [monotonicity #54]: #52
-#65 := [trans #57 #63]: #64
+#74 := [trans #57 #64]: #73
+#76 := [trans #74 #72]: #75
#49 := (iff #13 #48)
#46 := (iff #12 #45)
#43 := (= #11 #40)
@@ -6727,40 +6748,38 @@
#32 := [rewrite]: #31
#47 := [monotonicity #32 #44]: #46
#50 := [monotonicity #47]: #49
-#67 := [trans #50 #65]: #66
+#78 := [trans #50 #76]: #77
#29 := [asserted]: #13
-#68 := [mp #29 #67]: #61
-#1 := true
-#28 := [true-axiom]: true
-#142 := (or false #58)
-#143 := [th-lemma]: #142
-#144 := [unit-resolution #143 #28]: #58
-[unit-resolution #144 #68]: false
-unsat
-
-oSBTeiF6GKDeHPsMxXHXtQ 1161
+#79 := [mp #29 #78]: #70
+[unit-resolution #79 #160]: false
+unsat
+
+oSBTeiF6GKDeHPsMxXHXtQ 1064
#2 := false
#5 := 2::int
decl uf_1 :: int
#4 := uf_1
#6 := (mod uf_1 2::int)
-#55 := (>= #6 2::int)
+#122 := (>= #6 2::int)
+#123 := (not #122)
+#1 := true
+#27 := [true-axiom]: true
+#133 := (or false #123)
+#134 := [th-lemma]: #133
+#135 := [unit-resolution #134 #27]: #123
#9 := 3::int
+#29 := (* 2::int #6)
+#48 := (>= #29 3::int)
#10 := (+ uf_1 3::int)
#7 := (+ #6 #6)
#8 := (+ uf_1 #7)
#11 := (< #8 #10)
#12 := (not #11)
-#60 := (iff #12 #55)
+#55 := (iff #12 #48)
#35 := (+ 3::int uf_1)
-#29 := (* 2::int #6)
#32 := (+ uf_1 #29)
#38 := (< #32 #35)
#41 := (not #38)
-#58 := (iff #41 #55)
-#48 := (>= #29 3::int)
-#56 := (iff #48 #55)
-#57 := [rewrite]: #56
#53 := (iff #41 #48)
#46 := (not #48)
#45 := (not #46)
@@ -6771,7 +6790,6 @@
#44 := [rewrite]: #47
#50 := [monotonicity #44]: #49
#54 := [trans #50 #52]: #53
-#59 := [trans #54 #57]: #58
#42 := (iff #12 #41)
#39 := (iff #11 #38)
#36 := (= #10 #35)
@@ -6782,16 +6800,10 @@
#34 := [monotonicity #31]: #33
#40 := [monotonicity #34 #37]: #39
#43 := [monotonicity #40]: #42
-#61 := [trans #43 #59]: #60
+#56 := [trans #43 #54]: #55
#28 := [asserted]: #12
-#62 := [mp #28 #61]: #55
-#127 := (not #55)
-#1 := true
-#27 := [true-axiom]: true
-#137 := (or false #127)
-#138 := [th-lemma]: #137
-#139 := [unit-resolution #138 #27]: #127
-[unit-resolution #139 #62]: false
+#57 := [mp #28 #56]: #48
+[th-lemma #57 #135]: false
unsat
hqH9yBHvmZgES3pBkvsqVQ 2715
@@ -7344,52 +7356,99 @@
[mp #30 #96]: false
unsat
-sHpY0IFBgZUNhP56aRB+/w 1765
+sHpY0IFBgZUNhP56aRB+/w 2968
#2 := false
+#9 := 1::int
+decl ?x1!1 :: int
+#91 := ?x1!1
+#68 := -2::int
+#129 := (* -2::int ?x1!1)
+decl ?x2!0 :: int
+#90 := ?x2!0
+#7 := 2::int
+#128 := (* 2::int ?x2!0)
+#130 := (+ #128 #129)
+#131 := (<= #130 1::int)
+#136 := (not #131)
+#55 := 0::int
+#53 := -1::int
+#115 := (* -1::int ?x1!1)
+#116 := (+ ?x2!0 #115)
+#117 := (<= #116 0::int)
+#139 := (or #117 #136)
+#142 := (not #139)
+#92 := (* -2::int ?x2!0)
+#93 := (* 2::int ?x1!1)
+#94 := (+ #93 #92)
+#95 := (>= #94 -1::int)
+#96 := (not #95)
+#97 := (* -1::int ?x2!0)
+#98 := (+ ?x1!1 #97)
+#99 := (>= #98 0::int)
+#100 := (or #99 #96)
+#101 := (not #100)
+#143 := (iff #101 #142)
+#140 := (iff #100 #139)
+#137 := (iff #96 #136)
+#134 := (iff #95 #131)
+#122 := (+ #92 #93)
+#125 := (>= #122 -1::int)
+#132 := (iff #125 #131)
+#133 := [rewrite]: #132
+#126 := (iff #95 #125)
+#123 := (= #94 #122)
+#124 := [rewrite]: #123
+#127 := [monotonicity #124]: #126
+#135 := [trans #127 #133]: #134
+#138 := [monotonicity #135]: #137
+#120 := (iff #99 #117)
+#109 := (+ #97 ?x1!1)
+#112 := (>= #109 0::int)
+#118 := (iff #112 #117)
+#119 := [rewrite]: #118
+#113 := (iff #99 #112)
+#110 := (= #98 #109)
+#111 := [rewrite]: #110
+#114 := [monotonicity #111]: #113
+#121 := [trans #114 #119]: #120
+#141 := [monotonicity #121 #138]: #140
+#144 := [monotonicity #141]: #143
#5 := (:var 0 int)
-#7 := 2::int
-#11 := (* 2::int #5)
-#9 := 1::int
+#71 := (* -2::int #5)
#4 := (:var 1 int)
#8 := (* 2::int #4)
+#72 := (+ #8 #71)
+#70 := (>= #72 -1::int)
+#69 := (not #70)
+#57 := (* -1::int #5)
+#58 := (+ #4 #57)
+#56 := (>= #58 0::int)
+#75 := (or #56 #69)
+#78 := (forall (vars (?x1 int) (?x2 int)) #75)
+#81 := (not #78)
+#102 := (~ #81 #101)
+#103 := [sk]: #102
+#11 := (* 2::int #5)
#10 := (+ #8 1::int)
#12 := (< #10 #11)
#6 := (< #4 #5)
#13 := (implies #6 #12)
#14 := (forall (vars (?x1 int) (?x2 int)) #13)
#15 := (not #14)
-#91 := (iff #15 false)
+#84 := (iff #15 #81)
#32 := (+ 1::int #8)
#35 := (< #32 #11)
#41 := (not #6)
#42 := (or #41 #35)
#47 := (forall (vars (?x1 int) (?x2 int)) #42)
#50 := (not #47)
-#89 := (iff #50 false)
-#1 := true
-#84 := (not true)
-#87 := (iff #84 false)
-#88 := [rewrite]: #87
-#85 := (iff #50 #84)
-#82 := (iff #47 true)
-#77 := (forall (vars (?x1 int) (?x2 int)) true)
-#80 := (iff #77 true)
-#81 := [elim-unused]: #80
-#78 := (iff #47 #77)
-#75 := (iff #42 true)
-#55 := 0::int
-#53 := -1::int
-#57 := (* -1::int #5)
-#58 := (+ #4 #57)
-#56 := (>= #58 0::int)
+#82 := (iff #50 #81)
+#79 := (iff #47 #78)
+#76 := (iff #42 #75)
+#73 := (iff #35 #69)
+#74 := [rewrite]: #73
+#66 := (iff #41 #56)
#54 := (not #56)
-#69 := (or #56 #54)
-#73 := (iff #69 true)
-#74 := [rewrite]: #73
-#71 := (iff #42 #69)
-#70 := (iff #35 #54)
-#68 := [rewrite]: #70
-#66 := (iff #41 #56)
#61 := (not #54)
#64 := (iff #61 #56)
#65 := [rewrite]: #64
@@ -7398,12 +7457,9 @@
#60 := [rewrite]: #59
#63 := [monotonicity #60]: #62
#67 := [trans #63 #65]: #66
-#72 := [monotonicity #67 #68]: #71
-#76 := [trans #72 #74]: #75
-#79 := [quant-intro #76]: #78
-#83 := [trans #79 #81]: #82
-#86 := [monotonicity #83]: #85
-#90 := [trans #86 #88]: #89
+#77 := [monotonicity #67 #74]: #76
+#80 := [quant-intro #77]: #79
+#83 := [monotonicity #80]: #82
#51 := (iff #15 #50)
#48 := (iff #14 #47)
#45 := (iff #13 #42)
@@ -7419,54 +7475,89 @@
#46 := [trans #40 #44]: #45
#49 := [quant-intro #46]: #48
#52 := [monotonicity #49]: #51
-#92 := [trans #52 #90]: #91
+#85 := [trans #52 #83]: #84
#31 := [asserted]: #15
-[mp #31 #92]: false
-unsat
-
-7GSX+dyn9XwHWNcjJ4X1ww 1400
+#86 := [mp #31 #85]: #81
+#106 := [mp~ #86 #103]: #101
+#107 := [mp #106 #144]: #142
+#146 := [not-or-elim #107]: #131
+#108 := (not #117)
+#145 := [not-or-elim #107]: #108
+[th-lemma #145 #146]: false
+unsat
+
+7GSX+dyn9XwHWNcjJ4X1ww 2292
#2 := false
-#9 := (:var 0 int)
+#7 := 1::int
+decl ?x1!1 :: int
+#74 := ?x1!1
+#51 := -2::int
+#96 := (* -2::int ?x1!1)
+decl ?x2!0 :: int
+#73 := ?x2!0
#4 := 2::int
-#10 := (* 2::int #9)
-#7 := 1::int
+#95 := (* 2::int ?x2!0)
+#97 := (+ #95 #96)
+#166 := (<= #97 1::int)
+#94 := (= #97 1::int)
+#53 := -1::int
+#75 := (* -2::int ?x2!0)
+#76 := (* 2::int ?x1!1)
+#77 := (+ #76 #75)
+#78 := (= #77 -1::int)
+#79 := (not #78)
+#80 := (not #79)
+#110 := (iff #80 #94)
+#102 := (not #94)
+#105 := (not #102)
+#108 := (iff #105 #94)
+#109 := [rewrite]: #108
+#106 := (iff #80 #105)
+#103 := (iff #79 #102)
+#100 := (iff #78 #94)
+#88 := (+ #75 #76)
+#91 := (= #88 -1::int)
+#98 := (iff #91 #94)
+#99 := [rewrite]: #98
+#92 := (iff #78 #91)
+#89 := (= #77 #88)
+#90 := [rewrite]: #89
+#93 := [monotonicity #90]: #92
+#101 := [trans #93 #99]: #100
+#104 := [monotonicity #101]: #103
+#107 := [monotonicity #104]: #106
+#111 := [trans #107 #109]: #110
+#9 := (:var 0 int)
+#55 := (* -2::int #9)
#5 := (:var 1 int)
#6 := (* 2::int #5)
+#56 := (+ #6 #55)
+#54 := (= #56 -1::int)
+#58 := (not #54)
+#61 := (forall (vars (?x1 int) (?x2 int)) #58)
+#64 := (not #61)
+#81 := (~ #64 #80)
+#82 := [sk]: #81
+#10 := (* 2::int #9)
#8 := (+ #6 1::int)
#11 := (= #8 #10)
#12 := (not #11)
#13 := (forall (vars (?x1 int) (?x2 int)) #12)
#14 := (not #13)
-#74 := (iff #14 false)
+#67 := (iff #14 #64)
#31 := (+ 1::int #6)
#37 := (= #10 #31)
#42 := (not #37)
#45 := (forall (vars (?x1 int) (?x2 int)) #42)
#48 := (not #45)
-#72 := (iff #48 false)
-#1 := true
-#67 := (not true)
-#70 := (iff #67 false)
-#71 := [rewrite]: #70
-#68 := (iff #48 #67)
-#65 := (iff #45 true)
-#60 := (forall (vars (?x1 int) (?x2 int)) true)
-#63 := (iff #60 true)
-#64 := [elim-unused]: #63
-#61 := (iff #45 #60)
-#58 := (iff #42 true)
-#51 := (not false)
-#56 := (iff #51 true)
-#57 := [rewrite]: #56
-#52 := (iff #42 #51)
-#53 := (iff #37 false)
-#54 := [rewrite]: #53
-#55 := [monotonicity #54]: #52
-#59 := [trans #55 #57]: #58
-#62 := [quant-intro #59]: #61
-#66 := [trans #62 #64]: #65
-#69 := [monotonicity #66]: #68
-#73 := [trans #69 #71]: #72
+#65 := (iff #48 #64)
+#62 := (iff #45 #61)
+#59 := (iff #42 #58)
+#52 := (iff #37 #54)
+#57 := [rewrite]: #52
+#60 := [monotonicity #57]: #59
+#63 := [quant-intro #60]: #62
+#66 := [monotonicity #63]: #65
#49 := (iff #14 #48)
#46 := (iff #13 #45)
#43 := (iff #12 #42)
@@ -7482,9 +7573,19 @@
#44 := [monotonicity #41]: #43
#47 := [quant-intro #44]: #46
#50 := [monotonicity #47]: #49
-#75 := [trans #50 #73]: #74
+#68 := [trans #50 #66]: #67
#30 := [asserted]: #14
-[mp #30 #75]: false
+#69 := [mp #30 #68]: #64
+#85 := [mp~ #69 #82]: #80
+#86 := [mp #85 #111]: #94
+#168 := (or #102 #166)
+#169 := [th-lemma]: #168
+#170 := [unit-resolution #169 #86]: #166
+#167 := (>= #97 1::int)
+#171 := (or #102 #167)
+#172 := [th-lemma]: #171
+#173 := [unit-resolution #172 #86]: #167
+[th-lemma #173 #170]: false
unsat
oieid3+1h5s04LTQ9d796Q 2636
@@ -7960,38 +8061,62 @@
[unit-resolution #585 #307]: false
unsat
-4Dtb5Y1TTRPWZcHG9Griog 1594
+4Dtb5Y1TTRPWZcHG9Griog 2407
#2 := false
+#104 := -1::int
+decl ?x1!1 :: int
+#86 := ?x1!1
+#106 := -4::int
+#107 := (* -4::int ?x1!1)
+decl ?x2!0 :: int
+#85 := ?x2!0
+#7 := 6::int
+#105 := (* 6::int ?x2!0)
+#108 := (+ #105 #107)
+#168 := (<= #108 -1::int)
+#109 := (= #108 -1::int)
#12 := 1::int
+#33 := -6::int
+#87 := (* -6::int ?x2!0)
+#4 := 4::int
+#88 := (* 4::int ?x1!1)
+#89 := (+ #88 #87)
+#90 := (= #89 1::int)
+#112 := (iff #90 #109)
+#98 := (+ #87 #88)
+#101 := (= #98 1::int)
+#110 := (iff #101 #109)
+#111 := [rewrite]: #110
+#102 := (iff #90 #101)
+#99 := (= #89 #98)
+#100 := [rewrite]: #99
+#103 := [monotonicity #100]: #102
+#113 := [trans #103 #111]: #112
+#53 := (:var 0 int)
+#54 := (* -6::int #53)
#9 := (:var 1 int)
-#7 := 6::int
+#55 := (* 4::int #9)
+#56 := (+ #55 #54)
+#76 := (= #56 1::int)
+#74 := (exists (vars (?x1 int) (?x2 int)) #76)
+#91 := (~ #74 #90)
+#92 := [sk]: #91
#8 := (- 6::int)
#10 := (* #8 #9)
#5 := (:var 2 int)
-#4 := 4::int
#6 := (* 4::int #5)
#11 := (+ #6 #10)
#13 := (= #11 1::int)
#14 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #13)
#15 := (not #14)
#16 := (not #15)
-#82 := (iff #16 false)
-#53 := (:var 0 int)
-#33 := -6::int
-#54 := (* -6::int #53)
-#55 := (* 4::int #9)
-#56 := (+ #55 #54)
+#79 := (iff #16 #74)
#57 := (= 1::int #56)
#58 := (exists (vars (?x1 int) (?x2 int)) #57)
-#80 := (iff #58 false)
-#76 := (exists (vars (?x1 int) (?x2 int)) false)
-#78 := (iff #76 false)
-#79 := [elim-unused]: #78
-#77 := (iff #58 #76)
-#73 := (iff #57 false)
-#74 := [rewrite]: #73
-#75 := [quant-intro #74]: #77
-#81 := [trans #75 #79]: #80
+#77 := (iff #58 #74)
+#75 := (iff #57 #76)
+#73 := [rewrite]: #75
+#78 := [quant-intro #73]: #77
#71 := (iff #16 #58)
#63 := (not #58)
#66 := (not #63)
@@ -8025,9 +8150,20 @@
#65 := [monotonicity #62]: #64
#68 := [monotonicity #65]: #67
#72 := [trans #68 #70]: #71
-#83 := [trans #72 #81]: #82
+#80 := [trans #72 #78]: #79
#32 := [asserted]: #16
-[mp #32 #83]: false
+#81 := [mp #32 #80]: #74
+#95 := [mp~ #81 #92]: #90
+#96 := [mp #95 #113]: #109
+#170 := (not #109)
+#171 := (or #170 #168)
+#172 := [th-lemma]: #171
+#173 := [unit-resolution #172 #96]: #168
+#169 := (>= #108 -1::int)
+#174 := (or #170 #169)
+#175 := [th-lemma]: #174
+#176 := [unit-resolution #175 #96]: #169
+[th-lemma #176 #173]: false
unsat
dbOre63OdVavsqL3lG2ttw 2516
@@ -8445,12 +8581,12 @@
#46 := (+ #4 #45)
#44 := (>= #46 0::int)
#42 := (not #44)
-#57 := (or #44 #42)
-#61 := (iff #57 true)
+#60 := (or #44 #42)
+#61 := (iff #60 true)
#62 := [rewrite]: #61
-#59 := (iff #32 #57)
+#56 := (iff #32 #60)
#58 := (iff #11 #42)
-#56 := [rewrite]: #58
+#59 := [rewrite]: #58
#54 := (iff #31 #44)
#49 := (not #42)
#52 := (iff #49 #44)
@@ -8460,8 +8596,8 @@
#48 := [rewrite]: #47
#51 := [monotonicity #48]: #50
#55 := [trans #51 #53]: #54
-#60 := [monotonicity #55 #56]: #59
-#64 := [trans #60 #62]: #63
+#57 := [monotonicity #55 #59]: #56
+#64 := [trans #57 #62]: #63
#67 := [quant-intro #64]: #66
#71 := [trans #67 #69]: #70
#74 := [monotonicity #71]: #73
@@ -8764,7 +8900,7 @@
[mp #45 #150]: false
unsat
-iPZ87GNdi7uQw4ehEpbTPQ 6383
+iPZ87GNdi7uQw4ehEpbTPQ 7012
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -8779,19 +8915,19 @@
#295 := -1::int
#274 := (* -1::int #293)
#610 := (+ #24 #274)
-#594 := (<= #610 0::int)
+#315 := (<= #610 0::int)
#612 := (= #610 0::int)
-#606 := (>= #23 0::int)
-#237 := (= #293 0::int)
-#549 := (not #237)
-#588 := (<= #293 0::int)
-#457 := (not #588)
+#255 := (>= #23 0::int)
+#317 := (= #293 0::int)
+#522 := (not #317)
+#577 := (<= #293 0::int)
+#519 := (not #577)
#26 := 1::int
-#558 := (>= #293 1::int)
-#555 := (= #293 1::int)
+#553 := (>= #293 1::int)
+#552 := (= #293 1::int)
#27 := (uf_1 1::int)
-#589 := (uf_2 #27)
-#301 := (= #589 1::int)
+#420 := (uf_2 #27)
+#565 := (= #420 1::int)
#10 := (:var 0 int)
#12 := (uf_1 #10)
#626 := (pattern #12)
@@ -8847,57 +8983,57 @@
#87 := [mp #51 #86]: #82
#146 := [mp~ #87 #130]: #82
#632 := [mp #146 #631]: #627
-#609 := (not #627)
-#578 := (or #609 #301)
-#311 := (>= 1::int 0::int)
-#585 := (not #311)
-#586 := (= 1::int #589)
-#590 := (or #586 #585)
-#419 := (or #609 #590)
-#421 := (iff #419 #578)
-#564 := (iff #578 #578)
-#565 := [rewrite]: #564
-#577 := (iff #590 #301)
-#574 := (or #301 false)
-#571 := (iff #574 #301)
-#576 := [rewrite]: #571
-#575 := (iff #590 #574)
-#584 := (iff #585 false)
+#237 := (not #627)
+#442 := (or #237 #565)
+#578 := (>= 1::int 0::int)
+#419 := (not #578)
+#421 := (= 1::int #420)
+#563 := (or #421 #419)
+#443 := (or #237 #563)
+#550 := (iff #443 #442)
+#547 := (iff #442 #442)
+#548 := [rewrite]: #547
+#559 := (iff #563 #565)
+#554 := (or #565 false)
+#558 := (iff #554 #565)
+#556 := [rewrite]: #558
+#555 := (iff #563 #554)
+#400 := (iff #419 false)
#1 := true
-#582 := (not true)
-#583 := (iff #582 false)
-#580 := [rewrite]: #583
-#296 := (iff #585 #582)
-#303 := (iff #311 true)
-#581 := [rewrite]: #303
-#579 := [monotonicity #581]: #296
-#573 := [trans #579 #580]: #584
-#300 := (iff #586 #301)
-#302 := [rewrite]: #300
-#570 := [monotonicity #302 #573]: #575
-#572 := [trans #570 #576]: #577
-#563 := [monotonicity #572]: #421
-#566 := [trans #563 #565]: #421
-#420 := [quant-inst]: #419
-#560 := [mp #420 #566]: #578
-#442 := [unit-resolution #560 #632]: #301
-#443 := (= #293 #589)
+#567 := (not true)
+#569 := (iff #567 false)
+#398 := [rewrite]: #569
+#568 := (iff #419 #567)
+#560 := (iff #578 true)
+#561 := [rewrite]: #560
+#562 := [monotonicity #561]: #568
+#401 := [trans #562 #398]: #400
+#564 := (iff #421 #565)
+#566 := [rewrite]: #564
+#557 := [monotonicity #566 #401]: #555
+#441 := [trans #557 #556]: #559
+#452 := [monotonicity #441]: #550
+#551 := [trans #452 #548]: #550
+#402 := [quant-inst]: #443
+#436 := [mp #402 #551]: #442
+#524 := [unit-resolution #436 #632]: #565
+#526 := (= #293 #420)
#28 := (= #25 #27)
#129 := [asserted]: #28
-#436 := [monotonicity #129]: #443
-#451 := [trans #436 #442]: #555
-#453 := (not #555)
-#454 := (or #453 #558)
-#447 := [th-lemma]: #454
-#455 := [unit-resolution #447 #451]: #558
-#456 := (not #558)
-#458 := (or #456 #457)
-#459 := [th-lemma]: #458
-#552 := [unit-resolution #459 #455]: #457
-#553 := (or #549 #588)
-#540 := [th-lemma]: #553
-#542 := [unit-resolution #540 #552]: #549
-#603 := (or #237 #606)
+#527 := [monotonicity #129]: #526
+#528 := [trans #527 #524]: #552
+#529 := (not #552)
+#525 := (or #529 #553)
+#530 := [th-lemma]: #525
+#516 := [unit-resolution #530 #528]: #553
+#517 := (not #553)
+#520 := (or #517 #519)
+#521 := [th-lemma]: #520
+#518 := [unit-resolution #521 #516]: #519
+#502 := (or #522 #577)
+#503 := [th-lemma]: #502
+#505 := [unit-resolution #503 #518]: #522
+#300 := (or #255 #317)
#18 := (= #13 0::int)
#118 := (or #18 #70)
#633 := (forall (vars (?x3 int)) (:pat #626) #118)
@@ -8954,71 +9090,95 @@
#128 := [mp #88 #127]: #123
#149 := [mp~ #128 #134]: #123
#638 := [mp #149 #637]: #633
-#604 := (not #633)
-#602 := (or #604 #237 #606)
+#582 := (not #633)
+#296 := (or #582 #255 #317)
#204 := (>= #24 0::int)
-#601 := (or #237 #204)
-#605 := (or #604 #601)
-#317 := (iff #605 #602)
-#592 := (or #604 #603)
-#315 := (iff #592 #602)
-#316 := [rewrite]: #315
-#299 := (iff #605 #592)
-#242 := (iff #601 #603)
-#279 := (iff #204 #606)
-#280 := [rewrite]: #279
-#243 := [monotonicity #280]: #242
-#314 := [monotonicity #243]: #299
-#210 := [trans #314 #316]: #317
-#591 := [quant-inst]: #605
-#587 := [mp #591 #210]: #602
-#534 := [unit-resolution #587 #638]: #603
-#531 := [unit-resolution #534 #542]: #606
-#613 := (not #606)
-#607 := (or #613 #612)
-#251 := (or #609 #613 #612)
+#210 := (or #317 #204)
+#579 := (or #582 #210)
+#570 := (iff #579 #296)
+#580 := (or #582 #300)
+#574 := (iff #580 #296)
+#575 := [rewrite]: #574
+#584 := (iff #579 #580)
+#303 := (iff #210 #300)
+#606 := (* 1::int #23)
+#279 := (>= #606 0::int)
+#311 := (or #279 #317)
+#301 := (iff #311 #300)
+#256 := (iff #279 #255)
+#251 := (= #606 #23)
+#593 := [rewrite]: #251
+#257 := [monotonicity #593]: #256
+#302 := [monotonicity #257]: #301
+#586 := (iff #210 #311)
+#587 := (or #317 #279)
+#585 := (iff #587 #311)
+#589 := [rewrite]: #585
+#588 := (iff #210 #587)
+#280 := (iff #204 #279)
+#613 := [rewrite]: #280
+#310 := [monotonicity #613]: #588
+#590 := [trans #310 #589]: #586
+#581 := [trans #590 #302]: #303
+#573 := [monotonicity #581]: #584
+#571 := [trans #573 #575]: #570
+#583 := [quant-inst]: #579
+#576 := [mp #583 #571]: #296
+#506 := [unit-resolution #576 #638]: #300
+#507 := [unit-resolution #506 #505]: #255
+#258 := (not #255)
+#597 := (or #258 #612)
+#601 := (or #237 #258 #612)
#289 := (not #204)
#294 := (= #24 #293)
#291 := (or #294 #289)
-#593 := (or #609 #291)
-#597 := (iff #593 #251)
-#256 := (or #609 #607)
-#595 := (iff #256 #251)
-#596 := [rewrite]: #595
-#257 := (iff #593 #256)
-#608 := (iff #291 #607)
-#616 := (or #612 #613)
-#266 := (iff #616 #607)
-#271 := [rewrite]: #266
-#611 := (iff #291 #616)
-#614 := (iff #289 #613)
-#615 := [monotonicity #280]: #614
+#603 := (or #237 #291)
+#592 := (iff #603 #601)
+#243 := (or #237 #597)
+#605 := (iff #243 #601)
+#591 := [rewrite]: #605
+#604 := (iff #603 #243)
+#594 := (iff #291 #597)
+#614 := (not #279)
+#266 := (or #614 #612)
+#598 := (iff #266 #597)
+#595 := (iff #614 #258)
+#596 := [monotonicity #257]: #595
+#599 := [monotonicity #596]: #598
+#267 := (iff #291 #266)
+#611 := (or #612 #614)
+#271 := (iff #611 #266)
+#608 := [rewrite]: #271
+#617 := (iff #291 #611)
+#615 := (iff #289 #614)
+#616 := [monotonicity #613]: #615
#268 := (iff #294 #612)
#399 := [rewrite]: #268
-#617 := [monotonicity #399 #615]: #611
-#267 := [trans #617 #271]: #608
-#258 := [monotonicity #267]: #257
-#598 := [trans #258 #596]: #597
-#255 := [quant-inst]: #593
-#599 := [mp #255 #598]: #251
-#533 := [unit-resolution #599 #632]: #607
-#543 := [unit-resolution #533 #531]: #612
-#544 := (not #612)
-#545 := (or #544 #594)
-#541 := [th-lemma]: #545
-#546 := [unit-resolution #541 #543]: #594
-#600 := (>= #610 0::int)
-#535 := (or #544 #600)
-#536 := [th-lemma]: #535
-#537 := [unit-resolution #536 #543]: #600
-#557 := (<= #293 1::int)
-#538 := (or #453 #557)
-#532 := [th-lemma]: #538
-#539 := [unit-resolution #532 #451]: #557
-[th-lemma #455 #539 #537 #546]: false
-unsat
-
-kDuOn7kAggfP4Um8ghhv5A 5551
+#607 := [monotonicity #399 #616]: #617
+#609 := [trans #607 #608]: #267
+#600 := [trans #609 #599]: #594
+#602 := [monotonicity #600]: #604
+#299 := [trans #602 #591]: #592
+#242 := [quant-inst]: #603
+#314 := [mp #242 #299]: #601
+#508 := [unit-resolution #314 #632]: #597
+#509 := [unit-resolution #508 #507]: #612
+#510 := (not #612)
+#511 := (or #510 #315)
+#512 := [th-lemma]: #511
+#513 := [unit-resolution #512 #509]: #315
+#316 := (>= #610 0::int)
+#514 := (or #510 #316)
+#504 := [th-lemma]: #514
+#515 := [unit-resolution #504 #509]: #316
+#549 := (<= #293 1::int)
+#493 := (or #529 #549)
+#494 := [th-lemma]: #493
+#496 := [unit-resolution #494 #528]: #549
+[th-lemma #516 #496 #515 #513]: false
+unsat
+
+kDuOn7kAggfP4Um8ghhv5A 6068
#2 := false
#23 := 3::int
decl uf_2 :: (-> T1 int)
@@ -9041,13 +9201,13 @@
#632 := -1::int
#634 := (* -1::int #28)
#290 := (+ #26 #634)
-#623 := (>= #290 0::int)
+#609 := (>= #290 0::int)
#421 := (= #290 0::int)
-#302 := (>= #22 0::int)
-#625 := (= #28 0::int)
-#318 := (not #625)
-#322 := (<= #28 0::int)
-#324 := (not #322)
+#273 := (>= #22 0::int)
+#610 := (= #28 0::int)
+#588 := (not #610)
+#441 := (<= #28 0::int)
+#443 := (not #441)
#29 := 7::int
#143 := (>= #28 7::int)
#30 := (< #28 7::int)
@@ -9064,12 +9224,12 @@
#151 := [trans #147 #149]: #150
#133 := [asserted]: #31
#152 := [mp #133 #151]: #143
-#325 := (or #324 #141)
-#603 := [th-lemma]: #325
-#604 := [unit-resolution #603 #152]: #324
-#601 := (or #318 #322)
-#605 := [th-lemma]: #601
-#602 := [unit-resolution #605 #604]: #318
+#585 := (or #443 #141)
+#586 := [th-lemma]: #585
+#587 := [unit-resolution #586 #152]: #443
+#582 := (or #588 #441)
+#583 := [th-lemma]: #582
+#589 := [unit-resolution #583 #587]: #588
#10 := (:var 0 int)
#12 := (uf_1 #10)
#648 := (pattern #12)
@@ -9132,33 +9292,45 @@
#131 := [mp #91 #130]: #126
#172 := [mp~ #131 #155]: #126
#660 := [mp #172 #659]: #655
-#337 := (not #655)
-#338 := (or #337 #302 #625)
+#605 := (not #655)
+#602 := (or #605 #273 #610)
#315 := (>= #26 0::int)
-#264 := (or #625 #315)
-#339 := (or #337 #264)
-#611 := (iff #339 #338)
-#627 := (or #302 #625)
-#609 := (or #337 #627)
-#333 := (iff #609 #338)
-#607 := [rewrite]: #333
-#610 := (iff #339 #609)
-#321 := (iff #264 #627)
-#265 := (or #625 #302)
-#613 := (iff #265 #627)
-#614 := [rewrite]: #613
-#626 := (iff #264 #265)
-#635 := (iff #315 #302)
-#636 := [rewrite]: #635
-#624 := [monotonicity #636]: #626
-#336 := [trans #624 #614]: #321
-#332 := [monotonicity #336]: #610
-#608 := [trans #332 #607]: #611
-#231 := [quant-inst]: #339
-#612 := [mp #231 #608]: #338
-#606 := [unit-resolution #612 #660 #602]: #302
-#637 := (not #302)
-#293 := (or #637 #421)
+#332 := (or #610 #315)
+#606 := (or #605 #332)
+#599 := (iff #606 #602)
+#323 := (or #273 #610)
+#596 := (or #605 #323)
+#593 := (iff #596 #602)
+#598 := [rewrite]: #593
+#597 := (iff #606 #596)
+#318 := (iff #332 #323)
+#302 := 1::int
+#635 := (* 1::int #22)
+#636 := (>= #635 0::int)
+#333 := (or #610 #636)
+#603 := (iff #333 #323)
+#608 := (or #610 #273)
+#324 := (iff #608 #323)
+#325 := [rewrite]: #324
+#612 := (iff #333 #608)
+#615 := (iff #636 #273)
+#289 := (= #635 #22)
+#631 := [rewrite]: #289
+#277 := [monotonicity #631]: #615
+#322 := [monotonicity #277]: #612
+#604 := [trans #322 #325]: #603
+#607 := (iff #332 #333)
+#637 := (iff #315 #636)
+#638 := [rewrite]: #637
+#611 := [monotonicity #638]: #607
+#601 := [trans #611 #604]: #318
+#592 := [monotonicity #601]: #597
+#594 := [trans #592 #598]: #599
+#595 := [quant-inst]: #606
+#600 := [mp #595 #594]: #602
+#590 := [unit-resolution #600 #660 #589]: #273
+#278 := (not #273)
+#620 := (or #278 #421)
#55 := (= #10 #13)
#80 := (or #55 #74)
#649 := (forall (vars (?x2 int)) (:pat #648) #80)
@@ -9208,39 +9380,47 @@
#90 := [mp #54 #89]: #85
#169 := [mp~ #90 #134]: #85
#654 := [mp #169 #653]: #649
-#615 := (not #649)
-#277 := (or #615 #637 #421)
+#264 := (not #649)
+#265 := (or #264 #278 #421)
#243 := (not #315)
#317 := (= #26 #28)
#296 := (or #317 #243)
-#278 := (or #615 #296)
-#621 := (iff #278 #277)
-#280 := (or #615 #293)
-#619 := (iff #280 #277)
-#620 := [rewrite]: #619
-#617 := (iff #278 #280)
-#631 := (iff #296 #293)
-#639 := (or #421 #637)
-#630 := (iff #639 #293)
-#289 := [rewrite]: #630
-#629 := (iff #296 #639)
-#638 := (iff #243 #637)
-#633 := [monotonicity #636]: #638
+#626 := (or #264 #296)
+#337 := (iff #626 #265)
+#627 := (or #264 #620)
+#321 := (iff #627 #265)
+#336 := [rewrite]: #321
+#613 := (iff #626 #627)
+#623 := (iff #296 #620)
+#633 := (not #636)
+#288 := (or #421 #633)
+#622 := (iff #288 #620)
+#617 := (or #421 #278)
+#621 := (iff #617 #620)
+#616 := [rewrite]: #621
+#618 := (iff #288 #617)
+#279 := (iff #633 #278)
+#280 := [monotonicity #277]: #279
+#619 := [monotonicity #280]: #618
+#259 := [trans #619 #616]: #622
+#293 := (iff #296 #288)
+#639 := (iff #243 #633)
+#629 := [monotonicity #638]: #639
#628 := (iff #317 #421)
#301 := [rewrite]: #628
-#288 := [monotonicity #301 #633]: #629
-#273 := [trans #288 #289]: #631
-#618 := [monotonicity #273]: #617
-#616 := [trans #618 #620]: #621
-#279 := [quant-inst]: #278
-#622 := [mp #279 #616]: #277
-#595 := [unit-resolution #622 #654]: #293
-#596 := [unit-resolution #595 #606]: #421
-#597 := (not #421)
-#592 := (or #597 #623)
-#593 := [th-lemma]: #592
-#598 := [unit-resolution #593 #596]: #623
-[th-lemma #152 #598 #139]: false
+#630 := [monotonicity #301 #629]: #293
+#625 := [trans #630 #259]: #623
+#614 := [monotonicity #625]: #613
+#338 := [trans #614 #336]: #337
+#624 := [quant-inst]: #626
+#339 := [mp #624 #338]: #265
+#584 := [unit-resolution #339 #654]: #620
+#591 := [unit-resolution #584 #590]: #421
+#420 := (not #421)
+#422 := (or #420 #609)
+#423 := [th-lemma]: #422
+#576 := [unit-resolution #423 #591]: #609
+[th-lemma #152 #576 #139]: false
unsat
aiB004AWADNjynNrqCDsxw 9284
@@ -9916,7 +10096,7 @@
[th-lemma #623 #188 #601 #628]: false
unsat
-ZcLxnpFewGGQ0H47MfRXGw 11816
+ZcLxnpFewGGQ0H47MfRXGw 12389
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -9930,8 +10110,8 @@
#297 := (uf_2 #141)
#357 := (= #297 0::int)
#166 := (uf_1 0::int)
-#531 := (uf_2 #166)
-#537 := (= #531 0::int)
+#454 := (uf_2 #166)
+#515 := (= #454 0::int)
#10 := (:var 0 int)
#12 := (uf_1 #10)
#672 := (pattern #12)
@@ -9988,40 +10168,40 @@
#192 := [mp~ #95 #175]: #90
#678 := [mp #192 #677]: #673
#650 := (not #673)
-#528 := (or #650 #537)
-#529 := (>= 0::int 0::int)
-#530 := (not #529)
-#534 := (= 0::int #531)
-#535 := (or #534 #530)
-#508 := (or #650 #535)
-#509 := (iff #508 #528)
-#514 := (iff #528 #528)
-#515 := [rewrite]: #514
-#527 := (iff #535 #537)
-#520 := (or #537 false)
-#525 := (iff #520 #537)
-#526 := [rewrite]: #525
-#521 := (iff #535 #520)
-#519 := (iff #530 false)
+#468 := (or #650 #515)
+#528 := (>= 0::int 0::int)
+#508 := (not #528)
+#509 := (= 0::int #454)
+#490 := (or #509 #508)
+#469 := (or #650 #490)
+#471 := (iff #469 #468)
+#473 := (iff #468 #468)
+#474 := [rewrite]: #473
+#462 := (iff #490 #515)
+#495 := (or #515 false)
+#486 := (iff #495 #515)
+#507 := [rewrite]: #486
+#496 := (iff #490 #495)
+#492 := (iff #508 false)
#1 := true
-#512 := (not true)
-#517 := (iff #512 false)
-#518 := [rewrite]: #517
-#513 := (iff #530 #512)
-#538 := (iff #529 true)
-#511 := [rewrite]: #538
-#406 := [monotonicity #511]: #513
-#524 := [trans #406 #518]: #519
-#536 := (iff #534 #537)
-#532 := [rewrite]: #536
-#522 := [monotonicity #532 #524]: #521
-#523 := [trans #522 #526]: #527
-#490 := [monotonicity #523]: #509
-#510 := [trans #490 #515]: #509
-#454 := [quant-inst]: #508
-#516 := [mp #454 #510]: #528
-#394 := [unit-resolution #516 #678]: #537
-#355 := (= #297 #531)
+#491 := (not true)
+#483 := (iff #491 false)
+#485 := [rewrite]: #483
+#450 := (iff #508 #491)
+#516 := (iff #528 true)
+#484 := [rewrite]: #516
+#481 := [monotonicity #484]: #450
+#494 := [trans #481 #485]: #492
+#514 := (iff #509 #515)
+#510 := [rewrite]: #514
+#506 := [monotonicity #510 #494]: #496
+#463 := [trans #506 #507]: #462
+#472 := [monotonicity #463]: #471
+#475 := [trans #472 #474]: #471
+#470 := [quant-inst]: #469
+#476 := [mp #470 #475]: #468
+#351 := [unit-resolution #476 #678]: #515
+#287 := (= #297 #454)
#250 := (= #141 #166)
#26 := 2::int
#144 := (* 2::int #22)
@@ -10032,29 +10212,24 @@
#161 := (uf_1 #156)
#336 := (= #161 #166)
#327 := (not #336)
-#588 := (uf_2 #161)
-#555 := (= #588 0::int)
-#398 := (= #588 #531)
-#395 := [hypothesis]: #336
-#387 := [monotonicity #395]: #398
-#399 := [trans #387 #394]: #555
-#390 := (not #555)
-#547 := (<= #588 0::int)
-#403 := (not #547)
-#595 := (>= #150 0::int)
-#302 := -1::int
-#618 := (* -1::int #150)
-#624 := (+ #144 #618)
-#488 := (<= #624 0::int)
-#465 := (= #624 0::int)
-#609 := (>= #22 0::int)
-#442 := (= #22 0::int)
+#568 := (uf_2 #161)
+#537 := (= #568 0::int)
+#354 := (= #568 #454)
+#352 := [hypothesis]: #336
+#344 := [monotonicity #352]: #354
+#355 := [trans #344 #351]: #537
+#368 := (not #537)
+#527 := (<= #568 0::int)
+#375 := (not #527)
+#566 := (>= #150 0::int)
+#447 := (>= #22 0::int)
+#421 := (= #22 0::int)
#660 := (uf_1 #22)
-#495 := (uf_2 #660)
-#496 := (= #495 0::int)
-#612 := (not #609)
-#451 := [hypothesis]: #612
-#506 := (or #496 #609)
+#451 := (uf_2 #660)
+#452 := (= #451 0::int)
+#603 := (not #447)
+#424 := [hypothesis]: #603
+#455 := (or #447 #452)
#18 := (= #13 0::int)
#126 := (or #18 #78)
#679 := (forall (vars (?x3 int)) (:pat #672) #126)
@@ -10112,15 +10287,23 @@
#195 := [mp~ #136 #180]: #131
#684 := [mp #195 #683]: #679
#346 := (not #679)
-#462 := (or #346 #496 #609)
-#463 := (or #346 #506)
-#469 := (iff #463 #462)
-#470 := [rewrite]: #469
-#468 := [quant-inst]: #463
-#471 := [mp #468 #470]: #462
-#452 := [unit-resolution #471 #684]: #506
-#453 := [unit-resolution #452 #451]: #496
-#456 := (= #22 #495)
+#458 := (or #346 #447 #452)
+#453 := (or #452 #447)
+#459 := (or #346 #453)
+#434 := (iff #459 #458)
+#443 := (or #346 #455)
+#432 := (iff #443 #458)
+#433 := [rewrite]: #432
+#461 := (iff #459 #443)
+#456 := (iff #453 #455)
+#457 := [rewrite]: #456
+#431 := [monotonicity #457]: #461
+#436 := [trans #431 #433]: #434
+#460 := [quant-inst]: #459
+#437 := [mp #460 #436]: #458
+#420 := [unit-resolution #437 #684]: #455
+#425 := [unit-resolution #420 #424]: #452
+#405 := (= #22 #451)
#661 := (= uf_3 #660)
#4 := (:var 0 T1)
#5 := (uf_2 #4)
@@ -10151,117 +10334,136 @@
#663 := (not #665)
#653 := (or #663 #661)
#312 := [quant-inst]: #653
-#455 := [unit-resolution #312 #671]: #661
-#457 := [monotonicity #455]: #456
-#458 := [trans #457 #453]: #442
-#459 := (not #442)
-#460 := (or #459 #609)
-#443 := [th-lemma]: #460
-#461 := [unit-resolution #443 #451 #458]: false
-#431 := [lemma #461]: #609
-#613 := (or #465 #612)
-#615 := (or #650 #465 #612)
+#415 := [unit-resolution #312 #671]: #661
+#407 := [monotonicity #415]: #405
+#408 := [trans #407 #425]: #421
+#411 := (not #421)
+#412 := (or #411 #447)
+#416 := [th-lemma]: #412
+#409 := [unit-resolution #416 #424 #408]: false
+#417 := [lemma #409]: #447
+#302 := -1::int
+#618 := (* -1::int #150)
+#624 := (+ #144 #618)
+#595 := (<= #624 0::int)
+#465 := (= #624 0::int)
+#489 := (or #603 #465)
+#482 := (or #650 #603 #465)
#616 := (>= #144 0::int)
#617 := (not #616)
#622 := (= #144 #150)
#623 := (or #622 #617)
-#444 := (or #650 #623)
-#602 := (iff #444 #615)
-#447 := (or #650 #613)
-#603 := (iff #447 #615)
-#604 := [rewrite]: #603
-#600 := (iff #444 #447)
-#614 := (iff #623 #613)
-#606 := (iff #617 #612)
-#610 := (iff #616 #609)
-#611 := [rewrite]: #610
-#607 := [monotonicity #611]: #606
+#497 := (or #650 #623)
+#504 := (iff #497 #482)
+#500 := (or #650 #489)
+#502 := (iff #500 #482)
+#503 := [rewrite]: #502
+#493 := (iff #497 #500)
+#594 := (iff #623 #489)
+#609 := (* 1::int #22)
+#610 := (>= #609 0::int)
+#606 := (not #610)
+#614 := (or #465 #606)
+#498 := (iff #614 #489)
+#605 := (or #465 #603)
+#448 := (iff #605 #489)
+#596 := [rewrite]: #448
+#487 := (iff #614 #605)
+#604 := (iff #606 #603)
+#600 := (iff #610 #447)
+#444 := (= #609 #22)
+#446 := [rewrite]: #444
+#601 := [monotonicity #446]: #600
+#602 := [monotonicity #601]: #604
+#488 := [monotonicity #602]: #487
+#593 := [trans #488 #596]: #498
+#608 := (iff #623 #614)
+#607 := (iff #617 #606)
+#611 := (iff #616 #610)
+#612 := [rewrite]: #611
+#613 := [monotonicity #612]: #607
#466 := (iff #622 #465)
#467 := [rewrite]: #466
-#608 := [monotonicity #467 #607]: #614
-#601 := [monotonicity #608]: #600
-#605 := [trans #601 #604]: #602
-#446 := [quant-inst]: #444
-#487 := [mp #446 #605]: #615
-#439 := [unit-resolution #487 #678]: #613
-#435 := [unit-resolution #439 #431]: #465
-#440 := (not #465)
-#419 := (or #440 #488)
-#422 := [th-lemma]: #419
-#426 := [unit-resolution #422 #435]: #488
-#430 := (not #488)
-#433 := (or #595 #612 #430)
-#438 := [th-lemma]: #433
-#402 := [unit-resolution #438 #431 #426]: #595
-#590 := -3::int
-#579 := (* -1::int #588)
-#589 := (+ #150 #579)
-#553 := (<= #589 -3::int)
-#591 := (= #589 -3::int)
-#581 := (>= #150 -3::int)
+#615 := [monotonicity #467 #613]: #608
+#597 := [trans #615 #593]: #594
+#501 := [monotonicity #597]: #493
+#505 := [trans #501 #503]: #504
+#499 := [quant-inst]: #497
+#598 := [mp #499 #505]: #482
+#404 := [unit-resolution #598 #678]: #489
+#386 := [unit-resolution #404 #417]: #465
+#388 := (not #465)
+#389 := (or #388 #595)
+#390 := [th-lemma]: #389
+#391 := [unit-resolution #390 #386]: #595
+#395 := (not #595)
+#413 := (or #566 #603 #395)
+#403 := [th-lemma]: #413
+#373 := [unit-resolution #403 #391 #417]: #566
+#553 := -3::int
+#551 := (* -1::int #568)
+#552 := (+ #150 #551)
+#535 := (<= #552 -3::int)
+#554 := (= #552 -3::int)
+#557 := (>= #150 -3::int)
#644 := (>= #22 -1::int)
-#428 := (or #612 #644)
-#429 := [th-lemma]: #428
-#427 := [unit-resolution #429 #431]: #644
+#392 := (or #603 #644)
+#393 := [th-lemma]: #392
+#394 := [unit-resolution #393 #417]: #644
#646 := (not #644)
-#418 := (or #581 #646 #430)
-#421 := [th-lemma]: #418
-#423 := [unit-resolution #421 #426 #427]: #581
-#584 := (not #581)
-#573 := (or #584 #591)
-#562 := (or #650 #584 #591)
-#599 := (>= #156 0::int)
-#586 := (not #599)
-#580 := (= #156 #588)
-#577 := (or #580 #586)
-#563 := (or #650 #577)
-#549 := (iff #563 #562)
-#566 := (or #650 #573)
-#568 := (iff #566 #562)
-#548 := [rewrite]: #568
-#567 := (iff #563 #566)
-#571 := (iff #577 #573)
-#569 := (or #591 #584)
-#574 := (iff #569 #573)
-#575 := [rewrite]: #574
-#570 := (iff #577 #569)
-#578 := (iff #586 #584)
-#582 := (iff #599 #581)
-#583 := [rewrite]: #582
-#585 := [monotonicity #583]: #578
-#587 := (iff #580 #591)
-#592 := [rewrite]: #587
-#572 := [monotonicity #592 #585]: #570
-#576 := [trans #572 #575]: #571
-#564 := [monotonicity #576]: #567
-#551 := [trans #564 #548]: #549
-#565 := [quant-inst]: #563
-#552 := [mp #565 #551]: #562
-#424 := [unit-resolution #552 #678]: #573
-#420 := [unit-resolution #424 #423]: #591
-#425 := (not #591)
-#415 := (or #425 #553)
-#405 := [th-lemma]: #415
-#407 := [unit-resolution #405 #420]: #553
-#404 := (not #553)
-#401 := (not #595)
-#386 := (or #403 #401 #404)
-#388 := [th-lemma]: #386
-#389 := [unit-resolution #388 #407 #402]: #403
-#391 := (or #390 #547)
-#392 := [th-lemma]: #391
-#393 := [unit-resolution #392 #389]: #390
-#376 := [unit-resolution #393 #399]: false
-#378 := [lemma #376]: #327
+#396 := (or #557 #646 #395)
+#397 := [th-lemma]: #396
+#398 := [unit-resolution #397 #391 #394]: #557
+#560 := (not #557)
+#539 := (or #554 #560)
+#543 := (or #650 #554 #560)
+#567 := (>= #156 0::int)
+#564 := (not #567)
+#548 := (= #156 #568)
+#549 := (or #548 #564)
+#544 := (or #650 #549)
+#530 := (iff #544 #543)
+#546 := (or #650 #539)
+#533 := (iff #546 #543)
+#529 := [rewrite]: #533
+#541 := (iff #544 #546)
+#540 := (iff #549 #539)
+#550 := (iff #564 #560)
+#558 := (iff #567 #557)
+#559 := [rewrite]: #558
+#561 := [monotonicity #559]: #550
+#555 := (iff #548 #554)
+#556 := [rewrite]: #555
+#542 := [monotonicity #556 #561]: #540
+#547 := [monotonicity #542]: #541
+#531 := [trans #547 #529]: #530
+#545 := [quant-inst]: #544
+#534 := [mp #545 #531]: #543
+#387 := [unit-resolution #534 #678]: #539
+#399 := [unit-resolution #387 #398]: #554
+#376 := (not #554)
+#378 := (or #376 #535)
+#379 := [th-lemma]: #378
+#380 := [unit-resolution #379 #399]: #535
+#365 := (not #535)
+#364 := (not #566)
+#366 := (or #375 #364 #365)
+#358 := [th-lemma]: #366
+#367 := [unit-resolution #358 #380 #373]: #375
+#359 := (or #368 #527)
+#369 := [th-lemma]: #359
+#350 := [unit-resolution #369 #367]: #368
+#321 := [unit-resolution #350 #355]: false
+#323 := [lemma #321]: #327
#249 := (= #141 #161)
#334 := (not #249)
-#396 := (= #297 #588)
-#385 := [hypothesis]: #249
-#370 := [monotonicity #385]: #396
-#380 := (not #396)
-#434 := (+ #297 #579)
-#280 := (>= #434 0::int)
-#414 := (not #280)
+#343 := (= #297 #568)
+#322 := [hypothesis]: #249
+#333 := [monotonicity #322]: #343
+#315 := (not #343)
+#414 := (+ #297 #551)
+#401 := (>= #414 0::int)
+#372 := (not #401)
#303 := (* -1::int #297)
#304 := (+ #22 #303)
#356 := (>= #304 -1::int)
@@ -10290,21 +10492,21 @@
#256 := [trans #360 #362]: #363
#637 := [quant-inst]: #651
#633 := [mp #637 #256]: #648
-#408 := [unit-resolution #633 #678]: #649
-#411 := [unit-resolution #408 #427]: #641
-#412 := (not #641)
-#416 := (or #412 #356)
-#409 := [th-lemma]: #416
-#417 := [unit-resolution #409 #411]: #356
-#410 := [hypothesis]: #280
-#413 := [th-lemma #423 #410 #417 #407 #426]: false
-#400 := [lemma #413]: #414
-#381 := (or #380 #280)
-#382 := [th-lemma]: #381
-#377 := [unit-resolution #382 #400]: #380
-#371 := [unit-resolution #377 #370]: false
-#372 := [lemma #371]: #334
-#352 := (or #249 #250 #336)
+#381 := [unit-resolution #633 #678]: #649
+#382 := [unit-resolution #381 #394]: #641
+#383 := (not #641)
+#384 := (or #383 #356)
+#377 := [th-lemma]: #384
+#385 := [unit-resolution #377 #382]: #356
+#370 := [hypothesis]: #401
+#371 := [th-lemma #398 #370 #385 #380 #391]: false
+#374 := [lemma #371]: #372
+#328 := (or #315 #401)
+#329 := [th-lemma]: #328
+#332 := [unit-resolution #329 #374]: #315
+#316 := [unit-resolution #332 #333]: false
+#318 := [lemma #316]: #334
+#295 := (or #249 #250 #336)
#335 := (not #250)
#338 := (and #334 #335 #327)
#339 := (not #338)
@@ -10352,28 +10554,28 @@
#177 := [mp #137 #174]: #172
#326 := (or #169 #339)
#659 := [def-axiom]: #326
-#351 := [unit-resolution #659 #177]: #339
+#294 := [unit-resolution #659 #177]: #339
#314 := (or #338 #249 #250 #336)
#445 := [def-axiom]: #314
-#343 := [unit-resolution #445 #351]: #352
-#353 := [unit-resolution #343 #372 #378]: #250
-#321 := [monotonicity #353]: #355
-#323 := [trans #321 #394]: #357
-#368 := (not #357)
+#293 := [unit-resolution #445 #294]: #295
+#296 := [unit-resolution #293 #318 #323]: #250
+#290 := [monotonicity #296]: #287
+#285 := [trans #290 #351]: #357
+#310 := (not #357)
#620 := (<= #297 0::int)
-#364 := (not #620)
+#305 := (not #620)
#634 := (<= #304 -1::int)
-#374 := (or #412 #634)
-#373 := [th-lemma]: #374
-#375 := [unit-resolution #373 #411]: #634
-#365 := (not #634)
-#366 := (or #364 #612 #365)
-#358 := [th-lemma]: #366
-#367 := [unit-resolution #358 #375 #431]: #364
-#359 := (or #368 #620)
-#369 := [th-lemma]: #359
-#350 := [unit-resolution #369 #367]: #368
-[unit-resolution #350 #323]: false
+#319 := (or #383 #634)
+#298 := [th-lemma]: #319
+#300 := [unit-resolution #298 #382]: #634
+#306 := (not #634)
+#307 := (or #305 #603 #306)
+#308 := [th-lemma]: #307
+#309 := [unit-resolution #308 #300 #417]: #305
+#299 := (or #310 #620)
+#311 := [th-lemma]: #299
+#292 := [unit-resolution #311 #309]: #310
+[unit-resolution #292 #285]: false
unsat
ipe8HUL/t33OoeNl/z0smQ 4011
@@ -10539,7 +10741,7 @@
[th-lemma #656 #361 #261]: false
unsat
-eRjXXrQSzpzyc8Ro409d3Q 14366
+9FrZeHP8ZKJM+JmhfAjimQ 14889
#2 := false
#9 := 0::int
decl uf_2 :: (-> T1 int)
@@ -10551,46 +10753,46 @@
#38 := (* 4::int #37)
#39 := (uf_1 #38)
#40 := (uf_2 #39)
-#527 := (= #40 0::int)
-#976 := (not #527)
-#502 := (<= #40 0::int)
-#971 := (not #502)
+#504 := (= #40 0::int)
+#995 := (not #504)
+#475 := (<= #40 0::int)
+#990 := (not #475)
#22 := 1::int
#186 := (+ 1::int #40)
#189 := (uf_1 #186)
-#506 := (uf_2 #189)
-#407 := (<= #506 1::int)
-#876 := (not #407)
+#467 := (uf_2 #189)
+#380 := (<= #467 1::int)
+#893 := (not #380)
decl up_4 :: (-> T1 T1 bool)
#4 := (:var 0 T1)
-#408 := (up_4 #4 #189)
-#393 := (pattern #408)
-#413 := (= #4 #189)
-#414 := (not #408)
+#386 := (up_4 #4 #189)
+#374 := (pattern #386)
+#382 := (not #386)
+#385 := (= #4 #189)
#26 := (uf_1 1::int)
#27 := (= #4 #26)
-#392 := (or #27 #414 #413)
-#397 := (forall (vars (?x5 T1)) (:pat #393) #392)
-#383 := (not #397)
-#382 := (or #383 #407)
-#375 := (not #382)
+#845 := (or #27 #385 #382)
+#848 := (forall (vars (?x5 T1)) (:pat #374) #845)
+#851 := (not #848)
+#854 := (or #380 #851)
+#857 := (not #854)
decl up_3 :: (-> T1 bool)
#192 := (up_3 #189)
-#404 := (not #192)
-#841 := (or #404 #375)
+#840 := (not #192)
+#860 := (or #840 #857)
decl ?x5!0 :: (-> T1 T1)
-#422 := (?x5!0 #189)
-#434 := (= #189 #422)
-#417 := (up_4 #422 #189)
-#418 := (not #417)
-#415 := (= #26 #422)
-#847 := (or #415 #418 #434)
-#850 := (not #847)
-#853 := (or #192 #407 #850)
-#856 := (not #853)
-#844 := (not #841)
-#859 := (or #844 #856)
-#862 := (not #859)
+#392 := (?x5!0 #189)
+#397 := (up_4 #392 #189)
+#390 := (not #397)
+#396 := (= #26 #392)
+#395 := (= #189 #392)
+#866 := (or #395 #396 #390)
+#869 := (not #866)
+#872 := (or #192 #380 #869)
+#875 := (not #872)
+#863 := (not #860)
+#878 := (or #863 #875)
+#881 := (not #878)
#5 := (uf_2 #4)
#787 := (pattern #5)
#21 := (up_3 #4)
@@ -10769,64 +10971,59 @@
#314 := [mp #267 #313]: #311
#839 := [mp #314 #838]: #836
#589 := (not #836)
-#865 := (or #589 #862)
-#416 := (or #418 #415 #434)
-#419 := (not #416)
-#409 := (or #192 #407 #419)
-#410 := (not #409)
-#389 := (or #414 #27 #413)
-#394 := (forall (vars (?x5 T1)) (:pat #393) #389)
-#399 := (not #394)
-#401 := (or #407 #399)
-#402 := (not #401)
-#400 := (or #404 #402)
-#405 := (not #400)
-#388 := (or #405 #410)
-#391 := (not #388)
-#866 := (or #589 #391)
-#868 := (iff #866 #865)
-#870 := (iff #865 #865)
-#871 := [rewrite]: #870
-#863 := (iff #391 #862)
-#860 := (iff #388 #859)
-#857 := (iff #410 #856)
-#854 := (iff #409 #853)
-#851 := (iff #419 #850)
-#848 := (iff #416 #847)
-#849 := [rewrite]: #848
-#852 := [monotonicity #849]: #851
-#855 := [monotonicity #852]: #854
-#858 := [monotonicity #855]: #857
-#845 := (iff #405 #844)
-#842 := (iff #400 #841)
-#378 := (iff #402 #375)
-#376 := (iff #401 #382)
-#384 := (or #407 #383)
-#387 := (iff #384 #382)
-#374 := [rewrite]: #387
-#385 := (iff #401 #384)
-#380 := (iff #399 #383)
-#390 := (iff #394 #397)
-#395 := (iff #389 #392)
-#396 := [rewrite]: #395
-#398 := [quant-intro #396]: #390
-#381 := [monotonicity #398]: #380
-#386 := [monotonicity #381]: #385
-#377 := [trans #386 #374]: #376
-#840 := [monotonicity #377]: #378
-#843 := [monotonicity #840]: #842
-#846 := [monotonicity #843]: #845
-#861 := [monotonicity #846 #858]: #860
-#864 := [monotonicity #861]: #863
-#869 := [monotonicity #864]: #868
-#872 := [trans #869 #871]: #868
-#867 := [quant-inst]: #866
-#873 := [mp #867 #872]: #865
-#947 := [unit-resolution #873 #839]: #862
-#905 := (or #859 #841)
-#906 := [def-axiom]: #905
-#948 := [unit-resolution #906 #947]: #841
-#951 := (or #844 #375)
+#884 := (or #589 #881)
+#398 := (or #390 #396 #395)
+#383 := (not #398)
+#381 := (or #192 #380 #383)
+#384 := (not #381)
+#387 := (or #382 #27 #385)
+#376 := (forall (vars (?x5 T1)) (:pat #374) #387)
+#377 := (not #376)
+#375 := (or #380 #377)
+#378 := (not #375)
+#841 := (or #840 #378)
+#842 := (not #841)
+#843 := (or #842 #384)
+#844 := (not #843)
+#885 := (or #589 #844)
+#887 := (iff #885 #884)
+#889 := (iff #884 #884)
+#890 := [rewrite]: #889
+#882 := (iff #844 #881)
+#879 := (iff #843 #878)
+#876 := (iff #384 #875)
+#873 := (iff #381 #872)
+#870 := (iff #383 #869)
+#867 := (iff #398 #866)
+#868 := [rewrite]: #867
+#871 := [monotonicity #868]: #870
+#874 := [monotonicity #871]: #873
+#877 := [monotonicity #874]: #876
+#864 := (iff #842 #863)
+#861 := (iff #841 #860)
+#858 := (iff #378 #857)
+#855 := (iff #375 #854)
+#852 := (iff #377 #851)
+#849 := (iff #376 #848)
+#846 := (iff #387 #845)
+#847 := [rewrite]: #846
+#850 := [quant-intro #847]: #849
+#853 := [monotonicity #850]: #852
+#856 := [monotonicity #853]: #855
+#859 := [monotonicity #856]: #858
+#862 := [monotonicity #859]: #861
+#865 := [monotonicity #862]: #864
+#880 := [monotonicity #865 #877]: #879
+#883 := [monotonicity #880]: #882
+#888 := [monotonicity #883]: #887
+#891 := [trans #888 #890]: #887
+#886 := [quant-inst]: #885
+#892 := [mp #886 #891]: #884
+#966 := [unit-resolution #892 #839]: #881
+#924 := (or #878 #860)
+#925 := [def-axiom]: #924
+#967 := [unit-resolution #925 #966]: #860
+#970 := (or #863 #857)
#41 := (+ #40 1::int)
#42 := (uf_1 #41)
#43 := (up_3 #42)
@@ -10838,30 +11035,30 @@
#194 := [monotonicity #191]: #193
#185 := [asserted]: #43
#197 := [mp #185 #194]: #192
-#885 := (or #844 #404 #375)
-#886 := [def-axiom]: #885
-#952 := [unit-resolution #886 #197]: #951
-#953 := [unit-resolution #952 #948]: #375
-#877 := (or #382 #876)
-#878 := [def-axiom]: #877
-#954 := [unit-resolution #878 #953]: #876
+#904 := (or #863 #840 #857)
+#905 := [def-axiom]: #904
+#971 := [unit-resolution #905 #197]: #970
+#972 := [unit-resolution #971 #967]: #857
+#894 := (or #854 #893)
+#895 := [def-axiom]: #894
+#973 := [unit-resolution #895 #972]: #893
#542 := -1::int
-#508 := (* -1::int #506)
-#493 := (+ #40 #508)
-#438 := (>= #493 -1::int)
-#494 := (= #493 -1::int)
-#496 := (>= #40 -1::int)
-#451 := (= #506 0::int)
-#959 := (not #451)
-#432 := (<= #506 0::int)
-#955 := (not #432)
-#956 := (or #955 #407)
-#957 := [th-lemma]: #956
-#958 := [unit-resolution #957 #954]: #955
-#960 := (or #959 #432)
-#961 := [th-lemma]: #960
-#962 := [unit-resolution #961 #958]: #959
-#453 := (or #451 #496)
+#446 := (* -1::int #467)
+#447 := (+ #40 #446)
+#416 := (>= #447 -1::int)
+#438 := (= #447 -1::int)
+#453 := (>= #40 -1::int)
+#419 := (= #467 0::int)
+#978 := (not #419)
+#388 := (<= #467 0::int)
+#974 := (not #388)
+#975 := (or #974 #380)
+#976 := [th-lemma]: #975
+#977 := [unit-resolution #976 #973]: #974
+#979 := (or #978 #388)
+#980 := [th-lemma]: #979
+#981 := [unit-resolution #980 #977]: #978
+#409 := (or #419 #453)
#10 := (:var 0 int)
#12 := (uf_1 #10)
#795 := (pattern #12)
@@ -10924,28 +11121,28 @@
#145 := [mp #105 #144]: #140
#227 := [mp~ #145 #208]: #140
#807 := [mp #227 #806]: #802
-#514 := (not #802)
-#445 := (or #514 #451 #496)
-#504 := (>= #186 0::int)
-#452 := (or #451 #504)
-#456 := (or #514 #452)
-#429 := (iff #456 #445)
-#441 := (or #514 #453)
-#423 := (iff #441 #445)
-#428 := [rewrite]: #423
-#442 := (iff #456 #441)
-#454 := (iff #452 #453)
-#498 := (iff #504 #496)
-#487 := [rewrite]: #498
-#455 := [monotonicity #487]: #454
-#421 := [monotonicity #455]: #442
-#430 := [trans #421 #428]: #429
-#439 := [quant-inst]: #456
-#431 := [mp #439 #430]: #445
-#963 := [unit-resolution #431 #807]: #453
-#964 := [unit-resolution #963 #962]: #496
-#488 := (not #496)
-#490 := (or #494 #488)
+#496 := (not #802)
+#408 := (or #496 #419 #453)
+#476 := (>= #186 0::int)
+#407 := (or #419 #476)
+#414 := (or #496 #407)
+#404 := (iff #414 #408)
+#393 := (or #496 #409)
+#401 := (iff #393 #408)
+#402 := [rewrite]: #401
+#394 := (iff #414 #393)
+#410 := (iff #407 #409)
+#454 := (iff #476 #453)
+#455 := [rewrite]: #454
+#413 := [monotonicity #455]: #410
+#399 := [monotonicity #413]: #394
+#400 := [trans #399 #402]: #404
+#389 := [quant-inst]: #414
+#405 := [mp #389 #400]: #408
+#982 := [unit-resolution #405 #807]: #409
+#983 := [unit-resolution #982 #981]: #453
+#445 := (not #453)
+#441 := (or #438 #445)
#69 := (= #10 #13)
#94 := (or #69 #88)
#796 := (forall (vars (?x2 int)) (:pat #795) #94)
@@ -10995,48 +11192,48 @@
#104 := [mp #68 #103]: #99
#224 := [mp~ #104 #196]: #99
#801 := [mp #224 #800]: #796
-#530 := (not #796)
-#492 := (or #530 #494 #488)
-#505 := (not #504)
-#507 := (= #186 #506)
-#500 := (or #507 #505)
-#473 := (or #530 #500)
-#478 := (iff #473 #492)
-#475 := (or #530 #490)
-#477 := (iff #475 #492)
-#467 := [rewrite]: #477
-#466 := (iff #473 #475)
-#491 := (iff #500 #490)
-#489 := (iff #505 #488)
-#481 := [monotonicity #487]: #489
-#495 := (iff #507 #494)
-#497 := [rewrite]: #495
-#482 := [monotonicity #497 #481]: #491
-#476 := [monotonicity #482]: #466
-#444 := [trans #476 #467]: #478
-#474 := [quant-inst]: #473
-#446 := [mp #474 #444]: #492
-#965 := [unit-resolution #446 #801]: #490
-#966 := [unit-resolution #965 #964]: #494
-#967 := (not #494)
-#968 := (or #967 #438)
-#969 := [th-lemma]: #968
-#970 := [unit-resolution #969 #966]: #438
-#972 := (not #438)
-#973 := (or #971 #407 #972)
-#974 := [th-lemma]: #973
-#975 := [unit-resolution #974 #970 #954]: #971
-#977 := (or #976 #502)
-#978 := [th-lemma]: #977
-#979 := [unit-resolution #978 #975]: #976
-#553 := (>= #37 0::int)
-#546 := (not #553)
+#514 := (not #796)
+#423 := (or #514 #438 #445)
+#477 := (not #476)
+#478 := (= #186 #467)
+#444 := (or #478 #477)
+#428 := (or #514 #444)
+#434 := (iff #428 #423)
+#430 := (or #514 #441)
+#433 := (iff #430 #423)
+#422 := [rewrite]: #433
+#431 := (iff #428 #430)
+#442 := (iff #444 #441)
+#456 := (iff #477 #445)
+#439 := [monotonicity #455]: #456
+#451 := (iff #478 #438)
+#452 := [rewrite]: #451
+#421 := [monotonicity #452 #439]: #442
+#432 := [monotonicity #421]: #431
+#415 := [trans #432 #422]: #434
+#429 := [quant-inst]: #428
+#417 := [mp #429 #415]: #423
+#984 := [unit-resolution #417 #801]: #441
+#985 := [unit-resolution #984 #983]: #438
+#986 := (not #438)
+#987 := (or #986 #416)
+#988 := [th-lemma]: #987
+#989 := [unit-resolution #988 #985]: #416
+#991 := (not #416)
+#992 := (or #990 #380 #991)
+#993 := [th-lemma]: #992
+#994 := [unit-resolution #993 #989 #973]: #990
+#996 := (or #995 #475)
+#997 := [th-lemma]: #996
+#998 := [unit-resolution #997 #994]: #995
+#536 := (>= #37 0::int)
+#525 := (not #536)
#545 := (* -1::int #40)
#549 := (+ #38 #545)
#551 := (= #549 0::int)
-#984 := (not #551)
-#524 := (>= #549 0::int)
-#980 := (not #524)
+#1003 := (not #551)
+#503 := (>= #549 0::int)
+#999 := (not #503)
#201 := (>= #37 1::int)
#202 := (not #201)
#44 := (<= 1::int #37)
@@ -11047,55 +11244,79 @@
#204 := [monotonicity #200]: #203
#195 := [asserted]: #45
#205 := [mp #195 #204]: #202
-#981 := (or #980 #201 #407 #972)
-#982 := [th-lemma]: #981
-#983 := [unit-resolution #982 #205 #970 #954]: #980
-#985 := (or #984 #524)
-#986 := [th-lemma]: #985
-#987 := [unit-resolution #986 #983]: #984
-#548 := (or #551 #546)
-#531 := (or #530 #551 #546)
+#1000 := (or #999 #201 #380 #991)
+#1001 := [th-lemma]: #1000
+#1002 := [unit-resolution #1001 #205 #989 #973]: #999
+#1004 := (or #1003 #503)
+#1005 := [th-lemma]: #1004
+#1006 := [unit-resolution #1005 #1002]: #1003
+#527 := (or #525 #551)
+#515 := (or #514 #525 #551)
#403 := (>= #38 0::int)
#562 := (not #403)
#558 := (= #38 #40)
#563 := (or #558 #562)
-#534 := (or #530 #563)
-#537 := (iff #534 #531)
-#539 := (or #530 #548)
-#533 := (iff #539 #531)
-#536 := [rewrite]: #533
-#532 := (iff #534 #539)
-#538 := (iff #563 #548)
-#547 := (iff #562 #546)
-#541 := (iff #403 #553)
-#544 := [rewrite]: #541
-#543 := [monotonicity #544]: #547
-#552 := (iff #558 #551)
-#550 := [rewrite]: #552
-#528 := [monotonicity #550 #543]: #538
-#540 := [monotonicity #528]: #532
-#523 := [trans #540 #536]: #537
-#535 := [quant-inst]: #534
-#525 := [mp #535 #523]: #531
-#988 := [unit-resolution #525 #801]: #548
-#989 := [unit-resolution #988 #987]: #546
-#511 := (or #527 #553)
-#515 := (or #514 #527 #553)
-#509 := (or #527 #403)
-#516 := (or #514 #509)
+#516 := (or #514 #563)
#522 := (iff #516 #515)
-#518 := (or #514 #511)
+#518 := (or #514 #527)
#521 := (iff #518 #515)
#510 := [rewrite]: #521
#519 := (iff #516 #518)
-#512 := (iff #509 #511)
-#513 := [monotonicity #544]: #512
+#512 := (iff #563 #527)
+#553 := (* 1::int #37)
+#541 := (>= #553 0::int)
+#547 := (not #541)
+#531 := (or #547 #551)
+#509 := (iff #531 #527)
+#526 := (iff #547 #525)
+#537 := (iff #541 #536)
+#540 := (= #553 #37)
+#533 := [rewrite]: #540
+#523 := [monotonicity #533]: #537
+#524 := [monotonicity #523]: #526
+#511 := [monotonicity #524]: #509
+#539 := (iff #563 #531)
+#538 := (or #551 #547)
+#534 := (iff #538 #531)
+#535 := [rewrite]: #534
+#528 := (iff #563 #538)
+#543 := (iff #562 #547)
+#544 := (iff #403 #541)
+#546 := [rewrite]: #544
+#548 := [monotonicity #546]: #543
+#552 := (iff #558 #551)
+#550 := [rewrite]: #552
+#530 := [monotonicity #550 #548]: #528
+#532 := [trans #530 #535]: #539
+#513 := [trans #532 #511]: #512
#520 := [monotonicity #513]: #519
#499 := [trans #520 #510]: #522
#517 := [quant-inst]: #516
#501 := [mp #517 #499]: #515
-#990 := [unit-resolution #501 #807]: #511
-[unit-resolution #990 #989 #979]: false
+#1007 := [unit-resolution #501 #801]: #527
+#1008 := [unit-resolution #1007 #1006]: #525
+#508 := (or #504 #536)
+#498 := (or #496 #504 #536)
+#505 := (or #504 #403)
+#487 := (or #496 #505)
+#492 := (iff #487 #498)
+#489 := (or #496 #508)
+#491 := (iff #489 #498)
+#482 := [rewrite]: #491
+#481 := (iff #487 #489)
+#495 := (iff #505 #508)
+#506 := (or #504 #541)
+#493 := (iff #506 #508)
+#494 := [monotonicity #523]: #493
+#507 := (iff #505 #506)
+#500 := [monotonicity #546]: #507
+#497 := [trans #500 #494]: #495
+#490 := [monotonicity #497]: #481
+#473 := [trans #490 #482]: #492
+#488 := [quant-inst]: #487
+#474 := [mp #488 #473]: #498
+#1009 := [unit-resolution #474 #807]: #508
+[unit-resolution #1009 #1008 #998]: false
unsat
uq2MbDTgTG1nxWdwzZtWew 7
@@ -12233,7 +12454,7 @@
[mp #25 #34]: false
unsat
-YG20f6Uf93koN6rVg/alpA 9362
+YG20f6Uf93koN6rVg/alpA 9742
#2 := false
decl uf_1 :: (-> int T1)
#37 := 6::int
@@ -12248,18 +12469,18 @@
#35 := (uf_1 #34)
#36 := (uf_3 #35)
#39 := (= #36 #38)
-#476 := (uf_3 #38)
-#403 := (= #476 #38)
-#531 := (= #38 #476)
-#620 := (uf_2 #38)
+#484 := (uf_3 #38)
+#372 := (= #484 #38)
+#485 := (= #38 #484)
+#592 := (uf_2 #38)
#142 := -10::int
-#513 := (+ -10::int #620)
-#472 := (uf_1 #513)
-#503 := (uf_3 #472)
-#505 := (= #476 #503)
+#496 := (+ -10::int #592)
+#497 := (uf_1 #496)
+#498 := (uf_3 #497)
+#499 := (= #484 #498)
#22 := 10::int
-#507 := (>= #620 10::int)
-#514 := (ite #507 #505 #531)
+#500 := (>= #592 10::int)
+#501 := (ite #500 #499 #485)
#4 := (:var 0 T1)
#21 := (uf_3 #4)
#707 := (pattern #21)
@@ -12333,12 +12554,12 @@
#212 := [mp #207 #211]: #193
#713 := [mp #212 #712]: #708
#336 := (not #708)
-#518 := (or #336 #514)
-#528 := [quant-inst]: #518
-#477 := [unit-resolution #528 #713]: #514
-#529 := (not #507)
-#498 := (<= #620 6::int)
-#610 := (= #620 6::int)
+#463 := (or #336 #501)
+#464 := [quant-inst]: #463
+#444 := [unit-resolution #464 #713]: #501
+#473 := (not #500)
+#453 := (<= #592 6::int)
+#597 := (= #592 6::int)
#10 := (:var 0 int)
#12 := (uf_1 #10)
#694 := (pattern #12)
@@ -12396,79 +12617,79 @@
#201 := [mp~ #99 #183]: #94
#700 := [mp #201 #699]: #695
#673 := (not #695)
-#591 := (or #673 #610)
-#526 := (>= 6::int 0::int)
-#527 := (not #526)
-#617 := (= 6::int #620)
-#621 := (or #617 #527)
-#592 := (or #673 #621)
-#595 := (iff #592 #591)
-#597 := (iff #591 #591)
-#593 := [rewrite]: #597
-#600 := (iff #621 #610)
-#614 := (or #610 false)
-#605 := (iff #614 #610)
-#606 := [rewrite]: #605
-#603 := (iff #621 #614)
-#613 := (iff #527 false)
+#576 := (or #673 #597)
+#607 := (>= 6::int 0::int)
+#591 := (not #607)
+#594 := (= 6::int #592)
+#595 := (or #594 #591)
+#577 := (or #673 #595)
+#579 := (iff #577 #576)
+#581 := (iff #576 #576)
+#582 := [rewrite]: #581
+#574 := (iff #595 #597)
+#586 := (or #597 false)
+#571 := (iff #586 #597)
+#573 := [rewrite]: #571
+#590 := (iff #595 #586)
+#588 := (iff #591 false)
#1 := true
#663 := (not true)
#666 := (iff #663 false)
#667 := [rewrite]: #666
-#611 := (iff #527 #663)
-#599 := (iff #526 true)
-#601 := [rewrite]: #599
-#612 := [monotonicity #601]: #611
-#609 := [trans #612 #667]: #613
-#608 := (iff #617 #610)
-#602 := [rewrite]: #608
-#604 := [monotonicity #602 #609]: #603
-#607 := [trans #604 #606]: #600
-#596 := [monotonicity #607]: #595
-#598 := [trans #596 #593]: #595
-#594 := [quant-inst]: #592
-#584 := [mp #594 #598]: #591
-#478 := [unit-resolution #584 #700]: #610
-#453 := (not #610)
-#454 := (or #453 #498)
-#455 := [th-lemma]: #454
-#456 := [unit-resolution #455 #478]: #498
-#458 := (not #498)
-#459 := (or #458 #529)
-#460 := [th-lemma]: #459
-#302 := [unit-resolution #460 #456]: #529
-#508 := (not #514)
-#490 := (or #508 #507 #531)
-#491 := [def-axiom]: #490
-#461 := [unit-resolution #491 #302 #477]: #531
-#404 := [symm #461]: #403
-#405 := (= #36 #476)
+#585 := (iff #591 #663)
+#598 := (iff #607 true)
+#584 := [rewrite]: #598
+#587 := [monotonicity #584]: #585
+#589 := [trans #587 #667]: #588
+#596 := (iff #594 #597)
+#593 := [rewrite]: #596
+#570 := [monotonicity #593 #589]: #590
+#575 := [trans #570 #573]: #574
+#580 := [monotonicity #575]: #579
+#572 := [trans #580 #582]: #579
+#578 := [quant-inst]: #577
+#583 := [mp #578 #572]: #576
+#448 := [unit-resolution #583 #700]: #597
+#445 := (not #597)
+#446 := (or #445 #453)
+#442 := [th-lemma]: #446
+#447 := [unit-resolution #442 #448]: #453
+#437 := (not #453)
+#427 := (or #437 #473)
+#429 := [th-lemma]: #427
+#430 := [unit-resolution #429 #447]: #473
+#471 := (not #501)
+#477 := (or #471 #500 #485)
+#478 := [def-axiom]: #477
+#433 := [unit-resolution #478 #430 #444]: #485
+#373 := [symm #433]: #372
+#374 := (= #36 #484)
#649 := (uf_2 #35)
-#582 := (+ -10::int #649)
-#553 := (uf_1 #582)
-#556 := (uf_3 #553)
-#401 := (= #556 #476)
-#417 := (= #553 #38)
-#415 := (= #582 6::int)
+#554 := (+ -10::int #649)
+#549 := (uf_1 #554)
+#545 := (uf_3 #549)
+#381 := (= #545 #484)
+#395 := (= #549 #38)
+#394 := (= #554 6::int)
#335 := (uf_2 #31)
#647 := -1::int
-#502 := (* -1::int #335)
-#463 := (+ #33 #502)
-#464 := (<= #463 0::int)
-#486 := (= #33 #335)
-#445 := (= #32 #31)
-#574 := (= #31 #32)
-#575 := (+ -10::int #335)
-#576 := (uf_1 #575)
-#577 := (uf_3 #576)
-#578 := (= #32 #577)
-#579 := (>= #335 10::int)
-#580 := (ite #579 #578 #574)
-#572 := (or #336 #580)
-#583 := [quant-inst]: #572
-#457 := [unit-resolution #583 #713]: #580
-#562 := (not #579)
-#554 := (<= #335 4::int)
+#459 := (* -1::int #335)
+#460 := (+ #33 #459)
+#302 := (<= #460 0::int)
+#458 := (= #33 #335)
+#426 := (= #32 #31)
+#555 := (= #31 #32)
+#551 := (+ -10::int #335)
+#552 := (uf_1 #551)
+#553 := (uf_3 #552)
+#556 := (= #32 #553)
+#557 := (>= #335 10::int)
+#558 := (ite #557 #556 #555)
+#560 := (or #336 #558)
+#533 := [quant-inst]: #560
+#434 := [unit-resolution #533 #713]: #558
+#535 := (not #557)
+#531 := (<= #335 4::int)
#324 := (= #335 4::int)
#659 := (or #673 #324)
#678 := (>= 4::int 0::int)
@@ -12498,110 +12719,125 @@
#277 := [trans #383 #385]: #382
#367 := [quant-inst]: #660
#655 := [mp #367 #277]: #659
-#462 := [unit-resolution #655 #700]: #324
-#441 := (not #324)
-#444 := (or #441 #554)
-#448 := [th-lemma]: #444
-#450 := [unit-resolution #448 #462]: #554
-#451 := (not #554)
-#449 := (or #451 #562)
-#452 := [th-lemma]: #449
-#440 := [unit-resolution #452 #450]: #562
-#561 := (not #580)
-#566 := (or #561 #579 #574)
-#567 := [def-axiom]: #566
-#443 := [unit-resolution #567 #440 #457]: #574
-#446 := [symm #443]: #445
-#442 := [monotonicity #446]: #486
-#447 := (not #486)
-#437 := (or #447 #464)
-#427 := [th-lemma]: #437
-#429 := [unit-resolution #427 #442]: #464
-#471 := (>= #463 0::int)
-#430 := (or #447 #471)
-#433 := [th-lemma]: #430
-#434 := [unit-resolution #433 #442]: #471
-#560 := (>= #335 4::int)
-#438 := (or #441 #560)
-#431 := [th-lemma]: #438
-#439 := [unit-resolution #431 #462]: #560
+#438 := [unit-resolution #655 #700]: #324
+#431 := (not #324)
+#439 := (or #431 #531)
+#432 := [th-lemma]: #439
+#435 := [unit-resolution #432 #438]: #531
+#436 := (not #531)
+#422 := (or #436 #535)
+#424 := [th-lemma]: #422
+#425 := [unit-resolution #424 #435]: #535
+#534 := (not #558)
+#540 := (or #534 #557 #555)
+#541 := [def-axiom]: #540
+#423 := [unit-resolution #541 #425 #434]: #555
+#408 := [symm #423]: #426
+#410 := [monotonicity #408]: #458
+#411 := (not #458)
+#412 := (or #411 #302)
+#413 := [th-lemma]: #412
+#414 := [unit-resolution #413 #410]: #302
+#461 := (>= #460 0::int)
+#415 := (or #411 #461)
+#416 := [th-lemma]: #415
+#417 := [unit-resolution #416 #410]: #461
+#512 := (>= #335 4::int)
+#418 := (or #431 #512)
+#419 := [th-lemma]: #418
+#420 := [unit-resolution #419 #438]: #512
#651 := (* -1::int #649)
#648 := (+ #34 #651)
-#625 := (<= #648 0::int)
+#521 := (<= #648 0::int)
#652 := (= #648 0::int)
-#643 := (>= #33 0::int)
-#435 := (not #471)
-#432 := (not #560)
-#436 := (or #643 #432 #435)
-#422 := [th-lemma]: #436
-#424 := [unit-resolution #422 #439 #434]: #643
-#644 := (not #643)
-#489 := (or #644 #652)
-#628 := (or #673 #644 #652)
+#630 := (>= #33 0::int)
+#421 := (not #461)
+#409 := (not #512)
+#398 := (or #630 #409 #421)
+#400 := [th-lemma]: #398
+#401 := [unit-resolution #400 #420 #417]: #630
+#468 := (not #630)
+#623 := (or #468 #652)
+#509 := (or #673 #468 #652)
#370 := (>= #34 0::int)
#371 := (not #370)
#650 := (= #34 #649)
#364 := (or #650 #371)
-#629 := (or #673 #364)
-#469 := (iff #629 #628)
-#636 := (or #673 #489)
-#466 := (iff #636 #628)
-#468 := [rewrite]: #466
-#630 := (iff #629 #636)
-#633 := (iff #364 #489)
-#646 := (or #652 #644)
-#631 := (iff #646 #489)
-#632 := [rewrite]: #631
-#487 := (iff #364 #646)
-#645 := (iff #371 #644)
-#638 := (iff #370 #643)
-#639 := [rewrite]: #638
-#640 := [monotonicity #639]: #645
+#510 := (or #673 #364)
+#619 := (iff #510 #509)
+#470 := (or #673 #623)
+#615 := (iff #470 #509)
+#616 := [rewrite]: #615
+#618 := (iff #510 #470)
+#624 := (iff #364 #623)
+#643 := 1::int
+#638 := (* 1::int #33)
+#639 := (>= #638 0::int)
+#640 := (not #639)
+#632 := (or #640 #652)
+#625 := (iff #632 #623)
+#469 := (iff #640 #468)
+#637 := (iff #639 #630)
+#635 := (= #638 #33)
+#636 := [rewrite]: #635
+#466 := [monotonicity #636]: #637
+#622 := [monotonicity #466]: #469
+#626 := [monotonicity #622]: #625
+#628 := (iff #364 #632)
+#488 := (or #652 #640)
+#633 := (iff #488 #632)
+#634 := [rewrite]: #633
+#489 := (iff #364 #488)
+#646 := (iff #371 #640)
+#644 := (iff #370 #639)
+#645 := [rewrite]: #644
+#487 := [monotonicity #645]: #646
#641 := (iff #650 #652)
#642 := [rewrite]: #641
-#488 := [monotonicity #642 #640]: #487
-#634 := [trans #488 #632]: #633
-#637 := [monotonicity #634]: #630
-#622 := [trans #637 #468]: #469
-#635 := [quant-inst]: #629
-#623 := [mp #635 #622]: #628
-#425 := [unit-resolution #623 #700]: #489
-#423 := [unit-resolution #425 #424]: #652
-#426 := (not #652)
-#408 := (or #426 #625)
-#410 := [th-lemma]: #408
-#411 := [unit-resolution #410 #423]: #625
-#626 := (>= #648 0::int)
-#412 := (or #426 #626)
-#413 := [th-lemma]: #412
-#414 := [unit-resolution #413 #423]: #626
-#416 := [th-lemma #414 #411 #439 #450 #434 #429]: #415
-#418 := [monotonicity #416]: #417
-#402 := [monotonicity #418]: #401
-#557 := (= #36 #556)
-#581 := (= #35 #36)
-#558 := (>= #649 10::int)
-#559 := (ite #558 #557 #581)
-#533 := (or #336 #559)
-#534 := [quant-inst]: #533
-#419 := [unit-resolution #534 #713]: #559
-#420 := (not #625)
-#409 := (or #558 #420 #432 #435)
-#421 := [th-lemma]: #409
-#398 := [unit-resolution #421 #411 #439 #434]: #558
-#428 := (not #558)
-#535 := (not #559)
-#539 := (or #535 #428 #557)
-#540 := [def-axiom]: #539
-#400 := [unit-resolution #540 #398 #419]: #557
-#406 := [trans #400 #402]: #405
-#399 := [trans #406 #404]: #39
+#631 := [monotonicity #642 #487]: #489
+#629 := [trans #631 #634]: #628
+#627 := [trans #629 #626]: #624
+#520 := [monotonicity #627]: #618
+#504 := [trans #520 #616]: #619
+#511 := [quant-inst]: #510
+#519 := [mp #511 #504]: #509
+#402 := [unit-resolution #519 #700]: #623
+#403 := [unit-resolution #402 #401]: #652
+#404 := (not #652)
+#405 := (or #404 #521)
+#406 := [th-lemma]: #405
+#399 := [unit-resolution #406 #403]: #521
+#522 := (>= #648 0::int)
+#407 := (or #404 #522)
+#392 := [th-lemma]: #407
+#393 := [unit-resolution #392 #403]: #522
+#396 := [th-lemma #393 #399 #420 #435 #417 #414]: #394
+#397 := [monotonicity #396]: #395
+#391 := [monotonicity #397]: #381
+#550 := (= #36 #545)
+#559 := (= #35 #36)
+#530 := (>= #649 10::int)
+#476 := (ite #530 #550 #559)
+#536 := (or #336 #476)
+#537 := [quant-inst]: #536
+#386 := [unit-resolution #537 #713]: #476
+#387 := (not #521)
+#388 := (or #530 #387 #409 #421)
+#380 := [th-lemma]: #388
+#389 := [unit-resolution #380 #399 #420 #417]: #530
+#538 := (not #530)
+#532 := (not #476)
+#506 := (or #532 #538 #550)
+#513 := [def-axiom]: #506
+#390 := [unit-resolution #513 #389 #386]: #550
+#365 := [trans #390 #391]: #374
+#375 := [trans #365 #373]: #39
#40 := (not #39)
#182 := [asserted]: #40
-[unit-resolution #182 #399]: false
-unsat
-
-/fwo5o8vhLVHyS4oYEs4QA 10833
+[unit-resolution #182 #375]: false
+unsat
+
+/fwo5o8vhLVHyS4oYEs4QA 10902
#2 := false
#22 := 0::int
#8 := 2::int
@@ -12677,18 +12913,18 @@
#604 := [trans #593 #597]: #562
#603 := [quant-inst]: #596
#606 := [mp #603 #604]: #628
-#528 := [unit-resolution #606 #817]: #566
-#521 := (not #566)
-#464 := (or #521 #608)
-#456 := [th-lemma]: #464
-#465 := [unit-resolution #456 #528]: #608
+#516 := [unit-resolution #606 #817]: #566
+#498 := (not #566)
+#432 := (or #498 #608)
+#411 := [th-lemma]: #432
+#413 := [unit-resolution #411 #516]: #608
decl uf_10 :: int
#52 := uf_10
#57 := (mod uf_10 2::int)
#243 := (* -1::int #57)
#244 := (+ #56 #243)
#447 := (>= #244 0::int)
-#387 := (not #447)
+#372 := (not #447)
#245 := (= #244 0::int)
#248 := (not #245)
#218 := (* -1::int #55)
@@ -12708,9 +12944,9 @@
#662 := (not #672)
#1 := true
#81 := [true-axiom]: true
-#520 := (or false #662)
-#523 := [th-lemma]: #520
-#524 := [unit-resolution #523 #81]: #662
+#514 := (or false #662)
+#515 := [th-lemma]: #514
+#513 := [unit-resolution #515 #81]: #662
#701 := (* -1::int #613)
#204 := -2::int
#691 := (* -2::int #222)
@@ -12723,82 +12959,58 @@
#546 := [th-lemma]: #545
#548 := [unit-resolution #546 #81]: #692
#549 := (not #692)
-#497 := (or #549 #694)
-#482 := [th-lemma]: #497
-#483 := [unit-resolution #482 #548]: #694
+#482 := (or #549 #694)
+#483 := [th-lemma]: #482
+#484 := [unit-resolution #483 #548]: #694
#536 := (not #448)
-#395 := [hypothesis]: #536
+#382 := [hypothesis]: #536
#555 := (* -1::int uf_10)
#573 := (+ #51 #555)
#543 := (<= #573 0::int)
#53 := (= #51 uf_10)
#256 := (not #253)
#259 := (or #248 #256)
-#502 := 1::int
#731 := (div uf_10 2::int)
-#515 := (* -1::int #731)
-#513 := (+ #640 #515)
+#723 := (* -2::int #731)
+#521 := (* -2::int #624)
+#529 := (+ #521 #723)
#618 := (div #51 2::int)
-#514 := (* -1::int #618)
-#516 := (+ #514 #513)
-#498 := (+ #243 #516)
-#500 := (+ #56 #498)
-#501 := (+ uf_10 #500)
-#503 := (>= #501 1::int)
-#486 := (not #503)
-#361 := (<= #244 0::int)
+#583 := (* -2::int #618)
+#522 := (+ #583 #529)
+#528 := (* -2::int #57)
+#525 := (+ #528 #522)
+#524 := (* 2::int #56)
+#526 := (+ #524 #525)
+#523 := (* 2::int uf_10)
+#512 := (+ #523 #526)
+#520 := (>= #512 2::int)
#453 := (not #259)
-#517 := [hypothesis]: #453
+#519 := [hypothesis]: #453
#440 := (or #259 #245)
#451 := [def-axiom]: #440
-#519 := [unit-resolution #451 #517]: #245
-#478 := (or #248 #361)
-#470 := [th-lemma]: #478
-#479 := [unit-resolution #470 #519]: #361
-#449 := (>= #252 0::int)
+#470 := [unit-resolution #451 #519]: #245
+#465 := (or #248 #447)
+#466 := [th-lemma]: #465
+#457 := [unit-resolution #466 #470]: #447
+#544 := (>= #573 0::int)
+#441 := (not #544)
#452 := (or #259 #253)
#380 := [def-axiom]: #452
-#480 := [unit-resolution #380 #517]: #253
-#471 := (or #256 #449)
-#481 := [th-lemma]: #471
-#462 := [unit-resolution #481 #480]: #449
-#487 := (not #361)
-#485 := (not #449)
-#476 := (or #486 #485 #487)
-#607 := (<= #620 0::int)
-#529 := (or #521 #607)
-#522 := [th-lemma]: #529
-#525 := [unit-resolution #522 #528]: #607
-#723 := (* -2::int #731)
-#724 := (+ #243 #723)
-#718 := (+ uf_10 #724)
-#720 := (<= #718 0::int)
-#722 := (= #718 0::int)
-#526 := (or false #722)
-#512 := [th-lemma]: #526
-#504 := [unit-resolution #512 #81]: #722
-#505 := (not #722)
-#506 := (or #505 #720)
-#507 := [th-lemma]: #506
-#508 := [unit-resolution #507 #504]: #720
-#509 := [hypothesis]: #361
-#583 := (* -2::int #618)
-#584 := (+ #583 #640)
-#585 := (+ #51 #584)
-#587 := (<= #585 0::int)
-#582 := (= #585 0::int)
-#510 := (or false #582)
-#499 := [th-lemma]: #510
-#511 := [unit-resolution #499 #81]: #582
-#488 := (not #582)
-#490 := (or #488 #587)
-#491 := [th-lemma]: #490
-#492 := [unit-resolution #491 #511]: #587
-#493 := [hypothesis]: #503
+#481 := [unit-resolution #380 #519]: #253
+#467 := (or #256 #448)
+#434 := [th-lemma]: #467
+#436 := [unit-resolution #434 #481]: #448
+#532 := (or #543 #536)
+#695 := (>= #699 0::int)
+#550 := (or #549 #695)
+#393 := [th-lemma]: #550
+#551 := [unit-resolution #393 #548]: #695
+#547 := (not #543)
+#552 := [hypothesis]: #547
#649 := (* -2::int #60)
#644 := (+ #218 #649)
#650 := (+ #51 #644)
-#636 := (>= #650 0::int)
+#631 := (<= #650 0::int)
#623 := (= #650 0::int)
#43 := (uf_7 uf_2 #35)
#44 := (uf_6 #34 #43)
@@ -12859,33 +13071,6 @@
#630 := [quant-inst]: #629
#531 := [unit-resolution #630 #824]: #623
#534 := (not #623)
-#494 := (or #534 #636)
-#495 := [th-lemma]: #494
-#496 := [unit-resolution #495 #531]: #636
-#489 := [hypothesis]: #449
-#484 := [th-lemma #483 #489 #496 #493 #492 #509 #508 #525 #524]: false
-#477 := [lemma #484]: #476
-#463 := [unit-resolution #477 #462 #479]: #486
-#727 := (>= #718 0::int)
-#466 := (or #505 #727)
-#457 := [th-lemma]: #466
-#467 := [unit-resolution #457 #504]: #727
-#434 := (or #248 #447)
-#436 := [th-lemma]: #434
-#437 := [unit-resolution #436 #519]: #447
-#544 := (>= #573 0::int)
-#445 := (not #544)
-#428 := (or #256 #448)
-#441 := [th-lemma]: #428
-#442 := [unit-resolution #441 #480]: #448
-#532 := (or #543 #536)
-#695 := (>= #699 0::int)
-#550 := (or #549 #695)
-#393 := [th-lemma]: #550
-#551 := [unit-resolution #393 #548]: #695
-#547 := (not #543)
-#552 := [hypothesis]: #547
-#631 := (<= #650 0::int)
#538 := (or #534 #631)
#540 := [th-lemma]: #538
#541 := [unit-resolution #540 #531]: #631
@@ -12896,8 +13081,8 @@
#533 := [unit-resolution #530 #81]: #666
#535 := [th-lemma #533 #539 #541 #552 #551]: false
#537 := [lemma #535]: #532
-#443 := [unit-resolution #537 #442]: #543
-#429 := (or #547 #445)
+#437 := [unit-resolution #537 #436]: #543
+#444 := (or #547 #441)
#764 := (not #53)
#771 := (or #764 #259)
#262 := (iff #53 #259)
@@ -12950,65 +13135,118 @@
#438 := (or #764 #259 #433)
#439 := [def-axiom]: #438
#772 := [unit-resolution #439 #267]: #771
-#444 := [unit-resolution #772 #517]: #764
-#435 := (or #53 #547 #445)
-#446 := [th-lemma]: #435
-#431 := [unit-resolution #446 #444]: #429
-#432 := [unit-resolution #431 #443]: #445
+#428 := [unit-resolution #772 #519]: #764
+#442 := (or #53 #547 #441)
+#443 := [th-lemma]: #442
+#445 := [unit-resolution #443 #428]: #444
+#435 := [unit-resolution #445 #437]: #441
+#584 := (+ #583 #640)
+#585 := (+ #51 #584)
#588 := (>= #585 0::int)
-#411 := (or #488 #588)
-#413 := [th-lemma]: #411
-#418 := [unit-resolution #413 #511]: #588
-#419 := [th-lemma #418 #432 #437 #467 #465 #463]: false
-#420 := [lemma #419]: #259
+#582 := (= #585 0::int)
+#499 := (or false #582)
+#511 := [th-lemma]: #499
+#488 := [unit-resolution #511 #81]: #582
+#490 := (not #582)
+#446 := (or #490 #588)
+#429 := [th-lemma]: #446
+#431 := [unit-resolution #429 #488]: #588
+#724 := (+ #243 #723)
+#718 := (+ uf_10 #724)
+#727 := (>= #718 0::int)
+#722 := (= #718 0::int)
+#503 := (or false #722)
+#504 := [th-lemma]: #503
+#505 := [unit-resolution #504 #81]: #722
+#506 := (not #722)
+#418 := (or #506 #727)
+#419 := [th-lemma]: #418
+#420 := [unit-resolution #419 #505]: #727
+#421 := [th-lemma #420 #413 #431 #435 #457]: #520
+#485 := (not #520)
+#361 := (<= #244 0::int)
+#479 := (or #248 #361)
+#480 := [th-lemma]: #479
+#471 := [unit-resolution #480 #470]: #361
+#449 := (>= #252 0::int)
+#462 := (or #256 #449)
+#463 := [th-lemma]: #462
+#464 := [unit-resolution #463 #481]: #449
+#476 := (not #361)
+#487 := (not #449)
+#477 := (or #485 #487 #476)
+#607 := (<= #620 0::int)
+#500 := (or #498 #607)
+#501 := [th-lemma]: #500
+#502 := [unit-resolution #501 #516]: #607
+#720 := (<= #718 0::int)
+#507 := (or #506 #720)
+#508 := [th-lemma]: #507
+#509 := [unit-resolution #508 #505]: #720
+#510 := [hypothesis]: #361
+#587 := (<= #585 0::int)
+#491 := (or #490 #587)
+#492 := [th-lemma]: #491
+#493 := [unit-resolution #492 #488]: #587
+#494 := [hypothesis]: #520
+#636 := (>= #650 0::int)
+#495 := (or #534 #636)
+#496 := [th-lemma]: #495
+#489 := [unit-resolution #496 #531]: #636
+#497 := [hypothesis]: #449
+#486 := [th-lemma #484 #497 #489 #494 #493 #510 #509 #502 #513]: false
+#478 := [lemma #486]: #477
+#456 := [unit-resolution #478 #464 #471]: #485
+#422 := [unit-resolution #456 #421]: false
+#423 := [lemma #422]: #259
#427 := (or #53 #453)
#768 := (or #53 #453 #433)
#770 := [def-axiom]: #768
#557 := [unit-resolution #770 #267]: #427
-#406 := [unit-resolution #557 #420]: #53
-#377 := (or #764 #543)
-#381 := [th-lemma]: #377
-#382 := [unit-resolution #381 #406]: #543
-#385 := [th-lemma #496 #382 #395 #483 #524]: false
-#386 := [lemma #385]: #448
-#390 := (or #253 #536)
-#408 := [hypothesis]: #485
-#409 := (or #764 #544)
-#397 := [th-lemma]: #409
-#399 := [unit-resolution #397 #406]: #544
-#400 := [th-lemma #399 #408 #533 #551 #541]: false
-#403 := [lemma #400]: #449
-#392 := (or #253 #536 #485)
-#394 := [th-lemma]: #392
-#657 := [unit-resolution #394 #403]: #390
-#658 := [unit-resolution #657 #386]: #253
+#399 := [unit-resolution #557 #423]: #53
+#385 := (or #764 #543)
+#386 := [th-lemma]: #385
+#387 := [unit-resolution #386 #399]: #543
+#379 := [th-lemma #489 #387 #382 #484 #513]: false
+#388 := [lemma #379]: #448
+#381 := (or #253 #536)
+#397 := [hypothesis]: #487
+#400 := (or #764 #544)
+#403 := [th-lemma]: #400
+#398 := [unit-resolution #403 #399]: #544
+#404 := [th-lemma #398 #397 #533 #551 #541]: false
+#378 := [lemma #404]: #449
+#395 := (or #253 #536 #487)
+#377 := [th-lemma]: #395
+#658 := [unit-resolution #377 #378]: #381
+#653 := [unit-resolution #658 #388]: #253
#450 := (or #453 #248 #256)
#454 := [def-axiom]: #450
-#762 := [unit-resolution #454 #420]: #259
-#664 := [unit-resolution #762 #658]: #248
-#372 := (or #245 #387)
-#560 := (+ #57 #640)
-#610 := (>= #560 0::int)
-#742 := (= #57 #624)
-#424 := (= #624 #57)
-#405 := [monotonicity #406]: #424
-#407 := [symm #405]: #742
-#705 := (not #742)
-#706 := (or #705 #610)
-#568 := [th-lemma]: #706
-#569 := [unit-resolution #568 #407]: #610
-#398 := [hypothesis]: #487
-#404 := [th-lemma #525 #398 #569]: false
-#378 := [lemma #404]: #361
-#379 := (or #245 #487 #387)
-#388 := [th-lemma]: #379
-#369 := [unit-resolution #388 #378]: #372
-#370 := [unit-resolution #369 #664]: #387
-#708 := (<= #560 0::int)
-#373 := (or #705 #708)
-#374 := [th-lemma]: #373
-#375 := [unit-resolution #374 #407]: #708
-[th-lemma #375 #370 #465]: false
+#664 := [unit-resolution #454 #423]: #259
+#665 := [unit-resolution #664 #653]: #248
+#373 := (or #245 #372)
+#708 := (+ #57 #640)
+#705 := (>= #708 0::int)
+#560 := (= #57 #624)
+#408 := (= #624 #57)
+#406 := [monotonicity #399]: #408
+#409 := [symm #406]: #560
+#706 := (not #560)
+#655 := (or #706 #705)
+#569 := [th-lemma]: #655
+#570 := [unit-resolution #569 #409]: #705
+#383 := [hypothesis]: #476
+#384 := [th-lemma #502 #383 #570]: false
+#389 := [lemma #384]: #361
+#369 := (or #245 #476 #372)
+#370 := [th-lemma]: #369
+#374 := [unit-resolution #370 #389]: #373
+#375 := [unit-resolution #374 #665]: #372
+#610 := (<= #708 0::int)
+#371 := (or #706 #610)
+#376 := [th-lemma]: #371
+#363 := [unit-resolution #376 #409]: #610
+[th-lemma #363 #375 #413]: false
unsat
s8LL71+1HTN+eIuEYeKhUw 1251
--- a/src/HOL/SMT/Examples/SMT_Examples.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/SMT/Examples/SMT_Examples.thy Mon Feb 08 17:13:45 2010 +0100
@@ -17,7 +17,7 @@
the following option is set to "false":
*}
-declare [[smt_record=false]]
+declare [[smt_record=false]]
--- a/src/HOL/Tools/Function/size.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/Function/size.ML Mon Feb 08 17:13:45 2010 +0100
@@ -153,7 +153,7 @@
val ctxt = ProofContext.init thy';
- val simpset1 = HOL_basic_ss addsimps @{thm add_0} :: @{thm add_0_right} ::
+ val simpset1 = HOL_basic_ss addsimps @{thm add_0} :: @{thm Nat.add_0_right} ::
size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites;
val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2);
--- a/src/HOL/Tools/Qelim/cooper.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Mon Feb 08 17:13:45 2010 +0100
@@ -79,9 +79,9 @@
| Const (@{const_name Algebras.less_eq}, _) $ y $ z =>
if term_of x aconv y then Le (Thm.dest_arg ct)
else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox
-| Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_) =>
+| Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_) =>
if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox
-| Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_)) =>
+| Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$y$_)) =>
if term_of x aconv y then
NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox
| _ => Nox)
@@ -222,8 +222,8 @@
| lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name Algebras.less_eq}, T) $ s $ t)) =
lin vs (Const (@{const_name Algebras.less}, T) $ t $ s)
| lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
- | lin (vs as x::_) (Const(@{const_name Ring_and_Field.dvd},_)$d$t) =
- HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} (numeral1 abs d, lint vs t)
+ | lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) =
+ HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t)
| lin (vs as x::_) ((b as Const("op =",_))$s$t) =
(case lint vs (subC$t$s) of
(t as a$(m$c$y)$r) =>
@@ -253,7 +253,7 @@
| is_intrel _ = false;
fun linearize_conv ctxt vs ct = case term_of ct of
- Const(@{const_name Ring_and_Field.dvd},_)$d$t =>
+ Const(@{const_name Rings.dvd},_)$d$t =>
let
val th = binop_conv (lint_conv ctxt vs) ct
val (d',t') = Thm.dest_binop (Thm.rhs_of th)
@@ -278,7 +278,7 @@
| _ => dth
end
end
-| Const (@{const_name Not},_)$(Const(@{const_name Ring_and_Field.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct
+| Const (@{const_name Not},_)$(Const(@{const_name Rings.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct
| t => if is_intrel t
then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop))
RS eq_reflection
@@ -301,7 +301,7 @@
if x aconv y andalso member (op =)
[@{const_name Algebras.less}, @{const_name Algebras.less_eq}] s
then (ins (dest_numeral c) acc, dacc) else (acc,dacc)
- | Const(@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_) =>
+ | Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_) =>
if x aconv y then (acc,ins (dest_numeral c) dacc) else (acc,dacc)
| Const("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
| Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
@@ -343,7 +343,7 @@
if x=y andalso member (op =)
[@{const_name Algebras.less}, @{const_name Algebras.less_eq}] s
then cv (l div dest_numeral c) t else Thm.reflexive t
- | Const(@{const_name Ring_and_Field.dvd},_)$d$(r as (Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_)) =>
+ | Const(@{const_name Rings.dvd},_)$d$(r as (Const(@{const_name Algebras.plus},_)$(Const(@{const_name Algebras.times},_)$c$y)$_)) =>
if x=y then
let
val k = l div dest_numeral c
@@ -562,7 +562,7 @@
| Const("False",_) => F
| Const(@{const_name Algebras.less},_)$t1$t2 => Lt (Sub (i_of_term vs t1,i_of_term vs t2))
| Const(@{const_name Algebras.less_eq},_)$t1$t2 => Le (Sub(i_of_term vs t1,i_of_term vs t2))
- | Const(@{const_name Ring_and_Field.dvd},_)$t1$t2 =>
+ | Const(@{const_name Rings.dvd},_)$t1$t2 =>
(Dvd(HOLogic.dest_number t1 |> snd, i_of_term vs t2) handle _ => cooper "Reification: unsupported dvd") (* FIXME avoid handle _ *)
| @{term "op = :: int => _"}$t1$t2 => Eq (Sub (i_of_term vs t1,i_of_term vs t2))
| @{term "op = :: bool => _ "}$t1$t2 => Iff(qf_of_term ps vs t1,qf_of_term ps vs t2)
--- a/src/HOL/Tools/lin_arith.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/lin_arith.ML Mon Feb 08 17:13:45 2010 +0100
@@ -236,7 +236,7 @@
end handle Rat.DIVZERO => NONE;
fun of_lin_arith_sort thy U =
- Sign.of_sort thy (U, @{sort Ring_and_Field.linordered_idom});
+ Sign.of_sort thy (U, @{sort Rings.linordered_idom});
fun allows_lin_arith thy (discrete : string list) (U as Type (D, [])) : bool * bool =
if of_lin_arith_sort thy U then (true, member (op =) discrete D)
@@ -814,8 +814,8 @@
addsimps @{thms ring_distribs}
addsimps [@{thm if_True}, @{thm if_False}]
addsimps
- [@{thm "monoid_add_class.add_0_left"},
- @{thm "monoid_add_class.add_0_right"},
+ [@{thm add_0_left},
+ @{thm add_0_right},
@{thm "Zero_not_Suc"}, @{thm "Suc_not_Zero"}, @{thm "le_0_eq"}, @{thm "One_nat_def"},
@{thm "order_less_irrefl"}, @{thm "zero_neq_one"}, @{thm "zero_less_one"},
@{thm "zero_le_one"}, @{thm "zero_neq_one"} RS not_sym, @{thm "not_one_le_zero"},
--- a/src/HOL/Tools/nat_arith.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/nat_arith.ML Mon Feb 08 17:13:45 2010 +0100
@@ -50,8 +50,8 @@
val mk_sum = mk_norm_sum;
val dest_sum = dest_sum;
val prove_conv = Arith_Data.prove_conv2;
- val norm_tac1 = Arith_Data.simp_all_tac [@{thm "add_Suc"}, @{thm "add_Suc_right"},
- @{thm "add_0"}, @{thm "add_0_right"}];
+ val norm_tac1 = Arith_Data.simp_all_tac [@{thm add_Suc}, @{thm add_Suc_right},
+ @{thm add_0}, @{thm Nat.add_0_right}];
val norm_tac2 = Arith_Data.simp_all_tac @{thms add_ac};
fun norm_tac ss = norm_tac1 ss THEN norm_tac2 ss;
fun gen_uncancel_tac rule = let val rule' = rule RS @{thm subst_equals}
--- a/src/HOL/Tools/nat_numeral_simprocs.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML Mon Feb 08 17:13:45 2010 +0100
@@ -124,7 +124,7 @@
(*Simplify 1*n and n*1 to n*)
-val add_0s = map rename_numerals [@{thm add_0}, @{thm add_0_right}];
+val add_0s = map rename_numerals [@{thm add_0}, @{thm Nat.add_0_right}];
val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}];
(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*)
@@ -136,7 +136,7 @@
val simplify_meta_eq =
Arith_Data.simplify_meta_eq
- ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right},
+ ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm Nat.add_0_right},
@{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules);
@@ -290,8 +290,8 @@
structure DvdCancelNumeralFactor = CancelNumeralFactorFun
(open CancelNumeralFactorCommon
val prove_conv = Arith_Data.prove_conv
- val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
- val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
+ val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd}
+ val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT
val cancel = @{thm nat_mult_dvd_cancel1} RS trans
val neg_exchanges = false
)
@@ -411,8 +411,8 @@
structure DvdCancelFactor = ExtractCommonTermFun
(open CancelFactorCommon
val prove_conv = Arith_Data.prove_conv
- val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
- val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
+ val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd}
+ val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT
fun simp_conv _ _ = SOME @{thm nat_mult_dvd_cancel_disj}
);
--- a/src/HOL/Tools/numeral_simprocs.ML Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML Mon Feb 08 17:13:45 2010 +0100
@@ -373,7 +373,7 @@
[@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}] @ simps
fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
val simplify_meta_eq = Arith_Data.simplify_meta_eq
- [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
+ [@{thm add_0}, @{thm Nat.add_0_right}, @{thm mult_zero_left},
@{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}];
end
@@ -561,8 +561,8 @@
structure DvdCancelFactor = ExtractCommonTermFun
(open CancelFactorCommon
val prove_conv = Arith_Data.prove_conv
- val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
- val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} Term.dummyT
+ val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd}
+ val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} Term.dummyT
fun simp_conv _ _ = SOME @{thm dvd_mult_cancel_left}
);
@@ -581,11 +581,11 @@
["(l::'a::idom) * m = n",
"(l::'a::idom) = m * n"],
K EqCancelFactor.proc),
- ("linlinordered_ring_le_cancel_factor",
+ ("linordered_ring_le_cancel_factor",
["(l::'a::linordered_ring) * m <= n",
"(l::'a::linordered_ring) <= m * n"],
K LeCancelFactor.proc),
- ("linlinordered_ring_less_cancel_factor",
+ ("linordered_ring_less_cancel_factor",
["(l::'a::linordered_ring) * m < n",
"(l::'a::linordered_ring) < m * n"],
K LessCancelFactor.proc),
--- a/src/HOL/Word/BinGeneral.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Word/BinGeneral.thy Mon Feb 08 17:13:45 2010 +0100
@@ -742,7 +742,7 @@
lemma sb_inc_lem':
"(a::int) < - (2^k) \<Longrightarrow> a + 2^k + 2^(Suc k) <= (a + 2^k) mod 2^(Suc k)"
- by (rule iffD1 [OF less_diff_eq, THEN sb_inc_lem, simplified OrderedGroup.diff_0])
+ by (rule sb_inc_lem) simp
lemma sbintrunc_inc:
"x < - (2^n) ==> x + 2^(Suc n) <= sbintrunc n x"
--- a/src/HOL/Word/Word.thy Mon Feb 08 15:25:00 2010 +0100
+++ b/src/HOL/Word/Word.thy Mon Feb 08 17:13:45 2010 +0100
@@ -2,7 +2,7 @@
Author: Gerwin Klein, NICTA
*)
-header {* Word Library interafce *}
+header {* Word Library interface *}
theory Word
imports WordGenLib