--- a/NEWS Mon Apr 26 09:37:46 2010 -0700
+++ b/NEWS Mon Apr 26 09:45:22 2010 -0700
@@ -119,8 +119,12 @@
*** HOL ***
* Abstract algebra:
- * class division_by_zero includes division_ring;
+ * classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
+ include rule inverse 0 = 0 -- subsumes former division_by_zero class.
* numerous lemmas have been ported from field to division_ring;
+ * dropped theorem group group_simps, use algebra_simps instead;
+ * dropped theorem group ring_simps, use field_simps instead;
+ * proper theorem collection field_simps subsumes former theorem groups field_eq_simps and field_simps;
* dropped lemma eq_minus_self_iff which is a duplicate for equal_neg_zero.
INCOMPATIBILITY.
--- a/doc-src/IsarImplementation/Thy/Logic.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Logic.thy Mon Apr 26 09:45:22 2010 -0700
@@ -128,8 +128,7 @@
@{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
@{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
@{index_ML Sign.add_types: "(binding * int * mixfix) list -> theory -> theory"} \\
- @{index_ML Sign.add_tyabbrs_i: "
- (binding * string list * typ * mixfix) list -> theory -> theory"} \\
+ @{index_ML Sign.add_type_abbrev: "binding * string list * typ -> theory -> theory"} \\
@{index_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
@{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
@{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
@@ -168,9 +167,9 @@
type constructors @{text "\<kappa>"} with @{text "k"} arguments and
optional mixfix syntax.
- \item @{ML Sign.add_tyabbrs_i}~@{text "[(\<kappa>, \<^vec>\<alpha>, \<tau>, mx), \<dots>]"}
- defines a new type abbreviation @{text "(\<^vec>\<alpha>)\<kappa> = \<tau>"} with
- optional mixfix syntax.
+ \item @{ML Sign.add_type_abbrev}~@{text "(\<kappa>, \<^vec>\<alpha>,
+ \<tau>)"} defines a new type abbreviation @{text
+ "(\<^vec>\<alpha>)\<kappa> = \<tau>"}.
\item @{ML Sign.primitive_class}~@{text "(c, [c\<^isub>1, \<dots>,
c\<^isub>n])"} declares a new class @{text "c"}, together with class
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex Mon Apr 26 09:37:46 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Mon Apr 26 09:45:22 2010 -0700
@@ -139,8 +139,7 @@
\indexdef{}{ML}{Sign.subsort}\verb|Sign.subsort: theory -> sort * sort -> bool| \\
\indexdef{}{ML}{Sign.of\_sort}\verb|Sign.of_sort: theory -> typ * sort -> bool| \\
\indexdef{}{ML}{Sign.add\_types}\verb|Sign.add_types: (binding * int * mixfix) list -> theory -> theory| \\
- \indexdef{}{ML}{Sign.add\_tyabbrs\_i}\verb|Sign.add_tyabbrs_i: |\isasep\isanewline%
-\verb| (binding * string list * typ * mixfix) list -> theory -> theory| \\
+ \indexdef{}{ML}{Sign.add\_type\_abbrev}\verb|Sign.add_type_abbrev: binding * string list * typ -> theory -> theory| \\
\indexdef{}{ML}{Sign.primitive\_class}\verb|Sign.primitive_class: binding * class list -> theory -> theory| \\
\indexdef{}{ML}{Sign.primitive\_classrel}\verb|Sign.primitive_classrel: class * class -> theory -> theory| \\
\indexdef{}{ML}{Sign.primitive\_arity}\verb|Sign.primitive_arity: arity -> theory -> theory| \\
@@ -176,9 +175,7 @@
type constructors \isa{{\isasymkappa}} with \isa{k} arguments and
optional mixfix syntax.
- \item \verb|Sign.add_tyabbrs_i|~\isa{{\isacharbrackleft}{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec {\isasymalpha}{\isacharcomma}\ {\isasymtau}{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}}
- defines a new type abbreviation \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}} with
- optional mixfix syntax.
+ \item \verb|Sign.add_type_abbrev|~\isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec {\isasymalpha}{\isacharcomma}\ {\isasymtau}{\isacharparenright}} defines a new type abbreviation \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}}.
\item \verb|Sign.primitive_class|~\isa{{\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub n{\isacharbrackright}{\isacharparenright}} declares a new class \isa{c}, together with class
relations \isa{c\ {\isasymsubseteq}\ c\isactrlisub i}, for \isa{i\ {\isacharequal}\ {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ n}.
--- a/src/HOL/Big_Operators.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Big_Operators.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1033,12 +1033,12 @@
by (erule finite_induct) (auto simp add: insert_Diff_if)
lemma setprod_inversef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_ring_inverse_zero}"
shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
by (erule finite_induct) auto
lemma setprod_dividef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_ring_inverse_zero}"
shows "finite A
==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
apply (subgoal_tac
@@ -1140,7 +1140,7 @@
using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
by simp
then have ?thesis using a cA
- by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
+ by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)}
ultimately show ?thesis by blast
qed
--- a/src/HOL/Boogie/Tools/boogie_commands.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Mon Apr 26 09:45:22 2010 -0700
@@ -99,7 +99,7 @@
|> (fn ctxt1 => ctxt1
|> prepare
|-> (fn us => fn ctxt2 => ctxt2
- |> Proof.theorem_i NONE (fn thmss => fn ctxt =>
+ |> Proof.theorem NONE (fn thmss => fn ctxt =>
let val export = map (finish ctxt1) o ProofContext.export ctxt ctxt2
in after_qed (map export thmss) ctxt end) [map (rpair []) us]))
end
@@ -188,7 +188,7 @@
fun prove thy meth vc =
ProofContext.init thy
- |> Proof.theorem_i NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
+ |> Proof.theorem NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
|> Proof.apply meth
|> Seq.hd
|> Proof.global_done_proof
--- a/src/HOL/Complex.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Complex.thy Mon Apr 26 09:45:22 2010 -0700
@@ -99,7 +99,7 @@
subsection {* Multiplication and Division *}
-instantiation complex :: "{field, division_by_zero}"
+instantiation complex :: "{field, division_ring_inverse_zero}"
begin
definition
--- a/src/HOL/Decision_Procs/Decision_Procs.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Decision_Procs/Decision_Procs.thy Mon Apr 26 09:45:22 2010 -0700
@@ -8,4 +8,4 @@
"ex/Commutative_Ring_Ex" "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex"
begin
-end
\ No newline at end of file
+end
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Apr 26 09:45:22 2010 -0700
@@ -27,7 +27,7 @@
"tmsize (CNP n c a) = 3 + polysize c + tmsize a "
(* Semantics of terms tm *)
-consts Itm :: "'a::{ring_char_0,division_by_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
+consts Itm :: "'a::{ring_char_0,division_ring_inverse_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
primrec
"Itm vs bs (CP c) = (Ipoly vs c)"
"Itm vs bs (Bound n) = bs!n"
@@ -239,7 +239,7 @@
lemma tmadd[simp]: "Itm vs bs (tmadd (t,s)) = Itm vs bs (Add t s)"
apply (induct t s rule: tmadd.induct, simp_all add: Let_def)
apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
-apply (case_tac "n1 = n2", simp_all add: ring_simps)
+apply (case_tac "n1 = n2", simp_all add: field_simps)
apply (simp only: right_distrib[symmetric])
by (auto simp del: polyadd simp add: polyadd[symmetric])
@@ -259,7 +259,7 @@
"tmmul t = (\<lambda> i. Mul i t)"
lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
-by (induct t arbitrary: i rule: tmmul.induct, simp_all add: ring_simps)
+by (induct t arbitrary: i rule: tmmul.induct, simp_all add: field_simps)
lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
by (induct t arbitrary: i rule: tmmul.induct, auto )
@@ -270,7 +270,7 @@
by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def)
lemma tmmul_allpolys_npoly[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
definition tmneg :: "tm \<Rightarrow> tm" where
@@ -296,7 +296,7 @@
using tmneg_def by simp
lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp
lemma tmneg_allpolys_npoly[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
unfolding tmneg_def by auto
@@ -310,7 +310,7 @@
lemma tmsub_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmsub t s )"
using tmsub_def by simp
lemma tmsub_allpolys_npoly[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
unfolding tmsub_def by (simp add: isnpoly_def)
@@ -324,8 +324,8 @@
"simptm (CNP n c t) = (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))"
lemma polynate_stupid:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
- shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
+ shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_ring_inverse_zero, field})"
apply (subst polynate[symmetric])
apply simp
done
@@ -345,7 +345,7 @@
lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))"
by (simp_all add: isnpoly_def)
lemma simptm_allpolys_npoly[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "allpolys isnpoly (simptm p)"
by (induct p rule: simptm.induct, auto simp add: Let_def)
@@ -369,14 +369,14 @@
"tmbound 0 (snd (split0 t)) \<and> (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)"
apply (induct t rule: split0.induct)
apply simp
- apply (simp add: Let_def split_def ring_simps)
- apply (simp add: Let_def split_def ring_simps)
- apply (simp add: Let_def split_def ring_simps)
- apply (simp add: Let_def split_def ring_simps)
- apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def field_simps)
+ apply (simp add: Let_def split_def field_simps)
+ apply (simp add: Let_def split_def field_simps)
+ apply (simp add: Let_def split_def field_simps)
+ apply (simp add: Let_def split_def field_simps)
apply (simp add: Let_def split_def mult_assoc right_distrib[symmetric])
- apply (simp add: Let_def split_def ring_simps)
- apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def field_simps)
+ apply (simp add: Let_def split_def field_simps)
done
lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
@@ -387,7 +387,7 @@
qed
lemma split0_nb0:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "split0 t = (c',t') \<Longrightarrow> tmbound 0 t'"
proof-
fix c' t'
@@ -395,7 +395,7 @@
with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp
qed
-lemma split0_nb0'[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+lemma split0_nb0'[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows "tmbound0 (snd (split0 t))"
using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff)
@@ -418,7 +418,7 @@
lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
by (induct p rule: split0.induct, auto simp add: isnpoly_def Let_def split_def split0_stupid)
-lemma isnpoly_fst_split0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+lemma isnpoly_fst_split0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
shows
"allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
by (induct p rule: split0.induct,
@@ -447,7 +447,7 @@
by (induct p rule: fmsize.induct) simp_all
(* Semantics of formulae (fm) *)
-consts Ifm ::"'a::{division_by_zero,linordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
+consts Ifm ::"'a::{division_ring_inverse_zero,linordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
primrec
"Ifm vs bs T = True"
"Ifm vs bs F = False"
@@ -969,24 +969,24 @@
definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
-lemma simplt_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simplt_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "islin (simplt t)"
unfolding simplt_def
using split0_nb0'
by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
-lemma simple_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simple_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "islin (simple t)"
unfolding simple_def
using split0_nb0'
by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
-lemma simpeq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpeq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "islin (simpeq t)"
unfolding simpeq_def
using split0_nb0'
by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
-lemma simpneq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpneq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "islin (simpneq t)"
unfolding simpneq_def
using split0_nb0'
@@ -994,7 +994,7 @@
lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
by (cases "split0 s", auto)
-lemma split0_npoly: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma split0_npoly: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
and n: "allpolys isnpoly t"
shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))"
using n
@@ -1083,7 +1083,7 @@
apply (case_tac poly, auto)
done
-lemma simplt_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simplt_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
using split0 [of "simptm t" vs bs]
proof(simp add: simplt_def Let_def split_def)
@@ -1100,7 +1100,7 @@
fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb)
qed
-lemma simple_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simple_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
using split0 [of "simptm t" vs bs]
proof(simp add: simple_def Let_def split_def)
@@ -1117,7 +1117,7 @@
fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb)
qed
-lemma simpeq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpeq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
using split0 [of "simptm t" vs bs]
proof(simp add: simpeq_def Let_def split_def)
@@ -1134,7 +1134,7 @@
fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb)
qed
-lemma simpneq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpneq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
using split0 [of "simptm t" vs bs]
proof(simp add: simpneq_def Let_def split_def)
@@ -1267,7 +1267,7 @@
lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
by(induct p arbitrary: bs rule: simpfm.induct, auto)
-lemma simpfm_bound0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpfm_bound0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
by (induct p rule: simpfm.induct, auto)
@@ -1296,7 +1296,7 @@
lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)" by (simp add: disj_def)
lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)" by (simp add: conj_def)
-lemma assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "qfree p \<Longrightarrow> islin (simpfm p)"
apply (induct p rule: simpfm.induct)
apply (simp_all add: conj_lin disj_lin)
@@ -1698,11 +1698,11 @@
{assume c: "?N c > 0"
from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]
have px': "x < - ?Nt x s / ?N c"
- by (auto simp add: not_less ring_simps)
+ by (auto simp add: not_less field_simps)
{assume y: "y < - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
- hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
{assume y: "y > -?Nt x s / ?N c"
@@ -1715,11 +1715,11 @@
{assume c: "?N c < 0"
from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]
have px': "x > - ?Nt x s / ?N c"
- by (auto simp add: not_less ring_simps)
+ by (auto simp add: not_less field_simps)
{assume y: "y > - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
- hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
{assume y: "y < -?Nt x s / ?N c"
@@ -1743,11 +1743,11 @@
moreover
{assume c: "?N c > 0"
from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]
- have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps)
+ have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less field_simps)
{assume y: "y < - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
- hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
{assume y: "y > -?Nt x s / ?N c"
@@ -1759,11 +1759,11 @@
moreover
{assume c: "?N c < 0"
from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]
- have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps)
+ have px': "x >= - ?Nt x s / ?N c" by (simp add: field_simps)
{assume y: "y > - ?Nt x s / ?N c"
hence "y * ?N c < - ?Nt x s"
by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
- hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
moreover
{assume y: "y < -?Nt x s / ?N c"
@@ -1787,7 +1787,7 @@
moreover
{assume c: "?N c > 0" hence cnz: "?N c \<noteq> 0" by simp
from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
- have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+ have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
{assume y: "y < -?Nt x s / ?N c"
with ly have eu: "l < - ?Nt x s / ?N c" by auto
with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
@@ -1802,7 +1802,7 @@
moreover
{assume c: "?N c < 0" hence cnz: "?N c \<noteq> 0" by simp
from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
- have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+ have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
{assume y: "y < -?Nt x s / ?N c"
with ly have eu: "l < - ?Nt x s / ?N c" by auto
with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
@@ -1829,7 +1829,7 @@
moreover
{assume c: "?N c \<noteq> 0"
from yne c eq_divide_eq[of "y" "- ?Nt x s" "?N c"] have ?case
- by (simp add: ring_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
+ by (simp add: field_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
ultimately show ?case by blast
qed (auto simp add: nth_pos2 tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
@@ -1844,7 +1844,7 @@
lemma half_sum_eq: "(u + u) / (1+1) = (u::'a::{linordered_field})"
proof-
- have "(u + u) = (1 + 1) * u" by (simp add: ring_simps)
+ have "(u + u) = (1 + 1) * u" by (simp add: field_simps)
hence "(u + u) / (1+1) = (1 + 1)*u / (1 + 1)" by simp
with nonzero_mult_divide_cancel_left[OF one_plus_one_nonzero, of u] show ?thesis by simp
qed
@@ -1987,7 +1987,7 @@
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) = 0"
using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r= 0"
- by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+ by (simp add: field_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r = 0" using d by simp
finally have ?thesis using c d
@@ -2003,7 +2003,7 @@
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) = 0"
using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r= 0"
- by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+ by (simp add: field_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r = 0" using c by simp
finally have ?thesis using c d
apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
@@ -2014,19 +2014,19 @@
{assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r = 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) =0 "
using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r =0"
- using nonzero_mult_divide_cancel_left[OF dc] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ using nonzero_mult_divide_cancel_left [OF dc] c d
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex ring_simps)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex field_simps)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
done }
ultimately show ?thesis by blast
qed
@@ -2075,7 +2075,7 @@
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) \<noteq> 0"
using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r\<noteq> 0"
- by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+ by (simp add: field_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r \<noteq> 0" using d by simp
finally have ?thesis using c d
@@ -2091,7 +2091,7 @@
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) \<noteq> 0"
using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r \<noteq> 0"
- by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+ by (simp add: field_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r \<noteq> 0" using c by simp
finally have ?thesis using c d
apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
@@ -2102,7 +2102,7 @@
{assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r \<noteq> 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2110,11 +2110,11 @@
using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r \<noteq> 0"
using nonzero_mult_divide_cancel_left[OF dc] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex ring_simps)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex field_simps)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
done }
ultimately show ?thesis by blast
qed
@@ -2169,7 +2169,7 @@
from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2178,11 +2178,11 @@
using dc' dc'' mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r < 0"
using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using dc c d nc nd dc'
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- by (simp add: ring_simps order_less_not_sym[OF dc])}
+ by (simp add: field_simps order_less_not_sym[OF dc])}
moreover
{assume dc: "?c*?d < 0"
@@ -2191,7 +2191,7 @@
hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2201,78 +2201,78 @@
using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r < 0"
using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using dc c d nc nd
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- by (simp add: ring_simps order_less_not_sym[OF dc]) }
+ by (simp add: field_simps order_less_not_sym[OF dc]) }
moreover
{assume c: "?c > 0" and d: "?d=0"
from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
from c have c': "(1 + 1)*?c \<noteq> 0" by simp
- from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) < 0"
using c mult_less_cancel_left_disj[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
also have "\<dots> \<longleftrightarrow> - ?a*?t+ (1 + 1)*?c *?r < 0"
using nonzero_mult_divide_cancel_left[OF c'] c
- by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using c order_less_not_sym[OF c] less_imp_neq[OF c]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
moreover
{assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \<noteq> 0" by simp
from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
- from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) > 0"
using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r < 0"
using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using c order_less_not_sym[OF c] less_imp_neq[OF c]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
moreover
moreover
{assume c: "?c = 0" and d: "?d>0"
from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
from d have d': "(1 + 1)*?d \<noteq> 0" by simp
- from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) < 0"
using d mult_less_cancel_left_disj[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
also have "\<dots> \<longleftrightarrow> - ?a*?s+ (1 + 1)*?d *?r < 0"
using nonzero_mult_divide_cancel_left[OF d'] d
- by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using d order_less_not_sym[OF d] less_imp_neq[OF d]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps) }
moreover
{assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \<noteq> 0" by simp
from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
- from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) > 0"
using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r < 0"
using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using d order_less_not_sym[OF d] less_imp_neq[OF d]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
ultimately show ?thesis by blast
qed
@@ -2325,7 +2325,7 @@
from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2334,11 +2334,11 @@
using dc' dc'' mult_le_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r <= 0"
using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using dc c d nc nd dc'
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- by (simp add: ring_simps order_less_not_sym[OF dc])}
+ by (simp add: field_simps order_less_not_sym[OF dc])}
moreover
{assume dc: "?c*?d < 0"
@@ -2347,7 +2347,7 @@
hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
from add_frac_eq[OF c d, of "- ?t" "- ?s"]
have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0"
by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2357,78 +2357,78 @@
using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r <= 0"
using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using dc c d nc nd
- apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
- by (simp add: ring_simps order_less_not_sym[OF dc]) }
+ by (simp add: field_simps order_less_not_sym[OF dc]) }
moreover
{assume c: "?c > 0" and d: "?d=0"
from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
from c have c': "(1 + 1)*?c \<noteq> 0" by simp
- from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) <= 0"
using c mult_le_cancel_left[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
also have "\<dots> \<longleftrightarrow> - ?a*?t+ (1 + 1)*?c *?r <= 0"
using nonzero_mult_divide_cancel_left[OF c'] c
- by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using c order_less_not_sym[OF c] less_imp_neq[OF c]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
moreover
{assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \<noteq> 0" by simp
from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
- from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) >= 0"
using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r <= 0"
using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using c order_less_not_sym[OF c] less_imp_neq[OF c]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
moreover
moreover
{assume c: "?c = 0" and d: "?d>0"
from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
from d have d': "(1 + 1)*?d \<noteq> 0" by simp
- from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) <= 0"
using d mult_le_cancel_left[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
also have "\<dots> \<longleftrightarrow> - ?a*?s+ (1 + 1)*?d *?r <= 0"
using nonzero_mult_divide_cancel_left[OF d'] d
- by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using d order_less_not_sym[OF d] less_imp_neq[OF d]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
moreover
{assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \<noteq> 0" by simp
from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
- from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) >= 0"
using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r <= 0"
using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
- by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
finally have ?thesis using c d nc nd
- apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
apply (simp only: one_add_one_is_two[symmetric] of_int_add)
using d order_less_not_sym[OF d] less_imp_neq[OF d]
- by (simp add: ring_simps ) }
+ by (simp add: field_simps ) }
ultimately show ?thesis by blast
qed
@@ -2519,7 +2519,7 @@
lemma remdps_set[simp]: "set (remdps xs) = set xs"
by (induct xs rule: remdps.induct, auto)
-lemma simpfm_lin: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpfm_lin: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "qfree p \<Longrightarrow> islin (simpfm p)"
by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin)
@@ -2551,7 +2551,7 @@
{fix c t d s assume ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
from U_l ctU dsU have norm: "isnpoly c" "isnpoly d" by auto
from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
- have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: ring_simps)}
+ have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: field_simps)}
hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))" by clarsimp
{fix x assume xUp: "x \<in> set ?Up"
then obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
@@ -2616,7 +2616,7 @@
let ?s = "Itm vs (x # bs) s"
let ?t = "Itm vs (x # bs) t"
have eq2: "\<And>(x::'a). x + x = (1 + 1) * x"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
{assume "?c = 0 \<and> ?d = 0"
with ct have ?D by simp}
moreover
@@ -2747,12 +2747,12 @@
using lp tnb
by (induct p c t rule: msubstpos.induct, auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
-lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
shows "bound0 (msubstneg p c t)"
using lp tnb
by (induct p c t rule: msubstneg.induct, auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
-lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
shows "bound0 (msubst2 p c t)"
using lp tnb
by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
@@ -2899,14 +2899,14 @@
by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
from msubst2[OF lq norm2(1) z(1), of x bs]
msubst2[OF lq norm2(2) z(2), of x bs] H
- show ?rhs by (simp add: ring_simps)
+ show ?rhs by (simp add: field_simps)
next
assume H: ?rhs
hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
from msubst2[OF lq norm2(1) z(1), of x bs]
msubst2[OF lq norm2(2) z(2), of x bs] H
- show ?lhs by (simp add: ring_simps)
+ show ?lhs by (simp add: field_simps)
qed}
hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
by clarsimp
@@ -3156,54 +3156,54 @@
*} "Parametric QE for linear Arithmetic over fields, Version 2"
-lemma "\<exists>(x::'a::{division_by_zero,linordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
- apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "y::'a::{division_by_zero,linordered_field,number_ring}")
- apply (simp add: ring_simps)
+lemma "\<exists>(x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+ apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "y::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+ apply (simp add: field_simps)
apply (rule spec[where x=y])
- apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "z::'a::{division_by_zero,linordered_field,number_ring}")
+ apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "z::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
by simp
text{* Collins/Jones Problem *}
(*
-lemma "\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+lemma "\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
proof-
- have "(\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: ring_simps)
+ have "(\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: field_simps)
have "?rhs"
- apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "a::'a::{division_by_zero,linordered_field,number_ring}" "b::'a::{division_by_zero,linordered_field,number_ring}")
- apply (simp add: ring_simps)
+ apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "a::'a::{linordered_field, division_ring_inverse_zero, number_ring}" "b::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+ apply (simp add: field_simps)
oops
*)
(*
-lemma "ALL (x::'a::{division_by_zero,linordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "t::'a::{division_by_zero,linordered_field,number_ring}")
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "t::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
oops
*)
-lemma "\<exists>(x::'a::{division_by_zero,linordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
- apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "y::'a::{division_by_zero,linordered_field,number_ring}")
- apply (simp add: ring_simps)
+lemma "\<exists>(x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+ apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "y::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+ apply (simp add: field_simps)
apply (rule spec[where x=y])
- apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "z::'a::{division_by_zero,linordered_field,number_ring}")
+ apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "z::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
by simp
text{* Collins/Jones Problem *}
(*
-lemma "\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+lemma "\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
proof-
- have "(\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: ring_simps)
+ have "(\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: field_simps)
have "?rhs"
- apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "a::'a::{division_by_zero,linordered_field,number_ring}" "b::'a::{division_by_zero,linordered_field,number_ring}")
+ apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "a::'a::{linordered_field, division_ring_inverse_zero, number_ring}" "b::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
apply simp
oops
*)
(*
-lemma "ALL (x::'a::{division_by_zero,linordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "t::'a::{division_by_zero,linordered_field,number_ring}")
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "t::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
apply (simp add: field_simps linorder_neq_iff[symmetric])
apply ferrack
oops
--- a/src/HOL/Decision_Procs/Polynomial_List.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy Mon Apr 26 09:45:22 2010 -0700
@@ -283,11 +283,11 @@
apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe)
apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe)
apply (drule_tac x = "Suc (length q)" in spec)
-apply (auto simp add: ring_simps)
+apply (auto simp add: field_simps)
apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
apply (drule_tac x = m in spec)
-apply (auto simp add:ring_simps)
+apply (auto simp add:field_simps)
done
lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0, standard]
@@ -327,7 +327,7 @@
apply (drule_tac x = "a#i" in spec)
apply (auto simp only: poly_mult List.list.size)
apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
done
lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard]
@@ -413,7 +413,7 @@
by (auto intro!: ext)
lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
-by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult)
+by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib)
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Apr 26 09:45:22 2010 -0700
@@ -230,7 +230,7 @@
subsection{* Semantics of the polynomial representation *}
-consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_by_zero,field}"
+consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_ring_inverse_zero,field}"
primrec
"Ipoly bs (C c) = INum c"
"Ipoly bs (Bound n) = bs!n"
@@ -241,7 +241,7 @@
"Ipoly bs (Pw t n) = (Ipoly bs t) ^ n"
"Ipoly bs (CN c n p) = (Ipoly bs c) + (bs!n)*(Ipoly bs p)"
abbreviation
- Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_by_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+ Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_ring_inverse_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p"
lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i"
@@ -322,7 +322,7 @@
qed auto
lemma polyadd[simp]: "Ipoly bs (polyadd (p,q)) = (Ipoly bs p) + (Ipoly bs q)"
-by (induct p q rule: polyadd.induct, auto simp add: Let_def ring_simps right_distrib[symmetric] simp del: right_distrib)
+by (induct p q rule: polyadd.induct, auto simp add: Let_def field_simps right_distrib[symmetric] simp del: right_distrib)
lemma polyadd_norm: "\<lbrakk> isnpoly p ; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polyadd(p,q))"
using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
@@ -394,7 +394,7 @@
qed simp_all
lemma polymul_properties:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> min n0 n1"
shows "isnpolyh (p *\<^sub>p q) (min n0 n1)"
and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p)"
@@ -565,22 +565,22 @@
qed auto
lemma polymul[simp]: "Ipoly bs (p *\<^sub>p q) = (Ipoly bs p) * (Ipoly bs q)"
-by(induct p q rule: polymul.induct, auto simp add: ring_simps)
+by(induct p q rule: polymul.induct, auto simp add: field_simps)
lemma polymul_normh:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (p *\<^sub>p q) (min n0 n1)"
using polymul_properties(1) by blast
lemma polymul_eq0_iff:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p) "
using polymul_properties(2) by blast
lemma polymul_degreen:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 else degreen p m + degreen q m)"
using polymul_properties(3) by blast
lemma polymul_norm:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polymul (p,q))"
using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp
@@ -591,7 +591,7 @@
by (induct p arbitrary: n0, auto)
lemma monic_eqI: assumes np: "isnpolyh p n0"
- shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_by_zero,field})"
+ shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_ring_inverse_zero,field})"
unfolding monic_def Let_def
proof(cases "headconst p = 0\<^sub>N", simp_all add: headconst_zero[OF np])
let ?h = "headconst p"
@@ -629,13 +629,13 @@
lemma polysub_norm: "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polysub(p,q))"
using polyadd_norm polyneg_norm by (simp add: polysub_def)
-lemma polysub_same_0[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma polysub_same_0[simp]: assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "isnpolyh p n0 \<Longrightarrow> polysub (p, p) = 0\<^sub>p"
unfolding polysub_def split_def fst_conv snd_conv
by (induct p arbitrary: n0,auto simp add: Let_def Nsub0[simplified Nsub_def])
lemma polysub_0:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p -\<^sub>p q = 0\<^sub>p) = (p = q)"
unfolding polysub_def split_def fst_conv snd_conv
apply (induct p q arbitrary: n0 n1 rule:polyadd.induct, simp_all add: Nsub0[simplified Nsub_def])
@@ -657,7 +657,7 @@
done
text{* polypow is a power function and preserves normal forms *}
-lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{ring_char_0,division_by_zero,field})) ^ n"
+lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{field, division_ring_inverse_zero, ring_char_0})) ^ n"
proof(induct n rule: polypow.induct)
case 1 thus ?case by simp
next
@@ -688,7 +688,7 @@
qed
lemma polypow_normh:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "isnpolyh p n \<Longrightarrow> isnpolyh (polypow k p) n"
proof (induct k arbitrary: n rule: polypow.induct)
case (2 k n)
@@ -701,17 +701,17 @@
qed auto
lemma polypow_norm:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
by (simp add: polypow_normh isnpoly_def)
text{* Finally the whole normalization*}
-lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{ring_char_0,division_by_zero,field})"
+lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field, division_ring_inverse_zero, ring_char_0})"
by (induct p rule:polynate.induct, auto)
lemma polynate_norm[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "isnpoly (polynate p)"
by (induct p rule: polynate.induct, simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm) (simp_all add: isnpoly_def)
@@ -736,29 +736,29 @@
shows "isnpolyh (funpow k f p) n"
using f np by (induct k arbitrary: p, auto)
-lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
+lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {field, division_ring_inverse_zero, ring_char_0}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
by (induct n arbitrary: p, simp_all add: shift1_isnpoly shift1 power_Suc )
lemma shift1_isnpolyh: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> isnpolyh (shift1 p) 0"
using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by (simp add: shift1_def)
lemma funpow_shift1_1:
- "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
+ "(Ipoly bs (funpow n shift1 p) :: 'a :: {field, division_ring_inverse_zero, ring_char_0}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
by (simp add: funpow_shift1)
lemma poly_cmul[simp]: "Ipoly bs (poly_cmul c p) = Ipoly bs (Mul (C c) p)"
-by (induct p arbitrary: n0 rule: poly_cmul.induct, auto simp add: ring_simps)
+by (induct p arbitrary: n0 rule: poly_cmul.induct, auto simp add: field_simps)
lemma behead:
assumes np: "isnpolyh p n"
- shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {ring_char_0,division_by_zero,field})"
+ shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {field, division_ring_inverse_zero, ring_char_0})"
using np
proof (induct p arbitrary: n rule: behead.induct)
case (1 c p n) hence pn: "isnpolyh p n" by simp
from prems(2)[OF pn]
have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" .
then show ?case using "1.hyps" apply (simp add: Let_def,cases "behead p = 0\<^sub>p")
- by (simp_all add: th[symmetric] ring_simps power_Suc)
+ by (simp_all add: th[symmetric] field_simps power_Suc)
qed (auto simp add: Let_def)
lemma behead_isnpolyh:
@@ -981,7 +981,7 @@
by (simp add: head_eq_headn0)
lemma isnpolyh_zero_iff:
- assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_by_zero,field})"
+ assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_ring_inverse_zero,field})"
shows "p = 0\<^sub>p"
using nq eq
proof (induct "maxindex p" arbitrary: p n0 rule: less_induct)
@@ -1033,7 +1033,7 @@
lemma isnpolyh_unique:
assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1"
- shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_by_zero,field})) \<longleftrightarrow> p = q"
+ shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_ring_inverse_zero,field})) \<longleftrightarrow> p = q"
proof(auto)
assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a)= \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
hence "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp
@@ -1046,50 +1046,50 @@
text{* consequenses of unicity on the algorithms for polynomial normalization *}
-lemma polyadd_commute: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma polyadd_commute: assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p"
using isnpolyh_unique[OF polyadd_normh[OF np nq] polyadd_normh[OF nq np]] by simp
lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp
lemma one_normh: "isnpolyh 1\<^sub>p n" by simp
lemma polyadd_0[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" shows "p +\<^sub>p 0\<^sub>p = p" and "0\<^sub>p +\<^sub>p p = p"
using isnpolyh_unique[OF polyadd_normh[OF np zero_normh] np]
isnpolyh_unique[OF polyadd_normh[OF zero_normh np] np] by simp_all
lemma polymul_1[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" shows "p *\<^sub>p 1\<^sub>p = p" and "1\<^sub>p *\<^sub>p p = p"
using isnpolyh_unique[OF polymul_normh[OF np one_normh] np]
isnpolyh_unique[OF polymul_normh[OF one_normh np] np] by simp_all
lemma polymul_0[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" shows "p *\<^sub>p 0\<^sub>p = 0\<^sub>p" and "0\<^sub>p *\<^sub>p p = 0\<^sub>p"
using isnpolyh_unique[OF polymul_normh[OF np zero_normh] zero_normh]
isnpolyh_unique[OF polymul_normh[OF zero_normh np] zero_normh] by simp_all
lemma polymul_commute:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np:"isnpolyh p n0" and nq: "isnpolyh q n1"
shows "p *\<^sub>p q = q *\<^sub>p p"
-using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_by_zero,field}"] by simp
+using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_ring_inverse_zero,field}"] by simp
declare polyneg_polyneg[simp]
lemma isnpolyh_polynate_id[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np:"isnpolyh p n0" shows "polynate p = p"
- using isnpolyh_unique[where ?'a= "'a::{ring_char_0,division_by_zero,field}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{ring_char_0,division_by_zero,field}"] by simp
+ using isnpolyh_unique[where ?'a= "'a::{field, division_ring_inverse_zero, ring_char_0}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}"] by simp
lemma polynate_idempotent[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "polynate (polynate p) = polynate p"
using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] .
lemma poly_nate_polypoly': "poly_nate bs p = polypoly' bs (polynate p)"
unfolding poly_nate_def polypoly'_def ..
-lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{ring_char_0,division_by_zero,field}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
+lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{field, division_ring_inverse_zero, ring_char_0}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
using polypoly'_poly[OF polynate_norm[unfolded isnpoly_def], symmetric, of bs p]
unfolding poly_nate_polypoly' by (auto intro: ext)
@@ -1116,7 +1116,7 @@
qed
lemma degree_polysub_samehead:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q"
and d: "degree p = degree q"
shows "degree (p -\<^sub>p q) < degree p \<or> (p -\<^sub>p q = 0\<^sub>p)"
@@ -1226,7 +1226,7 @@
done
lemma polymul_head_polyeq:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "\<lbrakk>isnpolyh p n0; isnpolyh q n1 ; p \<noteq> 0\<^sub>p ; q \<noteq> 0\<^sub>p \<rbrakk> \<Longrightarrow> head (p *\<^sub>p q) = head p *\<^sub>p head q"
proof (induct p q arbitrary: n0 n1 rule: polymul.induct)
case (2 a b c' n' p' n0 n1)
@@ -1300,7 +1300,7 @@
by (induct p arbitrary: n0 rule: polyneg.induct, auto)
lemma degree_polymul:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and nq: "isnpolyh q n1"
shows "degree (p *\<^sub>p q) \<le> degree p + degree q"
using polymul_degreen[OF np nq, where m="0"] degree_eq_degreen0 by simp
@@ -1344,7 +1344,7 @@
qed
lemma polydivide_aux_properties:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and ns: "isnpolyh s n1"
and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \<noteq> 0\<^sub>p"
shows "polydivide_aux_dom (a,n,p,k,s) \<and>
@@ -1415,19 +1415,19 @@
from polyadd_normh[OF polymul_normh[OF np
polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp
- from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) =
+ from asp have "\<forall> (bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) =
Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
- hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) =
+ hence " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) =
Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r"
- by (simp add: ring_simps)
- hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ by (simp add: field_simps)
+ hence " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p *\<^sub>p p)
+ Ipoly bs p * Ipoly bs q + Ipoly bs r"
by (auto simp only: funpow_shift1_1)
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p)
- + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ + Ipoly bs q) + Ipoly bs r" by (simp add: field_simps)
+ hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
with isnpolyh_unique[OF nakks' nqr']
have "a ^\<^sub>p (k' - k) *\<^sub>p s =
@@ -1444,9 +1444,9 @@
apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
have dom: ?dths apply (rule polydivide_aux_real_domintros)
using ba dn' domsp by simp_all
- from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
- have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
- hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
+ from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}"]
+ have " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs s = Ipoly bs ?p'" by simp
+ hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
by (simp only: funpow_shift1_1) simp
hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
{assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
@@ -1501,17 +1501,17 @@
and dr: "degree r = 0 \<or> degree r < degree p"
and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
- {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
+ {fix bs:: "'a::{field, division_ring_inverse_zero, ring_char_0} list"
from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
- by (simp add: ring_simps power_Suc)
+ by (simp add: field_simps power_Suc)
hence "Ipoly bs a ^ (k' - k) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
by (simp add:kk'' funpow_shift1_1[where n="degree s - n" and p="p"])
hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
- by (simp add: ring_simps)}
- hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ by (simp add: field_simps)}
+ hence ieq:"\<forall>(bs :: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto
let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
@@ -1532,17 +1532,17 @@
apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
have dom: ?dths using sz ba dn' domsp
by - (rule polydivide_aux_real_domintros, simp_all)
- {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
+ {fix bs :: "'a::{field, division_ring_inverse_zero, ring_char_0} list"
from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p"
by (simp add: funpow_shift1_1[where n="degree s - n" and p="p"])
hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
}
- hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+ hence hth: "\<forall> (bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
from hth
have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)"
- using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
+ using isnpolyh_unique[where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
simplified ap] by simp
{assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
@@ -1566,7 +1566,7 @@
qed
lemma polydivide_properties:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \<noteq> 0\<^sub>p"
shows "(\<exists> k r. polydivide s p = (k,r) \<and> (\<exists>nr. isnpolyh r nr) \<and> (degree r = 0 \<or> degree r < degree p)
\<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow k (head p)) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
@@ -1698,11 +1698,11 @@
definition "swapnorm n m t = polynate (swap n m t)"
lemma swapnorm: assumes nbs: "n < length bs" and mbs: "m < length bs"
- shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{ring_char_0,division_by_zero,field})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+ shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{field, division_ring_inverse_zero, ring_char_0})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
using swap[OF prems] swapnorm_def by simp
lemma swapnorm_isnpoly[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
shows "isnpoly (swapnorm n m p)"
unfolding swapnorm_def by simp
--- a/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy Mon Apr 26 09:45:22 2010 -0700
@@ -7,147 +7,147 @@
begin
lemma
- "\<exists>(y::'a::{linordered_field,number_ring, division_by_zero}) <2. x + 3* y < 0 \<and> x - y >0"
+ "\<exists>(y::'a::{linordered_field, division_ring_inverse_zero, number_ring}) <2. x + 3* y < 0 \<and> x - y >0"
by ferrack
-lemma "~ (ALL x (y::'a::{linordered_field,number_ring, division_by_zero}). x < y --> 10*x < 11*y)"
+lemma "~ (ALL x (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). x < y --> 10*x < 11*y)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (10*(x + 5*y + -1) < 60*y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (10*(x + 5*y + -1) < 60*y)"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. x ~= y --> x < y"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x ~= y --> x < y"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX (y::'a::{linordered_field,number_ring, division_by_zero}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) < 0. (EX (y::'a::{linordered_field,number_ring, division_by_zero}) > 0. 7*x + y > 0 & x - y <= 9)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) < 0. (EX (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}) > 0. 7*x + y > 0 & x - y <= 9)"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)"
by ferrack
-lemma "EX x. (ALL (y::'a::{linordered_field,number_ring, division_by_zero}). y < 2 --> 2*(y - x) \<le> 0 )"
+lemma "EX x. (ALL (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y < 2 --> 2*(y - x) \<le> 0 )"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + y < z --> y >= z --> x < 0"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + y < z --> y >= z --> x < 0"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. abs (x + y) <= z --> (abs z = z)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. abs (x + y) <= z --> (abs z = z)"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (EX z>0. x+z = y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (EX z>0. x+z = y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z>0. abs (x - y) <= z )"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z>0. abs (x - y) <= z )"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero})>0. (ALL y. (EX z. 13* abs z \<noteq> abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring})>0. (ALL y. (EX z. 13* abs z \<noteq> abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \<noteq> 0 \<longrightarrow> (EX z. 5*x - 3*abs y <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \<noteq> 0 \<longrightarrow> (EX z. 5*x - 3*abs y <= 7*z))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)"
by ferrack
-lemma "~(ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))"
+lemma "~(ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)"
by ferrack
-lemma "EX z. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))"
+lemma "EX z. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))"
by ferrack
-lemma "EX z. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))"
+lemma "EX z. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))"
by ferrack
-lemma "EX y. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))"
+lemma "EX y. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (ALL y < x. (EX z > (x+y).
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (ALL y < x. (EX z > (x+y).
(ALL w. 5*w + 10*x - z >= y --> w + 7*x + 3*z >= 2*y)))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (ALL y. (EX z > y.
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (ALL y. (EX z > y.
(ALL w . w < 13 --> w + 10*x - z >= y --> 5*w + 7*x + 13*z >= 2*y)))"
by ferrack
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))"
by ferrack
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))"
by ferrack
end
--- a/src/HOL/Fields.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Fields.thy Mon Apr 26 09:45:22 2010 -0700
@@ -96,51 +96,45 @@
"\<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 add_divide_eq_iff:
+lemma add_divide_eq_iff [field_simps]:
"z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
by (simp add: add_divide_distrib)
-lemma divide_add_eq_iff:
+lemma divide_add_eq_iff [field_simps]:
"z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
by (simp add: add_divide_distrib)
-lemma diff_divide_eq_iff:
+lemma diff_divide_eq_iff [field_simps]:
"z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
by (simp add: diff_divide_distrib)
-lemma divide_diff_eq_iff:
+lemma divide_diff_eq_iff [field_simps]:
"z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
by (simp add: diff_divide_distrib)
-lemmas field_eq_simps[no_atp] = 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
- 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)
+ by (simp add: field_simps)
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)
+ by (simp add: field_simps)
+
+end
+
+class field_inverse_zero = field +
+ assumes field_inverse_zero: "inverse 0 = 0"
+begin
+
+subclass division_ring_inverse_zero proof
+qed (fact field_inverse_zero)
end
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})"
+ "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_ring_inverse_zero})"
proof cases
assume "a \<noteq> 0 & b \<noteq> 0"
thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_ac)
@@ -150,7 +144,7 @@
qed
lemma inverse_divide [simp]:
- "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
+ "inverse (a/b) = b / (a::'a::{field,division_ring_inverse_zero})"
by (simp add: divide_inverse mult_commute)
@@ -161,85 +155,85 @@
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})"
+ "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_ring_inverse_zero})"
apply (cases "b = 0")
apply simp_all
done
lemma mult_divide_mult_cancel_right:
- "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
+ "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_ring_inverse_zero})"
apply (cases "b = 0")
apply simp_all
done
lemma divide_divide_eq_right [simp,no_atp]:
- "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
+ "a / (b/c) = (a*c) / (b::'a::{field,division_ring_inverse_zero})"
by (simp add: divide_inverse mult_ac)
lemma divide_divide_eq_left [simp,no_atp]:
- "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
+ "(a / b) / (c::'a::{field,division_ring_inverse_zero}) = a / (b*c)"
by (simp add: divide_inverse mult_assoc)
text {*Special Cancellation Simprules for Division*}
lemma mult_divide_mult_cancel_left_if[simp,no_atp]:
-fixes c :: "'a :: {field,division_by_zero}"
+fixes c :: "'a :: {field,division_ring_inverse_zero}"
shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
by (simp add: mult_divide_mult_cancel_left)
text {* Division and Unary Minus *}
-lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
+lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_ring_inverse_zero})"
by (simp add: divide_inverse)
lemma divide_minus_right [simp, no_atp]:
- "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
+ "a / -(b::'a::{field,division_ring_inverse_zero}) = -(a / b)"
by (simp add: divide_inverse)
lemma minus_divide_divide:
- "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
+ "(-a)/(-b) = a / (b::'a::{field,division_ring_inverse_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)"
+ "((a::'a::{field,division_ring_inverse_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)"
+ "(b/c = (a::'a::{field,division_ring_inverse_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
by (force simp add: nonzero_divide_eq_eq)
lemma inverse_eq_1_iff [simp]:
- "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
+ "(inverse x = 1) = (x = (1::'a::{field,division_ring_inverse_zero}))"
by (insert inverse_eq_iff_eq [of x 1], simp)
lemma divide_eq_0_iff [simp,no_atp]:
- "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
+ "(a/b = 0) = (a=0 | b=(0::'a::{field,division_ring_inverse_zero}))"
by (simp add: divide_inverse)
lemma divide_cancel_right [simp,no_atp]:
- "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
+ "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (simp add: divide_inverse)
done
lemma divide_cancel_left [simp,no_atp]:
- "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
+ "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (simp add: divide_inverse)
done
lemma divide_eq_1_iff [simp,no_atp]:
- "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
+ "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_ring_inverse_zero}))"
apply (cases "b=0", simp)
apply (simp add: right_inverse_eq)
done
lemma one_eq_divide_iff [simp,no_atp]:
- "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
+ "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_ring_inverse_zero}))"
by (simp add: eq_commute [of 1])
@@ -391,7 +385,7 @@
"a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg)
-lemma pos_le_divide_eq: "0 < c ==> (a \<le> b/c) = (a*c \<le> b)"
+lemma pos_le_divide_eq [field_simps]: "0 < c ==> (a \<le> b/c) = (a*c \<le> b)"
proof -
assume less: "0<c"
hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
@@ -401,7 +395,7 @@
finally show ?thesis .
qed
-lemma neg_le_divide_eq: "c < 0 ==> (a \<le> b/c) = (b \<le> a*c)"
+lemma neg_le_divide_eq [field_simps]: "c < 0 ==> (a \<le> b/c) = (b \<le> a*c)"
proof -
assume less: "c<0"
hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
@@ -411,7 +405,7 @@
finally show ?thesis .
qed
-lemma pos_less_divide_eq:
+lemma pos_less_divide_eq [field_simps]:
"0 < c ==> (a < b/c) = (a*c < b)"
proof -
assume less: "0<c"
@@ -422,7 +416,7 @@
finally show ?thesis .
qed
-lemma neg_less_divide_eq:
+lemma neg_less_divide_eq [field_simps]:
"c < 0 ==> (a < b/c) = (b < a*c)"
proof -
assume less: "c<0"
@@ -433,7 +427,7 @@
finally show ?thesis .
qed
-lemma pos_divide_less_eq:
+lemma pos_divide_less_eq [field_simps]:
"0 < c ==> (b/c < a) = (b < a*c)"
proof -
assume less: "0<c"
@@ -444,7 +438,7 @@
finally show ?thesis .
qed
-lemma neg_divide_less_eq:
+lemma neg_divide_less_eq [field_simps]:
"c < 0 ==> (b/c < a) = (a*c < b)"
proof -
assume less: "c<0"
@@ -455,7 +449,7 @@
finally show ?thesis .
qed
-lemma pos_divide_le_eq: "0 < c ==> (b/c \<le> a) = (b \<le> a*c)"
+lemma pos_divide_le_eq [field_simps]: "0 < c ==> (b/c \<le> a) = (b \<le> a*c)"
proof -
assume less: "0<c"
hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
@@ -465,7 +459,7 @@
finally show ?thesis .
qed
-lemma neg_divide_le_eq: "c < 0 ==> (b/c \<le> a) = (a*c \<le> b)"
+lemma neg_divide_le_eq [field_simps]: "c < 0 ==> (b/c \<le> a) = (a*c \<le> b)"
proof -
assume less: "c<0"
hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
@@ -475,24 +469,15 @@
finally show ?thesis .
qed
-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[no_atp] = 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[no_atp] = group_simps
+lemmas sign_simps [no_atp] = algebra_simps
+ zero_less_mult_iff mult_less_0_iff
+
+lemmas (in -) sign_simps [no_atp] = algebra_simps
zero_less_mult_iff mult_less_0_iff
(* Only works once linear arithmetic is installed:
@@ -658,37 +643,46 @@
end
+class linordered_field_inverse_zero = linordered_field +
+ assumes linordered_field_inverse_zero: "inverse 0 = 0"
+begin
+
+subclass field_inverse_zero proof
+qed (fact linordered_field_inverse_zero)
+
+end
+
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}))"
+ else a \<le> (0::'a::{linordered_field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff)
done
lemma inverse_positive_iff_positive [simp]:
- "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))"
+ "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_ring_inverse_zero}))"
apply (cases "a = 0", simp)
apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
done
lemma inverse_negative_iff_negative [simp]:
- "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))"
+ "(inverse a < 0) = (a < (0::'a::{linordered_field,division_ring_inverse_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}))"
+ "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_ring_inverse_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}))"
+ "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_ring_inverse_zero}))"
by (simp add: linorder_not_less [symmetric])
lemma one_less_inverse_iff:
- "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))"
+ "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_ring_inverse_zero}))"
proof cases
assume "0 < x"
with inverse_less_iff_less [OF zero_less_one, of x]
@@ -706,22 +700,22 @@
qed
lemma one_le_inverse_iff:
- "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_by_zero}))"
+ "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_ring_inverse_zero}))"
by (force simp add: order_le_less one_less_inverse_iff)
lemma inverse_less_1_iff:
- "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_by_zero}))"
+ "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_ring_inverse_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}))"
+ "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_ring_inverse_zero}))"
by (simp add: linorder_not_less [symmetric] one_less_inverse_iff)
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}))"
+ else 0 \<le> (a::'a::{linordered_field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff)
done
@@ -730,7 +724,7 @@
"(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}))"
+ else a < (0::'a::{linordered_field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff)
done
@@ -739,7 +733,7 @@
"(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}))"
+ else 0 < (a::'a::{linordered_field,division_ring_inverse_zero}))"
apply (cases "c=0", simp)
apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff)
done
@@ -747,21 +741,21 @@
text {*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)"
+ "((0::'a::{linordered_field,division_ring_inverse_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})) =
+ "(a/b < (0::'a::{linordered_field,division_ring_inverse_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::'a::{linordered_field,division_ring_inverse_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})) =
+ "(a/b \<le> (0::'a::{linordered_field,division_ring_inverse_zero})) =
(0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
by (simp add: divide_inverse mult_le_0_iff)
@@ -770,13 +764,13 @@
text{*Simplify expressions equated with 1*}
lemma zero_eq_1_divide_iff [simp,no_atp]:
- "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)"
+ "((0::'a::{linordered_field,division_ring_inverse_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,no_atp]:
- "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)"
+ "(1/a = (0::'a::{linordered_field,division_ring_inverse_zero})) = (a = 0)"
apply (cases "a=0", simp)
apply (insert zero_neq_one [THEN not_sym])
apply (auto simp add: nonzero_divide_eq_eq)
@@ -794,16 +788,16 @@
declare divide_le_0_1_iff [simp,no_atp]
lemma divide_right_mono:
- "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_by_zero})"
+ "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_ring_inverse_zero})"
by (force simp add: divide_strict_right_mono order_le_less)
-lemma divide_right_mono_neg: "(a::'a::{linordered_field,division_by_zero}) <= b
+lemma divide_right_mono_neg: "(a::'a::{linordered_field,division_ring_inverse_zero}) <= b
==> c <= 0 ==> b / c <= a / c"
apply (drule divide_right_mono [of _ _ "- c"])
apply auto
done
-lemma divide_left_mono_neg: "(a::'a::{linordered_field,division_by_zero}) <= b
+lemma divide_left_mono_neg: "(a::'a::{linordered_field,division_ring_inverse_zero}) <= b
==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
apply (drule divide_left_mono [of _ _ "- c"])
apply (auto simp add: mult_commute)
@@ -814,22 +808,22 @@
text{*Simplify quotients that are compared with the value 1.*}
lemma le_divide_eq_1 [no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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 [no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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 [no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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 [no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
by (auto simp add: divide_less_eq)
@@ -837,76 +831,76 @@
text {*Conditional Simplification Rules: No Case Splits*}
lemma le_divide_eq_1_pos [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_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,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
by (auto simp add: less_divide_eq)
lemma less_divide_eq_1_neg [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
by (auto simp add: less_divide_eq)
lemma divide_less_eq_1_pos [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
by (auto simp add: divide_less_eq)
lemma divide_less_eq_1_neg [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
by (auto simp add: divide_less_eq)
lemma eq_divide_eq_1 [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "(1 = b/a) = ((a \<noteq> 0 & a = b))"
by (auto simp add: eq_divide_eq)
lemma divide_eq_eq_1 [simp,no_atp]:
- fixes a :: "'a :: {linordered_field,division_by_zero}"
+ fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
shows "(b/a = 1) = ((a \<noteq> 0 & a = b))"
by (auto simp add: divide_eq_eq)
lemma abs_inverse [simp]:
- "\<bar>inverse (a::'a::{linordered_field,division_by_zero})\<bar> =
+ "\<bar>inverse (a::'a::{linordered_field,division_ring_inverse_zero})\<bar> =
inverse \<bar>a\<bar>"
apply (cases "a=0", simp)
apply (simp add: nonzero_abs_inverse)
done
lemma abs_divide [simp]:
- "\<bar>a / (b::'a::{linordered_field,division_by_zero})\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
+ "\<bar>a / (b::'a::{linordered_field,division_ring_inverse_zero})\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
apply (cases "b=0", simp)
apply (simp add: nonzero_abs_divide)
done
-lemma abs_div_pos: "(0::'a::{linordered_field,division_by_zero}) < y ==>
+lemma abs_div_pos: "(0::'a::{linordered_field,division_ring_inverse_zero}) < y ==>
\<bar>x\<bar> / y = \<bar>x / y\<bar>"
apply (subst abs_divide)
apply (simp add: order_less_imp_le)
done
lemma field_le_mult_one_interval:
- fixes x :: "'a\<Colon>{linordered_field,division_by_zero}"
+ fixes x :: "'a\<Colon>{linordered_field,division_ring_inverse_zero}"
assumes *: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
shows "x \<le> y"
proof (cases "0 < x")
--- a/src/HOL/GCD.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/GCD.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1034,11 +1034,11 @@
thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
apply (simp add: bezw_non_0 gcd_non_0_nat)
apply (erule subst)
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
apply (subst mod_div_equality [of m n, symmetric])
(* applying simp here undoes the last substitution!
what is procedure cancel_div_mod? *)
- apply (simp only: ring_simps zadd_int [symmetric]
+ apply (simp only: field_simps zadd_int [symmetric]
zmult_int [symmetric])
done
qed
@@ -1389,7 +1389,7 @@
show "lcm (lcm n m) p = lcm n (lcm m p)"
by (rule lcm_unique_nat [THEN iffD1]) (metis dvd.order_trans lcm_unique_nat)
show "lcm m n = lcm n m"
- by (simp add: lcm_nat_def gcd_commute_nat ring_simps)
+ by (simp add: lcm_nat_def gcd_commute_nat field_simps)
qed
interpretation lcm_int!: abel_semigroup "lcm :: int \<Rightarrow> int \<Rightarrow> int"
--- a/src/HOL/Groebner_Basis.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Groebner_Basis.thy Mon Apr 26 09:45:22 2010 -0700
@@ -474,20 +474,20 @@
fieldgb "op +" "op *" "op ^" "0::'a::{field,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse)
lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp
-lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0"
+lemma divide_Numeral0: "(x::'a::{field,number_ring, division_ring_inverse_zero}) / Numeral0 = 0"
by simp
-lemma mult_frac_frac: "((x::'a::{field,division_by_zero}) / y) * (z / w) = (x*z) / (y*w)"
+lemma mult_frac_frac: "((x::'a::{field,division_ring_inverse_zero}) / y) * (z / w) = (x*z) / (y*w)"
by simp
-lemma mult_frac_num: "((x::'a::{field, division_by_zero}) / y) * z = (x*z) / y"
+lemma mult_frac_num: "((x::'a::{field, division_ring_inverse_zero}) / y) * z = (x*z) / y"
by simp
-lemma mult_num_frac: "((x::'a::{field, division_by_zero}) / y) * z = (x*z) / y"
+lemma mult_num_frac: "((x::'a::{field, division_ring_inverse_zero}) / y) * z = (x*z) / y"
by simp
lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp
-lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_by_zero}) / y + z = (x + z*y) / y"
+lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_ring_inverse_zero}) / y + z = (x + z*y) / y"
by (simp add: add_divide_distrib)
-lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_by_zero}) / y = (x + z*y) / y"
+lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_ring_inverse_zero}) / y = (x + z*y) / y"
by (simp add: add_divide_distrib)
ML {*
--- a/src/HOL/Groups.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Groups.thy Mon Apr 26 09:45:22 2010 -0700
@@ -12,13 +12,13 @@
subsection {* Fact collections *}
ML {*
-structure Algebra_Simps = Named_Thms(
- val name = "algebra_simps"
- val description = "algebra simplification rules"
+structure Ac_Simps = Named_Thms(
+ val name = "ac_simps"
+ val description = "associativity and commutativity simplification rules"
)
*}
-setup Algebra_Simps.setup
+setup Ac_Simps.setup
text{* The rewrites accumulated in @{text algebra_simps} deal with the
classical algebraic structures of groups, rings and family. They simplify
@@ -29,15 +29,28 @@
Of course it also works for fields, but it knows nothing about multiplicative
inverses or division. This is catered for by @{text field_simps}. *}
-
ML {*
-structure Ac_Simps = Named_Thms(
- val name = "ac_simps"
- val description = "associativity and commutativity simplification rules"
+structure Algebra_Simps = Named_Thms(
+ val name = "algebra_simps"
+ val description = "algebra simplification rules"
)
*}
-setup Ac_Simps.setup
+setup Algebra_Simps.setup
+
+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}. *}
+
+ML {*
+structure Field_Simps = Named_Thms(
+ val name = "field_simps"
+ val description = "algebra simplification rules for fields"
+)
+*}
+
+setup Field_Simps.setup
subsection {* Abstract structures *}
@@ -139,13 +152,13 @@
subsection {* Semigroups and Monoids *}
class semigroup_add = plus +
- assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
+ assumes add_assoc [algebra_simps, field_simps]: "(a + b) + c = a + (b + c)"
sublocale semigroup_add < add!: semigroup plus proof
qed (fact add_assoc)
class ab_semigroup_add = semigroup_add +
- assumes add_commute [algebra_simps]: "a + b = b + a"
+ assumes add_commute [algebra_simps, field_simps]: "a + b = b + a"
sublocale ab_semigroup_add < add!: abel_semigroup plus proof
qed (fact add_commute)
@@ -153,7 +166,7 @@
context ab_semigroup_add
begin
-lemmas add_left_commute [algebra_simps] = add.left_commute
+lemmas add_left_commute [algebra_simps, field_simps] = add.left_commute
theorems add_ac = add_assoc add_commute add_left_commute
@@ -162,13 +175,13 @@
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)"
+ assumes mult_assoc [algebra_simps, field_simps]: "(a * b) * c = a * (b * c)"
sublocale semigroup_mult < mult!: semigroup times proof
qed (fact mult_assoc)
class ab_semigroup_mult = semigroup_mult +
- assumes mult_commute [algebra_simps]: "a * b = b * a"
+ assumes mult_commute [algebra_simps, field_simps]: "a * b = b * a"
sublocale ab_semigroup_mult < mult!: abel_semigroup times proof
qed (fact mult_commute)
@@ -176,7 +189,7 @@
context ab_semigroup_mult
begin
-lemmas mult_left_commute [algebra_simps] = mult.left_commute
+lemmas mult_left_commute [algebra_simps, field_simps] = mult.left_commute
theorems mult_ac = mult_assoc mult_commute mult_left_commute
@@ -371,7 +384,7 @@
lemma add_diff_cancel: "a + b - b = a"
by (simp add: diff_minus add_assoc)
-declare diff_minus[symmetric, algebra_simps]
+declare diff_minus[symmetric, algebra_simps, field_simps]
lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
proof
@@ -402,7 +415,7 @@
then show "b = c" by simp
qed
-lemma uminus_add_conv_diff[algebra_simps]:
+lemma uminus_add_conv_diff[algebra_simps, field_simps]:
"- a + b = b - a"
by (simp add:diff_minus add_commute)
@@ -414,22 +427,22 @@
"- (a - b) = b - a"
by (simp add: diff_minus add_commute)
-lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c"
+lemma add_diff_eq[algebra_simps, field_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"
+lemma diff_add_eq[algebra_simps, field_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"
+lemma diff_eq_eq[algebra_simps, field_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"
+lemma eq_diff_eq[algebra_simps, field_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)"
+lemma diff_diff_eq[algebra_simps, field_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"
+lemma diff_diff_eq2[algebra_simps, field_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"
@@ -750,35 +763,29 @@
finally show ?thesis .
qed
-lemma diff_less_eq[algebra_simps]: "a - b < c \<longleftrightarrow> a < c + b"
+lemma diff_less_eq[algebra_simps, field_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"
+lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"
apply (subst less_iff_diff_less_0 [of "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"
+lemma diff_le_eq[algebra_simps, field_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"
+lemma le_diff_eq[algebra_simps, field_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[no_atp] = algebra_simps
-
end
-text{*Legacy - use @{text algebra_simps} *}
-lemmas group_simps[no_atp] = algebra_simps
-
class linordered_ab_semigroup_add =
linorder + ordered_ab_semigroup_add
--- a/src/HOL/HOL.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/HOL.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1869,7 +1869,7 @@
proof
assume "PROP ?ofclass"
show "PROP ?eq"
- by (tactic {* ALLGOALS (rtac (Drule.unconstrainTs @{thm equals_eq})) *})
+ by (tactic {* ALLGOALS (rtac (Thm.unconstrain_allTs @{thm equals_eq})) *})
(fact `PROP ?ofclass`)
next
assume "PROP ?eq"
--- a/src/HOL/Import/HOL/real.imp Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Import/HOL/real.imp Mon Apr 26 09:45:22 2010 -0700
@@ -251,7 +251,7 @@
"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_INV_0" > "Rings.division_ring_inverse_zero_class.inverse_zero"
"REAL_INVINV" > "Rings.nonzero_inverse_inverse_eq"
"REAL_INV1" > "Rings.inverse_1"
"REAL_INJ" > "RealDef.real_of_nat_inject"
--- a/src/HOL/Import/HOL/realax.imp Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Import/HOL/realax.imp Mon Apr 26 09:45:22 2010 -0700
@@ -101,7 +101,7 @@
"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_INV_0" > "Rings.division_ring_inverse_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 Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Int.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1995,15 +1995,15 @@
text{*Division By @{text "-1"}*}
lemma divide_minus1 [simp]:
- "x/-1 = -(x::'a::{field,division_by_zero,number_ring})"
+ "x/-1 = -(x::'a::{field,division_ring_inverse_zero,number_ring})"
by simp
lemma minus1_divide [simp]:
- "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
+ "-1 / (x::'a::{field,division_ring_inverse_zero,number_ring}) = - (1/x)"
by (simp add: divide_inverse)
lemma half_gt_zero_iff:
- "(0 < r/2) = (0 < (r::'a::{linordered_field,division_by_zero,number_ring}))"
+ "(0 < r/2) = (0 < (r::'a::{linordered_field,division_ring_inverse_zero,number_ring}))"
by auto
lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2, standard]
--- a/src/HOL/Lattices.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Lattices.thy Mon Apr 26 09:45:22 2010 -0700
@@ -365,13 +365,9 @@
subsection {* Bounded lattices and boolean algebras *}
-class bounded_lattice = lattice + top + bot
+class bounded_lattice_bot = lattice + bot
begin
-lemma dual_bounded_lattice:
- "bounded_lattice (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>"
- by unfold_locales (auto simp add: less_le_not_le)
-
lemma inf_bot_left [simp]:
"\<bottom> \<sqinter> x = \<bottom>"
by (rule inf_absorb1) simp
@@ -380,6 +376,23 @@
"x \<sqinter> \<bottom> = \<bottom>"
by (rule inf_absorb2) simp
+lemma sup_bot_left [simp]:
+ "\<bottom> \<squnion> x = x"
+ by (rule sup_absorb2) simp
+
+lemma sup_bot_right [simp]:
+ "x \<squnion> \<bottom> = x"
+ by (rule sup_absorb1) simp
+
+lemma sup_eq_bot_iff [simp]:
+ "x \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
+ by (simp add: eq_iff)
+
+end
+
+class bounded_lattice_top = lattice + top
+begin
+
lemma sup_top_left [simp]:
"\<top> \<squnion> x = \<top>"
by (rule sup_absorb1) simp
@@ -396,21 +409,18 @@
"x \<sqinter> \<top> = x"
by (rule inf_absorb1) simp
-lemma sup_bot_left [simp]:
- "\<bottom> \<squnion> x = x"
- by (rule sup_absorb2) simp
-
-lemma sup_bot_right [simp]:
- "x \<squnion> \<bottom> = x"
- by (rule sup_absorb1) simp
-
lemma inf_eq_top_iff [simp]:
"x \<sqinter> y = \<top> \<longleftrightarrow> x = \<top> \<and> y = \<top>"
by (simp add: eq_iff)
-lemma sup_eq_bot_iff [simp]:
- "x \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
- by (simp add: eq_iff)
+end
+
+class bounded_lattice = bounded_lattice_bot + bounded_lattice_top
+begin
+
+lemma dual_bounded_lattice:
+ "bounded_lattice (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>"
+ by unfold_locales (auto simp add: less_le_not_le)
end
--- a/src/HOL/Library/Abstract_Rat.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Abstract_Rat.thy Mon Apr 26 09:45:22 2010 -0700
@@ -184,7 +184,7 @@
lemma isnormNum_unique[simp]:
assumes na: "isnormNum x" and nb: "isnormNum y"
- shows "((INum x ::'a::{ring_char_0,field, division_by_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
+ shows "((INum x ::'a::{ring_char_0,field, division_ring_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
proof
have "\<exists> a b a' b'. x = (a,b) \<and> y = (a',b')" by auto
then obtain a b a' b' where xy[simp]: "x = (a,b)" "y=(a',b')" by blast
@@ -217,7 +217,7 @@
qed
-lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> (INum x = (0::'a::{ring_char_0, field,division_by_zero})) = (x = 0\<^sub>N)"
+lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> (INum x = (0::'a::{ring_char_0, field,division_ring_inverse_zero})) = (x = 0\<^sub>N)"
unfolding INum_int(2)[symmetric]
by (rule isnormNum_unique, simp_all)
@@ -245,7 +245,7 @@
done
-lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{ring_char_0,field, division_by_zero})"
+lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{ring_char_0,field, division_ring_inverse_zero})"
proof-
have "\<exists> a b. x = (a,b)" by auto
then obtain a b where x[simp]: "x = (a,b)" by blast
@@ -260,7 +260,7 @@
ultimately show ?thesis by blast
qed
-lemma INum_normNum_iff: "(INum x ::'a::{field, division_by_zero, ring_char_0}) = INum y \<longleftrightarrow> normNum x = normNum y" (is "?lhs = ?rhs")
+lemma INum_normNum_iff: "(INum x ::'a::{field, division_ring_inverse_zero, ring_char_0}) = INum y \<longleftrightarrow> normNum x = normNum y" (is "?lhs = ?rhs")
proof -
have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
by (simp del: normNum)
@@ -268,7 +268,7 @@
finally show ?thesis by simp
qed
-lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {ring_char_0,division_by_zero,field})"
+lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {ring_char_0,division_ring_inverse_zero,field})"
proof-
let ?z = "0:: 'a"
have " \<exists> a b. x = (a,b)" " \<exists> a' b'. y = (a',b')" by auto
@@ -300,7 +300,7 @@
ultimately show ?thesis by blast
qed
-lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {ring_char_0,division_by_zero,field}) "
+lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {ring_char_0,division_ring_inverse_zero,field}) "
proof-
let ?z = "0::'a"
have " \<exists> a b. x = (a,b)" " \<exists> a' b'. y = (a',b')" by auto
@@ -323,16 +323,16 @@
lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
by (simp add: Nneg_def split_def INum_def)
-lemma Nsub[simp]: shows "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {ring_char_0,division_by_zero,field})"
+lemma Nsub[simp]: shows "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {ring_char_0,division_ring_inverse_zero,field})"
by (simp add: Nsub_def split_def)
-lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: {division_by_zero,field}) / (INum x)"
+lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: {division_ring_inverse_zero,field}) / (INum x)"
by (simp add: Ninv_def INum_def split_def)
-lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {ring_char_0, division_by_zero,field})" by (simp add: Ndiv_def)
+lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {ring_char_0, division_ring_inverse_zero,field})" by (simp add: Ndiv_def)
lemma Nlt0_iff[simp]: assumes nx: "isnormNum x"
- shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})< 0) = 0>\<^sub>N x "
+ shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})< 0) = 0>\<^sub>N x "
proof-
have " \<exists> a b. x = (a,b)" by simp
then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -345,7 +345,7 @@
qed
lemma Nle0_iff[simp]:assumes nx: "isnormNum x"
- shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
+ shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
proof-
have " \<exists> a b. x = (a,b)" by simp
then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -357,7 +357,7 @@
ultimately show ?thesis by blast
qed
-lemma Ngt0_iff[simp]:assumes nx: "isnormNum x" shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})> 0) = 0<\<^sub>N x"
+lemma Ngt0_iff[simp]:assumes nx: "isnormNum x" shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})> 0) = 0<\<^sub>N x"
proof-
have " \<exists> a b. x = (a,b)" by simp
then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -369,7 +369,7 @@
ultimately show ?thesis by blast
qed
lemma Nge0_iff[simp]:assumes nx: "isnormNum x"
- shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) \<ge> 0) = 0\<le>\<^sub>N x"
+ shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) \<ge> 0) = 0\<le>\<^sub>N x"
proof-
have " \<exists> a b. x = (a,b)" by simp
then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -382,7 +382,7 @@
qed
lemma Nlt_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y"
- shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) < INum y) = (x <\<^sub>N y)"
+ shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) < INum y) = (x <\<^sub>N y)"
proof-
let ?z = "0::'a"
have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)" using nx ny by simp
@@ -391,7 +391,7 @@
qed
lemma Nle_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y"
- shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})\<le> INum y) = (x \<le>\<^sub>N y)"
+ shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})\<le> INum y) = (x \<le>\<^sub>N y)"
proof-
have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))" using nx ny by simp
also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))" using Nle0_iff[OF Nsub_normN[OF ny]] by simp
@@ -399,7 +399,7 @@
qed
lemma Nadd_commute:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "x +\<^sub>N y = y +\<^sub>N x"
proof-
have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
@@ -408,7 +408,7 @@
qed
lemma [simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "(0, b) +\<^sub>N y = normNum y"
and "(a, 0) +\<^sub>N y = normNum y"
and "x +\<^sub>N (0, b) = normNum x"
@@ -420,7 +420,7 @@
done
lemma normNum_nilpotent_aux[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
assumes nx: "isnormNum x"
shows "normNum x = x"
proof-
@@ -432,7 +432,7 @@
qed
lemma normNum_nilpotent[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "normNum (normNum x) = normNum x"
by simp
@@ -440,11 +440,11 @@
by (simp_all add: normNum_def)
lemma normNum_Nadd:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
lemma Nadd_normNum1[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "normNum x +\<^sub>N y = x +\<^sub>N y"
proof-
have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
@@ -454,7 +454,7 @@
qed
lemma Nadd_normNum2[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "x +\<^sub>N normNum y = x +\<^sub>N y"
proof-
have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
@@ -464,7 +464,7 @@
qed
lemma Nadd_assoc:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
proof-
have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
@@ -476,7 +476,7 @@
by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
lemma Nmul_assoc:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
assumes nx: "isnormNum x" and ny:"isnormNum y" and nz:"isnormNum z"
shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
proof-
@@ -487,7 +487,7 @@
qed
lemma Nsub0:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
assumes x: "isnormNum x" and y:"isnormNum y" shows "(x -\<^sub>N y = 0\<^sub>N) = (x = y)"
proof-
{ fix h :: 'a
@@ -502,7 +502,7 @@
by (simp_all add: Nmul_def Let_def split_def)
lemma Nmul_eq0[simp]:
- assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
assumes nx:"isnormNum x" and ny: "isnormNum y"
shows "(x*\<^sub>N y = 0\<^sub>N) = (x = 0\<^sub>N \<or> y = 0\<^sub>N)"
proof-
--- a/src/HOL/Library/Binomial.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Binomial.thy Mon Apr 26 09:45:22 2010 -0700
@@ -236,10 +236,10 @@
have th1: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) =
(\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
apply (rule setprod_reindex_cong[where f = "Suc"])
- using n0 by (auto simp add: expand_fun_eq ring_simps)
+ using n0 by (auto simp add: expand_fun_eq field_simps)
have ?thesis apply (simp add: pochhammer_def)
unfolding setprod_insert[OF th0, unfolded eq]
- using th1 by (simp add: ring_simps)}
+ using th1 by (simp add: field_simps)}
ultimately show ?thesis by blast
qed
@@ -378,10 +378,10 @@
by simp
from n h th0
have "fact k * fact (n - k) * (n choose k) = k * (fact h * fact (m - h) * (m choose h)) + (m - h) * (fact k * fact (m - k) * (m choose k))"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
also have "\<dots> = (k + (m - h)) * fact m"
using H[rule_format, OF mn hm'] H[rule_format, OF mn km]
- by (simp add: ring_simps)
+ by (simp add: field_simps)
finally have ?ths using h n km by simp}
moreover have "n=0 \<or> k = 0 \<or> k = n \<or> (EX m h. n=Suc m \<and> k = Suc h \<and> h < m)" using kn by presburger
ultimately show ?ths by blast
@@ -391,13 +391,13 @@
assumes kn: "k \<le> n"
shows "(of_nat (n choose k) :: 'a::field_char_0) = of_nat (fact n) / (of_nat (fact k) * of_nat (fact (n - k)))"
using binomial_fact_lemma[OF kn]
- by (simp add: field_eq_simps of_nat_mult [symmetric])
+ by (simp add: field_simps of_nat_mult [symmetric])
lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k"
proof-
{assume kn: "k > n"
from kn binomial_eq_0[OF kn] have ?thesis
- by (simp add: gbinomial_pochhammer field_eq_simps
+ by (simp add: gbinomial_pochhammer field_simps
pochhammer_of_nat_eq_0_iff)}
moreover
{assume "k=0" then have ?thesis by simp}
@@ -414,13 +414,13 @@
apply clarsimp
apply (presburger)
apply presburger
- by (simp add: expand_fun_eq ring_simps of_nat_add[symmetric] del: of_nat_add)
+ by (simp add: expand_fun_eq field_simps of_nat_add[symmetric] del: of_nat_add)
have th0: "finite {1..n - Suc h}" "finite {n - h .. n}"
"{1..n - Suc h} \<inter> {n - h .. n} = {}" and eq3: "{1..n - Suc h} \<union> {n - h .. n} = {1..n}" using h kn by auto
from eq[symmetric]
have ?thesis using kn
apply (simp add: binomial_fact[OF kn, where ?'a = 'a]
- gbinomial_pochhammer field_eq_simps pochhammer_Suc_setprod)
+ gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc)
unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
unfolding mult_assoc[symmetric]
@@ -449,9 +449,9 @@
have "?r = ((- 1) ^n * pochhammer (- a) n / of_nat (fact n)) * (of_nat n - (- a + of_nat n))"
unfolding gbinomial_pochhammer
pochhammer_Suc fact_Suc of_nat_mult right_diff_distrib power_Suc
- by (simp add: field_eq_simps del: of_nat_Suc)
+ by (simp add: field_simps del: of_nat_Suc)
also have "\<dots> = ?l" unfolding gbinomial_pochhammer
- by (simp add: ring_simps)
+ by (simp add: field_simps)
finally show ?thesis ..
qed
@@ -482,17 +482,17 @@
have "of_nat (fact (Suc k)) * (a gchoose k + (a gchoose (Suc k))) = ((a gchoose Suc h) * of_nat (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0\<Colon>nat..Suc h}. a - of_nat i)"
unfolding h
- apply (simp add: ring_simps del: fact_Suc)
+ apply (simp add: field_simps del: fact_Suc)
unfolding gbinomial_mult_fact'
apply (subst fact_Suc)
unfolding of_nat_mult
apply (subst mult_commute)
unfolding mult_assoc
unfolding gbinomial_mult_fact
- by (simp add: ring_simps)
+ by (simp add: field_simps)
also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc
- by (simp add: ring_simps h)
+ by (simp add: field_simps h)
also have "\<dots> = (\<Prod>i\<in>{0..k}. (a + 1) - of_nat i)"
using eq0
unfolding h setprod_nat_ivl_1_Suc
--- a/src/HOL/Library/Bit.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Bit.thy Mon Apr 26 09:45:22 2010 -0700
@@ -49,7 +49,7 @@
subsection {* Type @{typ bit} forms a field *}
-instantiation bit :: "{field, division_by_zero}"
+instantiation bit :: "{field, division_ring_inverse_zero}"
begin
definition plus_bit_def:
--- a/src/HOL/Library/Formal_Power_Series.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Formal_Power_Series.thy Mon Apr 26 09:45:22 2010 -0700
@@ -588,7 +588,7 @@
from k have "real k > - log y x" by simp
then have "ln y * real k > - ln x" unfolding log_def
using ln_gt_zero_iff[OF yp] y1
- by (simp add: minus_divide_left field_simps field_eq_simps del:minus_divide_left[symmetric])
+ by (simp add: minus_divide_left field_simps del:minus_divide_left[symmetric])
then have "ln y * real k + ln x > 0" by simp
then have "exp (real k * ln y + ln x) > exp 0"
by (simp add: mult_ac)
@@ -596,7 +596,7 @@
unfolding exp_zero exp_add exp_real_of_nat_mult
exp_ln[OF xp] exp_ln[OF yp] by simp
then have "x > (1/y)^k" using yp
- by (simp add: field_simps field_eq_simps nonzero_power_divide)
+ by (simp add: field_simps nonzero_power_divide)
then show ?thesis using kp by blast
qed
lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def)
@@ -693,7 +693,7 @@
from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} =
- (f$0) * (inverse f)$n"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
unfolding fps_mult_nth ifn ..
also have "\<dots> = f$0 * natfun_inverse f n
@@ -766,7 +766,7 @@
lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n+1)" by (simp add: fps_deriv_def)
lemma fps_deriv_linear[simp]: "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g"
- unfolding fps_eq_iff fps_add_nth fps_const_mult_left fps_deriv_nth by (simp add: ring_simps)
+ unfolding fps_eq_iff fps_add_nth fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
lemma fps_deriv_mult[simp]:
fixes f :: "('a :: comm_ring_1) fps"
@@ -817,7 +817,7 @@
unfolding s0 s1
unfolding setsum_addf[symmetric] setsum_right_distrib
apply (rule setsum_cong2)
- by (auto simp add: of_nat_diff ring_simps)
+ by (auto simp add: of_nat_diff field_simps)
finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .}
then show ?thesis unfolding fps_eq_iff by auto
qed
@@ -878,7 +878,7 @@
proof-
have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0" by simp
also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f-g)$0)" unfolding fps_deriv_eq_0_iff ..
- finally show ?thesis by (simp add: ring_simps)
+ finally show ?thesis by (simp add: field_simps)
qed
lemma fps_deriv_eq_iff_ex: "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>(c::'a::{idom,semiring_char_0}). f = fps_const c + g)"
@@ -929,7 +929,7 @@
qed
lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f:: ('a::comm_semiring_1) fps)) $ 0 = of_nat (fact k) * f$(k)"
- by (induct k arbitrary: f) (auto simp add: ring_simps of_nat_mult)
+ by (induct k arbitrary: f) (auto simp add: field_simps of_nat_mult)
subsection {* Powers*}
@@ -943,7 +943,7 @@
case (Suc n)
note h = Suc.hyps[OF `a$0 = 1`]
show ?case unfolding power_Suc fps_mult_nth
- using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: ring_simps)
+ using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: field_simps)
qed
lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
@@ -1005,7 +1005,7 @@
case 0 thus ?case by (simp add: power_0)
next
case (Suc n)
- have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: ring_simps power_Suc)
+ have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps power_Suc)
also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" by (simp add: fps_mult_nth)
also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
apply (rule setsum_mono_zero_right)
@@ -1045,8 +1045,8 @@
qed
lemma fps_deriv_power: "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)"
- apply (induct n, auto simp add: power_Suc ring_simps fps_const_add[symmetric] simp del: fps_const_add)
- by (case_tac n, auto simp add: power_Suc ring_simps)
+ apply (induct n, auto simp add: power_Suc field_simps fps_const_add[symmetric] simp del: fps_const_add)
+ by (case_tac n, auto simp add: power_Suc field_simps)
lemma fps_inverse_deriv:
fixes a:: "('a :: field) fps"
@@ -1060,11 +1060,11 @@
with inverse_mult_eq_1[OF a0]
have "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) = 0"
unfolding power2_eq_square
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
by (simp add: mult_assoc[symmetric])
hence "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * inverse a ^ 2 = 0 - fps_deriv a * inverse a ^ 2"
by simp
- then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: ring_simps)
+ then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: field_simps)
qed
lemma fps_inverse_mult:
@@ -1084,7 +1084,7 @@
from inverse_mult_eq_1[OF ab0]
have "inverse (a*b) * (a*b) * inverse a * inverse b = 1 * inverse a * inverse b" by simp
then have "inverse (a*b) * (inverse a * a) * (inverse b * b) = inverse a * inverse b"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
then have ?thesis using inverse_mult_eq_1[OF a0] inverse_mult_eq_1[OF b0] by simp}
ultimately show ?thesis by blast
qed
@@ -1105,7 +1105,7 @@
assumes a0: "b$0 \<noteq> 0"
shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b ^ 2"
using fps_inverse_deriv[OF a0]
- by (simp add: fps_divide_def ring_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
+ by (simp add: fps_divide_def field_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
@@ -1121,7 +1121,7 @@
proof-
have eq: "(1 + X) * ?r = 1"
unfolding minus_one_power_iff
- by (auto simp add: ring_simps fps_eq_iff)
+ by (auto simp add: field_simps fps_eq_iff)
show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
qed
@@ -1185,7 +1185,7 @@
next
case (Suc k)
note th = Suc.hyps[symmetric]
- have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
+ have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: field_simps)
also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
using th
unfolding fps_sub_nth by simp
@@ -1209,10 +1209,10 @@
definition "XD = op * X o fps_deriv"
lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: ('a::comm_ring_1) fps)"
- by (simp add: XD_def ring_simps)
+ by (simp add: XD_def field_simps)
lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
- by (simp add: XD_def ring_simps)
+ by (simp add: XD_def field_simps)
lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) = fps_const c * XD a + fps_const d * XD (b :: ('a::comm_ring_1) fps)"
by simp
@@ -1226,7 +1226,7 @@
lemma fps_mult_XD_shift:
"(XD ^^ k) (a:: ('a::{comm_ring_1}) fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
- by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def)
+ by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff field_simps del: One_nat_def)
subsubsection{* Rule 3 is trivial and is given by @{text fps_times_def}*}
subsubsection{* Rule 5 --- summation and "division" by (1 - X)*}
@@ -1688,7 +1688,7 @@
then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
by (simp add: natpermute_max_card[OF nz, simplified])
also have "\<dots> = a$n - setsum ?f ?Pnknn"
- unfolding n1 using r00 a0 by (simp add: field_eq_simps fps_radical_def del: of_nat_Suc)
+ unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
@@ -1764,7 +1764,7 @@
shows "a = b / c"
proof-
from eq have "a * c * inverse c = b * inverse c" by simp
- hence "a * (inverse c * c) = b/c" by (simp only: field_eq_simps divide_inverse)
+ hence "a * (inverse c * c) = b/c" by (simp only: field_simps divide_inverse)
then show "a = b/c" unfolding field_inverse[OF c0] by simp
qed
@@ -1837,7 +1837,7 @@
show "a$(xs !i) = ?r$(xs!i)" .
qed
have th00: "\<And>(x::'a). of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
- by (simp add: field_eq_simps del: of_nat_Suc)
+ by (simp add: field_simps del: of_nat_Suc)
from H have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff)
also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
unfolding fps_power_nth_Suc
@@ -1854,7 +1854,7 @@
then have "a$n = ?r $n"
apply (simp del: of_nat_Suc)
unfolding fps_radical_def n1
- by (simp add: field_eq_simps n1 th00 del: of_nat_Suc)}
+ by (simp add: field_simps n1 th00 del: of_nat_Suc)}
ultimately show "a$n = ?r $ n" by (cases n, auto)
qed}
then have "a = ?r" by (simp add: fps_eq_iff)}
@@ -2018,11 +2018,11 @@
proof-
{fix n
have "(fps_deriv (a oo b))$n = setsum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
- by (simp add: fps_compose_def ring_simps setsum_right_distrib del: of_nat_Suc)
+ by (simp add: fps_compose_def field_simps setsum_right_distrib del: of_nat_Suc)
also have "\<dots> = setsum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
- by (simp add: ring_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
+ by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
- unfolding fps_mult_left_const_nth by (simp add: ring_simps)
+ unfolding fps_mult_left_const_nth by (simp add: field_simps)
also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
unfolding fps_mult_nth ..
also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
@@ -2170,7 +2170,7 @@
by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0')
lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
- by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_addf)
+ by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_addf)
lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
proof-
@@ -2212,7 +2212,7 @@
apply (simp add: fps_mult_nth setsum_right_distrib)
apply (subst setsum_commute)
apply (rule setsum_cong2)
- by (auto simp add: ring_simps)
+ by (auto simp add: field_simps)
also have "\<dots> = ?l"
apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
apply (rule setsum_cong2)
@@ -2312,7 +2312,7 @@
qed
lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
- by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_negf[symmetric])
+ by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
lemma fps_compose_sub_distrib:
shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
@@ -2469,7 +2469,7 @@
proof-
let ?r = "fps_inv"
have ra0: "?r a $ 0 = 0" by (simp add: fps_inv_def)
- from a1 have ra1: "?r a $ 1 \<noteq> 0" by (simp add: fps_inv_def field_eq_simps)
+ from a1 have ra1: "?r a $ 1 \<noteq> 0" by (simp add: fps_inv_def field_simps)
have X0: "X$0 = 0" by simp
from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
then have "?r (?r a) oo ?r a oo a = X oo a" by simp
@@ -2486,7 +2486,7 @@
proof-
let ?r = "fps_ginv"
from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def)
- from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0" by (simp add: fps_ginv_def field_eq_simps)
+ from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0" by (simp add: fps_ginv_def field_simps)
from fps_ginv[OF rca0 rca1]
have "?r b (?r c a) oo ?r c a = b" .
then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp
@@ -2534,8 +2534,8 @@
proof-
{fix n
have "?l$n = ?r $ n"
- apply (auto simp add: E_def field_eq_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
- by (simp add: of_nat_mult ring_simps)}
+ apply (auto simp add: E_def field_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
+ by (simp add: of_nat_mult field_simps)}
then show ?thesis by (simp add: fps_eq_iff)
qed
@@ -2545,15 +2545,15 @@
proof-
{assume d: ?lhs
from d have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
- by (simp add: fps_deriv_def fps_eq_iff field_eq_simps del: of_nat_Suc)
+ by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
{fix n have "a$n = a$0 * c ^ n/ (of_nat (fact n))"
apply (induct n)
apply simp
unfolding th
using fact_gt_zero_nat
- apply (simp add: field_eq_simps del: of_nat_Suc fact_Suc)
+ apply (simp add: field_simps del: of_nat_Suc fact_Suc)
apply (drule sym)
- by (simp add: ring_simps of_nat_mult power_Suc)}
+ by (simp add: field_simps of_nat_mult power_Suc)}
note th' = this
have ?rhs
by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro : th')}
@@ -2570,7 +2570,7 @@
lemma E_add_mult: "E (a + b) = E (a::'a::field_char_0) * E b" (is "?l = ?r")
proof-
have "fps_deriv (?r) = fps_const (a+b) * ?r"
- by (simp add: fps_const_add[symmetric] ring_simps del: fps_const_add)
+ by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
then have "?r = ?l" apply (simp only: E_unique_ODE)
by (simp add: fps_mult_nth E_def)
then show ?thesis ..
@@ -2618,13 +2618,13 @@
(is "inverse ?l = ?r")
proof-
have th: "?l * ?r = 1"
- by (auto simp add: ring_simps fps_eq_iff minus_one_power_iff)
+ by (auto simp add: field_simps fps_eq_iff minus_one_power_iff)
have th': "?l $ 0 \<noteq> 0" by (simp add: )
from fps_inverse_unique[OF th' th] show ?thesis .
qed
lemma E_power_mult: "(E (c::'a::field_char_0))^n = E (of_nat n * c)"
- by (induct n, auto simp add: ring_simps E_add_mult power_Suc)
+ by (induct n, auto simp add: field_simps E_add_mult power_Suc)
lemma radical_E:
assumes r: "r (Suc k) 1 = 1"
@@ -2649,18 +2649,18 @@
text{* The generalized binomial theorem as a consequence of @{thm E_add_mult} *}
lemma gbinomial_theorem:
- "((a::'a::{field_char_0, division_by_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+ "((a::'a::{field_char_0, division_ring_inverse_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
proof-
from E_add_mult[of a b]
have "(E (a + b)) $ n = (E a * E b)$n" by simp
then have "(a + b) ^ n = (\<Sum>i\<Colon>nat = 0\<Colon>nat..n. a ^ i * b ^ (n - i) * (of_nat (fact n) / of_nat (fact i * fact (n - i))))"
- by (simp add: field_eq_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
+ by (simp add: field_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
then show ?thesis
apply simp
apply (rule setsum_cong2)
apply simp
apply (frule binomial_fact[where ?'a = 'a, symmetric])
- by (simp add: field_eq_simps of_nat_mult)
+ by (simp add: field_simps of_nat_mult)
qed
text{* And the nat-form -- also available from Binomial.thy *}
@@ -2683,7 +2683,7 @@
by (simp add: L_def fps_eq_iff del: of_nat_Suc)
lemma L_nth: "L c $ n = (if n=0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
- by (simp add: L_def field_eq_simps)
+ by (simp add: L_def field_simps)
lemma L_0[simp]: "L c $ 0 = 0" by (simp add: L_def)
lemma L_E_inv:
@@ -2694,9 +2694,9 @@
have b0: "?b $ 0 = 0" by simp
have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
have "fps_deriv (E a - 1) oo fps_inv (E a - 1) = (fps_const a * (E a - 1) + fps_const a) oo fps_inv (E a - 1)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
also have "\<dots> = fps_const a * (X + 1)" apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
- by (simp add: ring_simps)
+ by (simp add: field_simps)
finally have eq: "fps_deriv (E a - 1) oo fps_inv (E a - 1) = fps_const a * (X + 1)" .
from fps_inv_deriv[OF b0 b1, unfolded eq]
have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
@@ -2713,7 +2713,7 @@
shows "L c + L d = fps_const (c+d) * L (c*d)"
(is "?r = ?l")
proof-
- from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_eq_simps)
+ from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + X)"
by (simp add: fps_deriv_L fps_const_add[symmetric] algebra_simps del: fps_const_add)
also have "\<dots> = fps_deriv ?l"
@@ -2743,7 +2743,7 @@
have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
apply (simp only: fps_divide_def mult_assoc[symmetric] inverse_mult_eq_1[OF x10])
- by (simp add: ring_simps)
+ by (simp add: field_simps)
finally have eq: "?l = ?r \<longleftrightarrow> ?lhs" by simp
moreover
{assume h: "?l = ?r"
@@ -2752,8 +2752,8 @@
from lrn
have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n"
- apply (simp add: ring_simps del: of_nat_Suc)
- by (cases n, simp_all add: field_eq_simps del: of_nat_Suc)
+ apply (simp add: field_simps del: of_nat_Suc)
+ by (cases n, simp_all add: field_simps del: of_nat_Suc)
}
note th0 = this
{fix n have "a$n = (c gchoose n) * a$0"
@@ -2762,24 +2762,24 @@
next
case (Suc m)
thus ?case unfolding th0
- apply (simp add: field_eq_simps del: of_nat_Suc)
+ apply (simp add: field_simps del: of_nat_Suc)
unfolding mult_assoc[symmetric] gbinomial_mult_1
- by (simp add: ring_simps)
+ by (simp add: field_simps)
qed}
note th1 = this
have ?rhs
apply (simp add: fps_eq_iff)
apply (subst th1)
- by (simp add: ring_simps)}
+ by (simp add: field_simps)}
moreover
{assume h: ?rhs
have th00:"\<And>x y. x * (a$0 * y) = a$0 * (x*y)" by (simp add: mult_commute)
have "?l = ?r"
apply (subst h)
apply (subst (2) h)
- apply (clarsimp simp add: fps_eq_iff ring_simps)
+ apply (clarsimp simp add: fps_eq_iff field_simps)
unfolding mult_assoc[symmetric] th00 gbinomial_mult_1
- by (simp add: ring_simps gbinomial_mult_1)}
+ by (simp add: field_simps gbinomial_mult_1)}
ultimately show ?thesis by blast
qed
@@ -2798,9 +2798,9 @@
have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)" by simp
also have "\<dots> = inverse (1 + X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
unfolding fps_binomial_deriv
- by (simp add: fps_divide_def ring_simps)
+ by (simp add: fps_divide_def field_simps)
also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
- by (simp add: ring_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
+ by (simp add: field_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
by (simp add: fps_divide_def)
have "?P = fps_const (?P$0) * ?b (c + d)"
@@ -2880,7 +2880,7 @@
using kn pochhammer_minus'[where k=k and n=n and b=b]
apply (simp add: pochhammer_same)
using bn0
- by (simp add: field_eq_simps power_add[symmetric])}
+ by (simp add: field_simps power_add[symmetric])}
moreover
{assume nk: "k \<noteq> n"
have m1nk: "?m1 n = setprod (%i. - 1) {0..m}"
@@ -2905,7 +2905,7 @@
unfolding m1nk
unfolding m h pochhammer_Suc_setprod
- apply (simp add: field_eq_simps del: fact_Suc id_def)
+ apply (simp add: field_simps del: fact_Suc id_def)
unfolding fact_altdef_nat id_def
unfolding of_nat_setprod
unfolding setprod_timesf[symmetric]
@@ -2942,10 +2942,10 @@
apply auto
done
then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}"
- using nz' by (simp add: field_eq_simps)
+ using nz' by (simp add: field_simps)
have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
using bnz0
- by (simp add: field_eq_simps)
+ by (simp add: field_simps)
also have "\<dots> = b gchoose (n - k)"
unfolding th1 th2
using kn' by (simp add: gbinomial_def)
@@ -2959,15 +2959,15 @@
note th00 = this
have "?r = ((a + b) gchoose n) * (of_nat (fact n)/ (?m1 n * pochhammer (- b) n))"
unfolding gbinomial_pochhammer
- using bn0 by (auto simp add: field_eq_simps)
+ using bn0 by (auto simp add: field_simps)
also have "\<dots> = ?l"
unfolding gbinomial_Vandermonde[symmetric]
apply (simp add: th00)
unfolding gbinomial_pochhammer
- using bn0 apply (simp add: setsum_left_distrib setsum_right_distrib field_eq_simps)
+ using bn0 apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
apply (rule setsum_cong2)
apply (drule th00(2))
- by (simp add: field_eq_simps power_add[symmetric])
+ by (simp add: field_simps power_add[symmetric])
finally show ?thesis by simp
qed
@@ -2992,7 +2992,7 @@
have nz: "pochhammer c n \<noteq> 0" using c
by (simp add: pochhammer_eq_0_iff)
from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
- show ?thesis using nz by (simp add: field_eq_simps setsum_right_distrib)
+ show ?thesis using nz by (simp add: field_simps setsum_right_distrib)
qed
subsubsection{* Formal trigonometric functions *}
@@ -3014,11 +3014,11 @@
using en by (simp add: fps_sin_def)
also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
unfolding fact_Suc of_nat_mult
- by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+ by (simp add: field_simps del: of_nat_add of_nat_Suc)
also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
- by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+ by (simp add: field_simps del: of_nat_add of_nat_Suc)
finally have "?lhs $n = ?rhs$n" using en
- by (simp add: fps_cos_def ring_simps power_Suc )}
+ by (simp add: fps_cos_def field_simps power_Suc )}
then show "?lhs $ n = ?rhs $ n"
by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
qed
@@ -3038,13 +3038,13 @@
using en by (simp add: fps_cos_def)
also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
unfolding fact_Suc of_nat_mult
- by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+ by (simp add: field_simps del: of_nat_add of_nat_Suc)
also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
- by (simp add: field_eq_simps del: of_nat_add of_nat_Suc)
+ by (simp add: field_simps del: of_nat_add of_nat_Suc)
also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
unfolding th0 unfolding th1[OF en] by simp
finally have "?lhs $n = ?rhs$n" using en
- by (simp add: fps_sin_def ring_simps power_Suc)}
+ by (simp add: fps_sin_def field_simps power_Suc)}
then show "?lhs $ n = ?rhs $ n"
by (cases "even n", simp_all add: fps_deriv_def fps_sin_def
fps_cos_def)
@@ -3055,7 +3055,7 @@
proof-
have "fps_deriv ?lhs = 0"
apply (simp add: fps_deriv_power fps_sin_deriv fps_cos_deriv power_Suc)
- by (simp add: ring_simps fps_const_neg[symmetric] del: fps_const_neg)
+ by (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
then have "?lhs = fps_const (?lhs $ 0)"
unfolding fps_deriv_eq_0_iff .
also have "\<dots> = 1"
@@ -3177,7 +3177,7 @@
have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
show ?thesis
using fps_sin_cos_sum_of_squares[of c]
- apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] ring_simps power2_eq_square del: fps_const_neg)
+ apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] field_simps power2_eq_square del: fps_const_neg)
unfolding right_distrib[symmetric]
by simp
qed
@@ -3252,7 +3252,7 @@
subsection {* Hypergeometric series *}
-definition "F as bs (c::'a::{field_char_0, division_by_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
+definition "F as bs (c::'a::{field_char_0, division_ring_inverse_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
lemma F_nth[simp]: "F as bs c $ n = (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
by (simp add: F_def)
@@ -3321,9 +3321,9 @@
by (simp add: fps_eq_iff fps_integral_def)
lemma F_minus_nat:
- "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
+ "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
(pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)"
- "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
+ "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
(pochhammer (- of_nat (m + n)) k * of_nat (fact k)) else 0)"
by (auto simp add: pochhammer_eq_0_iff)
--- a/src/HOL/Library/Fraction_Field.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Fraction_Field.thy Mon Apr 26 09:45:22 2010 -0700
@@ -267,7 +267,7 @@
end
-instance fract :: (idom) division_by_zero
+instance fract :: (idom) division_ring_inverse_zero
proof
show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
(simp add: fract_collapse)
@@ -450,7 +450,7 @@
by simp
with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
by (simp add: mult_le_cancel_right)
- with neq show ?thesis by (simp add: ring_simps)
+ with neq show ?thesis by (simp add: field_simps)
qed
qed
show "q < r ==> 0 < s ==> s * q < s * r"
--- a/src/HOL/Library/Numeral_Type.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Numeral_Type.thy Mon Apr 26 09:45:22 2010 -0700
@@ -213,7 +213,7 @@
lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)"
apply (intro_classes, unfold definitions)
-apply (simp_all add: Rep_simps zmod_simps ring_simps)
+apply (simp_all add: Rep_simps zmod_simps field_simps)
done
end
--- a/src/HOL/Library/Polynomial.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Polynomial.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1093,10 +1093,10 @@
apply (cases "r = 0")
apply (cases "r' = 0")
apply (simp add: pdivmod_rel_def)
-apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq)
+apply (simp add: pdivmod_rel_def field_simps degree_mult_eq)
apply (cases "r' = 0")
apply (simp add: pdivmod_rel_def degree_mult_eq)
-apply (simp add: pdivmod_rel_def ring_simps)
+apply (simp add: pdivmod_rel_def field_simps)
apply (simp add: degree_mult_eq degree_add_less)
done
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Mon Apr 26 09:45:22 2010 -0700
@@ -1282,9 +1282,9 @@
fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
val concl = Thm.dest_arg o cprop_of
val shuffle1 =
- fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
+ fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps) })
val shuffle2 =
- fconv_rule (rewr_conv @{lemma "(x + a == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps)})
+ fconv_rule (rewr_conv @{lemma "(x + a == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)})
fun substitutable_monomial fvs tm = case term_of tm of
Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm)
else raise Failure "substitutable_monomial"
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1381,7 +1381,7 @@
apply(rule_tac x=x in bexI) apply assumption+ apply(rule continuous_on_intros)+
unfolding frontier_cball subset_eq Ball_def image_iff apply(rule,rule,erule bexE)
unfolding vector_dist_norm apply(simp add: * norm_minus_commute) . note x = this
- hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:group_simps)
+ hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:algebra_simps)
hence "a = x" unfolding scaleR_left_distrib[THEN sym] by auto
thus False using x using assms by auto qed
@@ -1394,7 +1394,7 @@
"interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
(\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i))"
apply rule unfolding Cart_eq interval_bij_def vector_component_simps
- by(auto simp add:group_simps field_simps add_divide_distrib[THEN sym])
+ by(auto simp add: field_simps add_divide_distrib[THEN sym])
lemma continuous_interval_bij:
"continuous (at x) (interval_bij (a,b::real^'n) (u,v))"
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 26 09:45:22 2010 -0700
@@ -645,7 +645,7 @@
using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
apply - apply(rule add_mono) by auto
- hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps) }
+ hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps) }
thus ?thesis unfolding convex_on_def by auto
qed
@@ -653,7 +653,7 @@
assumes "0 \<le> (c::real)" "convex_on s f"
shows "convex_on s (\<lambda>x. c * f x)"
proof-
- have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps)
+ have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: field_simps)
show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto
qed
@@ -1059,7 +1059,7 @@
proof-
have fin:"finite {a,b,c}" "finite {b,c}" "finite {c}" by auto
have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
- "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: ring_simps)
+ "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: field_simps)
show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and *
unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto
apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp
@@ -2309,7 +2309,7 @@
} moreover
{ fix a b assume "\<not> u * a + v * b \<le> a"
hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps)
- hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps)
+ hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_simps)
hence "u * a + v * b \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
using as(3-) dimindex_ge_1 by auto qed
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1,11 +1,12 @@
-(* Title: HOL/Library/Convex_Euclidean_Space.thy
- Author: John Harrison
- Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+(* Title: HOL/Multivariate_Analysis/Derivative.thy
+ Author: John Harrison
+ Translation from HOL Light: Robert Himmelmann, TU Muenchen
+*)
header {* Multivariate calculus in Euclidean space. *}
theory Derivative
- imports Brouwer_Fixpoint RealVector
+imports Brouwer_Fixpoint RealVector
begin
@@ -40,7 +41,7 @@
show ?l unfolding deriv_def LIM_def apply safe apply(drule as,safe)
apply(rule_tac x=d in exI,safe) apply(erule_tac x="xa + x" in allE)
unfolding vector_dist_norm diff_0_right norm_scaleR
- unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps *) qed
+ unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:algebra_simps *) qed
lemma FDERIV_conv_has_derivative:"FDERIV f (x::'a::{real_normed_vector,perfect_space}) :> f' = (f has_derivative f') (at x)" (is "?l = ?r") proof
assume ?l note as = this[unfolded fderiv_def]
@@ -50,14 +51,14 @@
thus "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
dist ((1 / norm (xa - netlimit (at x))) *\<^sub>R (f xa - (f (netlimit (at x)) + f' (xa - netlimit (at x))))) (0) < e"
apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa - x" in allE)
- unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed next
+ unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq) qed next
assume ?r note as = this[unfolded has_derivative_def]
show ?l unfolding fderiv_def LIM_def apply-apply(rule,rule as[THEN conjunct1]) proof(rule,rule)
fix e::real assume "e>0"
guess d using as[THEN conjunct2,unfolded Lim_at,rule_format,OF`e>0`] ..
thus "\<exists>s>0. \<forall>xa. xa \<noteq> 0 \<and> dist xa 0 < s \<longrightarrow> dist (norm (f (x + xa) - f x - f' xa) / norm xa) 0 < e" apply-
apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa + x" in allE)
- unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed qed
+ unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq add.commute) qed qed
subsection {* These are the only cases we'll care about, probably. *}
@@ -76,7 +77,7 @@
(\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm(x' - x) \<and> norm(x' - x) < d
\<longrightarrow> norm(f x' - f x - f'(x' - x)) / norm(x' - x) < e)"
unfolding has_derivative_within Lim_within vector_dist_norm
- unfolding diff_0_right norm_mul by(simp add: group_simps)
+ unfolding diff_0_right norm_mul by (simp add: diff_diff_eq)
lemma has_derivative_at':
"(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -186,14 +187,14 @@
note as = assms[unfolded has_derivative_def]
show ?thesis unfolding has_derivative_def apply(rule,rule bounded_linear_add)
using Lim_add[OF as(1)[THEN conjunct2] as(2)[THEN conjunct2]] and as
- by(auto simp add:group_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
+ by (auto simp add:algebra_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
lemma has_derivative_add_const:"(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net"
apply(drule has_derivative_add) apply(rule has_derivative_const) by auto
lemma has_derivative_sub:
"(f has_derivative f') net \<Longrightarrow> (g has_derivative g') net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_derivative (\<lambda>h. f'(h) - g'(h))) net"
- apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:group_simps)
+ apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:algebra_simps)
lemma has_derivative_setsum: assumes "finite s" "\<forall>a\<in>s. ((f a) has_derivative (f' a)) net"
shows "((\<lambda>x. setsum (\<lambda>a. f a x) s) has_derivative (\<lambda>h. setsum (\<lambda>a. f' a h) s)) net"
@@ -393,8 +394,8 @@
case False hence "norm (f y - (f x + f' (y - x))) < e * norm (y - x)" using as(4)[rule_format, OF `y\<in>s`]
unfolding vector_dist_norm diff_0_right norm_mul using as(3)
using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded vector_dist_norm]
- by(auto simp add:linear_0 linear_sub group_simps)
- thus ?thesis by(auto simp add:group_simps) qed qed next
+ by (auto simp add: linear_0 linear_sub)
+ thus ?thesis by(auto simp add:algebra_simps) qed qed next
assume ?rhs thus ?lhs unfolding has_derivative_within Lim_within apply-apply(erule conjE,rule,assumption)
apply(rule,erule_tac x="e/2" in allE,rule,erule impE) defer apply(erule exE,rule_tac x=d in exI)
apply(erule conjE,rule,assumption,rule,rule) unfolding vector_dist_norm diff_0_right norm_scaleR
@@ -402,7 +403,7 @@
fix e d y assume "bounded_linear f'" "0 < e" "0 < d" "y \<in> s" "0 < norm (y - x) \<and> norm (y - x) < d"
"norm (f y - f x - f' (y - x)) \<le> e / 2 * norm (y - x)"
thus "\<bar>1 / norm (y - x)\<bar> * norm (f y - (f x + f' (y - x))) < e"
- apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:group_simps) qed auto qed
+ apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:algebra_simps) qed auto qed
lemma has_derivative_at_alt:
"(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -437,8 +438,8 @@
hence 1:"norm (f y - f x - f' (y - x)) \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x))" using d1 d2 d by auto
have "norm (f y - f x) \<le> norm (f y - f x - f' (y - x)) + norm (f' (y - x))"
- using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:group_simps)
- also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:group_simps)
+ using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:algebra_simps)
+ also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:algebra_simps)
also have "\<dots> \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x)) + B1 * norm (y - x)" apply(rule add_right_mono) using d1 d2 d as by auto
also have "\<dots> \<le> norm (y - x) + B1 * norm (y - x)" by auto
also have "\<dots> = norm (y - x) * (1 + B1)" by(auto simp add:field_simps)
@@ -455,8 +456,8 @@
interpret g': bounded_linear g' using assms(2) by auto
interpret f': bounded_linear f' using assms(1) by auto
have "norm (- g' (f' (y - x)) + g' (f y - f x)) = norm (g' (f y - f x - f' (y - x)))"
- by(auto simp add:group_simps f'.diff g'.diff g'.add)
- also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:group_simps)
+ by(auto simp add:algebra_simps f'.diff g'.diff g'.add)
+ also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:algebra_simps)
also have "\<dots> \<le> B2 * (e / 2 / B2 * norm (y - x))" apply(rule mult_left_mono) using as d1 d2 d B2 by auto
also have "\<dots> \<le> e / 2 * norm (y - x)" using B2 by auto
finally have 5:"norm (- g' (f' (y - x)) + g' (f y - f x)) \<le> e / 2 * norm (y - x)" by auto
@@ -537,7 +538,7 @@
unfolding scaleR_right_distrib by auto
also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))"
unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto
- also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by(auto simp add:group_simps)
+ also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by (auto simp add: add.commute ab_diff_minus)
finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] using norm_basis[of i] unfolding vector_dist_norm
unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib by auto qed qed
@@ -625,7 +626,7 @@
have ***:"\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"])
using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left
- unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding group_simps by (auto intro: mult_pos_pos)
+ unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos)
qed
subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
@@ -729,7 +730,7 @@
shows "norm(f x - f y) \<le> B * norm(x - y)" proof-
let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)"
have *:"\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
- using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:group_simps)
+ using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:algebra_simps)
hence 1:"continuous_on {0..1} (f \<circ> ?p)" apply- apply(rule continuous_on_intros continuous_on_vmul)+
unfolding continuous_on_eq_continuous_within apply(rule,rule differentiable_imp_continuous_within)
unfolding differentiable_def apply(rule_tac x="f' xa" in exI)
@@ -864,7 +865,7 @@
assumes "compact t" "convex t" "t \<noteq> {}" "continuous_on t f"
"\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" "x\<in>s"
shows "\<exists>y\<in>t. f y = x" proof-
- have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:group_simps)
+ have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:algebra_simps)
show ?thesis unfolding * apply(rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"])
apply(rule continuous_on_intros assms)+ using assms(4-6) by auto qed
@@ -909,8 +910,8 @@
finally have *:"norm (x + g' (z - f x) - x) < e0" by auto
have **:"f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto
have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> norm (f (x + g' (z - f x)) - z) + norm (f x - y)"
- using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:group_simps)
- also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding group_simps ** by auto
+ using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:algebra_simps)
+ also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding algebra_simps ** by auto
also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" using as(1)[unfolded mem_cball vector_dist_norm] by auto
also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" using * and B by(auto simp add:field_simps)
also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" by auto
@@ -985,7 +986,7 @@
(* we know for some other reason that the inverse function exists, it's OK. *}
lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g ==> bounded_linear (\<lambda>x. f x - g x)"
- using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:group_simps)
+ using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:algebra_simps)
lemma has_derivative_locally_injective: fixes f::"real^'n \<Rightarrow> real^'m"
assumes "a \<in> s" "open s" "bounded_linear g'" "g' o f'(a) = id"
@@ -1006,7 +1007,7 @@
show "\<forall>x\<in>ball a d. \<forall>x'\<in>ball a d. f x' = f x \<longrightarrow> x' = x" proof(intro strip)
fix x y assume as:"x\<in>ball a d" "y\<in>ball a d" "f x = f y"
def ph \<equiv> "\<lambda>w. w - g'(f w - f x)" have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))"
- unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:group_simps)
+ unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:algebra_simps)
have "norm (ph x - ph y) \<le> (1/2) * norm (x - y)"
apply(rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"])
apply(rule_tac[!] ballI) proof- fix u assume u:"u \<in> ball a d" hence "u\<in>s" using d d2 by auto
@@ -1022,7 +1023,7 @@
unfolding linear_conv_bounded_linear by(rule assms(3) **)+
also have "\<dots> \<le> onorm g' * k" apply(rule mult_left_mono)
using d1[THEN conjunct2,rule_format,of u] using onorm_neg[OF **(1)[unfolded linear_linear]]
- using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:group_simps)
+ using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:algebra_simps)
also have "\<dots> \<le> 1/2" unfolding k_def by auto
finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" by assumption qed
moreover have "norm (ph y - ph x) = norm (y - x)" apply(rule arg_cong[where f=norm])
@@ -1041,7 +1042,7 @@
fix x assume "x\<in>s" show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)"
by(rule has_derivative_intros assms(2)[rule_format] `x\<in>s`)+
{ fix h have "norm (f' m x h - f' n x h) \<le> norm (f' m x h - g' x h) + norm (f' n x h - g' x h)"
- using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:group_simps)
+ using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:algebra_simps)
also have "\<dots> \<le> e * norm h+ e * norm h" using assms(3)[rule_format,OF `N\<le>m` `x\<in>s`, of h] assms(3)[rule_format,OF `N\<le>n` `x\<in>s`, of h]
by(auto simp add:field_simps)
finally have "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" by auto }
@@ -1085,7 +1086,7 @@
have "eventually (\<lambda>xa. norm (f n x - f n y - (f xa x - f xa y)) \<le> e * norm (x - y)) sequentially"
unfolding eventually_sequentially apply(rule_tac x=N in exI) proof(rule,rule)
fix m assume "N\<le>m" thus "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)"
- using N[rule_format, of n m x y] and as by(auto simp add:group_simps) qed
+ using N[rule_format, of n m x y] and as by(auto simp add:algebra_simps) qed
thus "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" apply-
apply(rule Lim_norm_ubound[OF trivial_limit_sequentially, where f="\<lambda>m. (f n x - f n y) - (f m x - f m y)"])
apply(rule Lim_sub Lim_const g[rule_format] as)+ by assumption qed qed
@@ -1122,10 +1123,10 @@
have "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" using d1 and as by auto ultimately
have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)"
using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"]
- by (auto simp add:group_simps) moreover
+ by (auto simp add:algebra_simps) moreover
have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" using N1 `x\<in>s` by auto
ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
- using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:group_simps)
+ using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:algebra_simps)
qed qed qed qed
subsection {* Can choose to line up antiderivatives if we want. *}
@@ -1276,7 +1277,7 @@
unfolding has_vector_derivative_def using has_derivative_id by auto
lemma has_vector_derivative_cmul: "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
- unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:group_simps)
+ unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:algebra_simps)
lemma has_vector_derivative_cmul_eq: assumes "c \<noteq> 0"
shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)"
--- a/src/HOL/Multivariate_Analysis/Determinants.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy Mon Apr 26 09:45:22 2010 -0700
@@ -55,7 +55,7 @@
done
(* FIXME: In Finite_Set there is a useless further assumption *)
-lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
+lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_ring_inverse_zero, field})"
apply (erule finite_induct)
apply (simp)
apply simp
@@ -352,13 +352,13 @@
apply (rule setprod_insert)
apply simp
by blast
- also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
+ also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: field_simps)
also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
unfolding setprod_insert[OF th3] by simp
finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] .
then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
qed
lemma det_row_mul:
@@ -389,14 +389,14 @@
apply (rule setprod_insert)
apply simp
by blast
- also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
+ also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: field_simps)
also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
unfolding th1 by (simp add: mult_ac)
also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
unfolding setprod_insert[OF th3] by simp
finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
qed
lemma det_row_0:
@@ -604,7 +604,7 @@
have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
unfolding setprod_timesf ..
then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
- setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
+ setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: field_simps)
qed
lemma det_mul:
@@ -681,7 +681,7 @@
using permutes_in_image[OF q] by vector
show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
- by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
+ by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose)
qed
}
then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
@@ -772,7 +772,7 @@
have fUk: "finite ?Uk" by simp
have kUk: "k \<notin> ?Uk" by simp
have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
- by (vector ring_simps)
+ by (vector field_simps)
have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext)
have "(\<chi> i. row i A) = A" by (vector row_def)
then have thd1: "det (\<chi> i. row i A) = det A" by simp
@@ -793,7 +793,7 @@
unfolding thd0
unfolding det_row_mul
unfolding th001[of k "\<lambda>i. row i A"]
- unfolding thd1 by (simp add: ring_simps)
+ unfolding thd1 by (simp add: field_simps)
qed
lemma cramer_lemma:
@@ -901,7 +901,7 @@
have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
proof-
fix x:: 'a
- have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps)
+ have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: field_simps)
have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
apply (subst eq_iff_diff_eq_0) by simp
have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
@@ -929,7 +929,7 @@
unfolding dot_norm_neg dist_norm[symmetric]
unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
note fc = this
- show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc ring_simps)
+ show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc field_simps)
qed
lemma isometry_linear:
@@ -980,7 +980,7 @@
using H(5-9)
apply (simp add: norm_eq norm_eq_1)
apply (simp add: inner_simps smult_conv_scaleR) unfolding *
- by (simp add: ring_simps) }
+ by (simp add: field_simps) }
note th0 = this
let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)"
{fix x:: "real ^'n" assume nx: "norm x = 1"
@@ -1079,7 +1079,7 @@
unfolding permutes_sing
apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
apply (simp add: arith_simps(31)[symmetric] del: arith_simps(31))
- by (simp add: ring_simps)
+ by (simp add: field_simps)
qed
end
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 26 09:45:22 2010 -0700
@@ -257,14 +257,14 @@
| "vector_power x (Suc n) = x * vector_power x n"
instance cart :: (semiring,finite) semiring
- apply (intro_classes) by (vector ring_simps)+
+ apply (intro_classes) by (vector field_simps)+
instance cart :: (semiring_0,finite) semiring_0
- apply (intro_classes) by (vector ring_simps)+
+ apply (intro_classes) by (vector field_simps)+
instance cart :: (semiring_1,finite) semiring_1
apply (intro_classes) by vector
instance cart :: (comm_semiring,finite) comm_semiring
- apply (intro_classes) by (vector ring_simps)+
+ apply (intro_classes) by (vector field_simps)+
instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)
instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
@@ -278,7 +278,7 @@
instance cart :: (real_algebra,finite) real_algebra
apply intro_classes
- apply (simp_all add: vector_scaleR_def ring_simps)
+ apply (simp_all add: vector_scaleR_def field_simps)
apply vector
apply vector
done
@@ -318,19 +318,19 @@
lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
by (vector mult_assoc)
lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
- by (vector ring_simps)
+ by (vector field_simps)
lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
- by (vector ring_simps)
+ by (vector field_simps)
lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
- by (vector ring_simps)
+ by (vector field_simps)
lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
- by (vector ring_simps)
+ by (vector field_simps)
lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
by (simp add: Cart_eq)
@@ -752,7 +752,7 @@
lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"
proof-
have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith
- thus ?thesis by (simp add: ring_simps power2_eq_square)
+ thus ?thesis by (simp add: field_simps power2_eq_square)
qed
lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"
@@ -827,7 +827,7 @@
lemma norm_triangle_sub:
fixes x y :: "'a::real_normed_vector"
shows "norm x \<le> norm y + norm (x - y)"
- using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
+ using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps)
lemma component_le_norm: "\<bar>x$i\<bar> <= norm x"
apply (simp add: norm_vector_def)
@@ -901,7 +901,7 @@
unfolding power2_norm_eq_inner inner_simps inner_commute by auto
lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"
- unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:group_simps)
+ unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:algebra_simps)
text{* Equality of vectors in terms of @{term "op \<bullet>"} products. *}
@@ -912,7 +912,7 @@
assume ?rhs
then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" by simp
hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" by (simp add: inner_simps inner_commute)
- then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps inner_simps inner_commute)
+ then have "(x - y) \<bullet> (x - y) = 0" by (simp add: field_simps inner_simps inner_commute)
then show "x = y" by (simp)
qed
@@ -933,7 +933,7 @@
by (rule order_trans [OF norm_triangle_ineq add_mono])
lemma ge_iff_diff_ge_0: "(a::'a::linordered_ring) \<ge> b == a - b \<ge> 0"
- by (simp add: ring_simps)
+ by (simp add: field_simps)
lemma pth_1:
fixes x :: "'a::real_normed_vector"
@@ -1433,15 +1433,15 @@
shows "linear f" using assms unfolding linear_def by auto
lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
- by (vector linear_def Cart_eq ring_simps)
+ by (vector linear_def Cart_eq field_simps)
lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"
- by (vector linear_def Cart_eq ring_simps)
+ by (vector linear_def Cart_eq field_simps)
lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"
- by (vector linear_def Cart_eq ring_simps)
+ by (vector linear_def Cart_eq field_simps)
lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"
by (simp add: linear_def)
@@ -1463,7 +1463,7 @@
shows "linear (\<lambda>x. f x $ k *s v)"
using lf
apply (auto simp add: linear_def )
- by (vector ring_simps)+
+ by (vector field_simps)+
lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"
unfolding linear_def
@@ -1539,7 +1539,7 @@
unfolding norm_mul
apply (simp only: mult_commute)
apply (rule mult_mono)
- by (auto simp add: ring_simps) }
+ by (auto simp add: field_simps) }
then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]
have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
@@ -1565,7 +1565,7 @@
{fix x::"real ^ 'n"
have "norm (f x) \<le> ?K * norm x"
using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp
- apply (auto simp add: ring_simps split add: abs_split)
+ apply (auto simp add: field_simps split add: abs_split)
apply (erule order_trans, simp)
done
}
@@ -1634,12 +1634,12 @@
lemma bilinear_lzero:
fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
using bilinear_ladd[OF bh, of 0 0 x]
- by (simp add: eq_add_iff ring_simps)
+ by (simp add: eq_add_iff field_simps)
lemma bilinear_rzero:
fixes h :: "'a::ring^_ \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
using bilinear_radd[OF bh, of x 0 0 ]
- by (simp add: eq_add_iff ring_simps)
+ by (simp add: eq_add_iff field_simps)
lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ _)) z = h x z - h y z"
by (simp add: diff_def bilinear_ladd bilinear_lneg)
@@ -1680,7 +1680,7 @@
apply (rule real_setsum_norm_le)
using fN fM
apply simp
- apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] ring_simps)
+ apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps)
apply (rule mult_mono)
apply (auto simp add: zero_le_mult_iff component_le_norm)
apply (rule mult_mono)
@@ -1770,7 +1770,7 @@
by (simp add: linear_cmul[OF lf])
finally have "f x \<bullet> y = x \<bullet> ?w"
apply (simp only: )
- apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
+ apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps)
done}
}
then show ?thesis unfolding adjoint_def
@@ -1835,7 +1835,7 @@
lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
- by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
+ by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
lemma matrix_mul_lid:
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
@@ -1954,7 +1954,7 @@
where "matrix f = (\<chi> i j. (f(basis j))$i)"
lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ _))"
- by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
+ by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf)
lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n)"
apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
@@ -2008,7 +2008,7 @@
proof-
have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
have "a = a * (u + v)" unfolding uv by simp
- hence th: "u * a + v * a = a" by (simp add: ring_simps)
+ hence th: "u * a + v * a = a" by (simp add: field_simps)
from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_strict_left_mono)
from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_strict_left_mono)
from xa ya u v have "u * x + v * y < u * a + v * a"
@@ -2031,7 +2031,7 @@
shows "u * x + v * y \<le> a"
proof-
from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
- also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)
+ also have "\<dots> \<le> (u + v) * a" by (simp add: field_simps)
finally show ?thesis unfolding uv by simp
qed
@@ -2052,7 +2052,7 @@
shows "x <= y + z"
proof-
have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y by (simp add: mult_nonneg_nonneg)
- with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)
+ with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square field_simps)
from y z have yz: "y + z \<ge> 0" by arith
from power2_le_imp_le[OF th yz] show ?thesis .
qed
@@ -2543,9 +2543,9 @@
from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
using mult_left_mono[OF p Suc.prems] by simp
- finally show ?case by (simp add: real_of_nat_Suc ring_simps)
+ finally show ?case by (simp add: real_of_nat_Suc field_simps)
qed
lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
@@ -2611,10 +2611,10 @@
from geometric_sum[OF x1, of "Suc n", unfolded x1']
have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
unfolding atLeastLessThanSuc_atLeastAtMost
- using x1' apply (auto simp only: field_eq_simps)
- apply (simp add: ring_simps)
+ using x1' apply (auto simp only: field_simps)
+ apply (simp add: field_simps)
done
- then have ?thesis by (simp add: ring_simps) }
+ then have ?thesis by (simp add: field_simps) }
ultimately show ?thesis by metis
qed
@@ -2633,7 +2633,7 @@
from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
then show ?thesis unfolding sum_gp_basic using mn
- by (simp add: ring_simps power_add[symmetric])
+ by (simp add: field_simps power_add[symmetric])
qed
lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
@@ -2646,7 +2646,7 @@
{assume x: "x = 1" hence ?thesis by simp}
moreover
{assume x: "x \<noteq> 1" hence nz: "1 - x \<noteq> 0" by simp
- from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_eq_simps)}
+ from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)}
ultimately have ?thesis by metis
}
ultimately show ?thesis by metis
@@ -2655,7 +2655,7 @@
lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
(if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
unfolding sum_gp[of x m "m + n"] power_Suc
- by (simp add: ring_simps power_add)
+ by (simp add: field_simps power_add)
subsection{* A bit of linear algebra. *}
@@ -2929,14 +2929,14 @@
apply (simp only: )
apply (rule span_add[unfolded mem_def])
apply assumption+
- apply (vector ring_simps)
+ apply (vector field_simps)
apply (clarsimp simp add: mem_def)
apply (rule_tac x= "c*k" in exI)
apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")
apply (simp only: )
apply (rule span_mul[unfolded mem_def])
apply assumption
- by (vector ring_simps)
+ by (vector field_simps)
ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
qed
@@ -3082,7 +3082,7 @@
setsum_clauses(2)[OF fS] cong del: if_weak_cong)
also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
- by (vector ring_simps)
+ by (vector field_simps)
also have "\<dots> = c*s x + y"
by (simp add: add_commute u)
finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
@@ -3119,7 +3119,7 @@
from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"
using fS aS
- apply (simp add: vector_smult_lneg setsum_clauses ring_simps)
+ apply (simp add: vector_smult_lneg setsum_clauses field_simps)
apply (subst (2) ua[symmetric])
apply (rule setsum_cong2)
by auto
@@ -3652,7 +3652,7 @@
from C(1) have fC: "finite ?C" by simp
from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
{fix x k
- have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
+ have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps)
have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
apply (simp only: vector_ssub_ldistrib th0)
apply (rule span_add_eq)
@@ -3872,7 +3872,7 @@
using z .
{fix k assume k: "z - k *s a \<in> span b"
have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
- by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
+ by (simp add: field_simps vector_sadd_rdistrib[symmetric])
from span_sub[OF th0 k]
have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
{assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
@@ -3886,7 +3886,7 @@
let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
{fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
- by (vector ring_simps)
+ by (vector field_simps)
have addh: "?h (x + y) = ?h x + ?h y"
apply (rule conjunct2[OF h, rule_format, symmetric])
apply (rule span_add[OF x y])
@@ -3899,14 +3899,14 @@
moreover
{fix x:: "'a^'n" and c:: 'a assume x: "x \<in> span (insert a b)"
have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
- by (vector ring_simps)
+ by (vector field_simps)
have hc: "?h (c *s x) = c * ?h x"
apply (rule conjunct2[OF h, rule_format, symmetric])
apply (metis span_mul x)
by (metis tha span_mul x conjunct1[OF h])
have "?g (c *s x) = c*s ?g x"
unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
- by (vector ring_simps)}
+ by (vector field_simps)}
moreover
{fix x assume x: "x \<in> (insert a b)"
{assume xa: "x = a"
@@ -4284,7 +4284,7 @@
fix j
have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
- by (simp add: ring_simps)
+ by (simp add: field_simps)
have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
apply (rule setsum_cong[OF refl])
@@ -4627,7 +4627,7 @@
from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
"infnorm y \<le> infnorm (x - y) + infnorm x"
- by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
+ by (simp_all add: field_simps infnorm_neg diff_def[symmetric])
from th[OF ths] show ?thesis .
qed
@@ -4726,9 +4726,9 @@
using x y
unfolding inner_simps smult_conv_scaleR
unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: inner_commute)
- apply (simp add: ring_simps) by metis
+ apply (simp add: field_simps) by metis
also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
- by (simp add: ring_simps inner_commute)
+ by (simp add: field_simps inner_commute)
also have "\<dots> \<longleftrightarrow> ?lhs" using x y
apply simp
by metis
@@ -4774,7 +4774,7 @@
also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"
unfolding norm_cauchy_schwarz_eq[symmetric]
unfolding power2_norm_eq_inner inner_simps
- by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute ring_simps)
+ by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
finally have ?thesis .}
ultimately show ?thesis by blast
qed
@@ -4860,10 +4860,10 @@
unfolding vector_smult_assoc
unfolding norm_mul
apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
apply simp
apply simp
done
--- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1131,7 +1131,7 @@
guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this
guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this
let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
- using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute)
+ using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:algebra_simps norm_minus_commute)
also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto
finally show False by auto
@@ -1244,7 +1244,7 @@
unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,THEN sym]
by(rule setsum_cong2,auto)
have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
- unfolding * by(auto simp add:group_simps) also let ?res = "\<dots>"
+ unfolding * by(auto simp add:algebra_simps) also let ?res = "\<dots>"
from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto
have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq])
apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto
@@ -1268,7 +1268,7 @@
lemma has_integral_sub:
shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
- using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto
+ using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding algebra_simps by auto
lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
by(rule integral_unique has_integral_0)+
@@ -1703,7 +1703,7 @@
proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
- using p using assms by(auto simp add:group_simps)
+ using p using assms by(auto simp add:algebra_simps)
qed qed
show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
@@ -1711,7 +1711,7 @@
proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
- using p using assms by(auto simp add:group_simps) qed qed qed qed
+ using p using assms by(auto simp add:algebra_simps) qed qed qed qed
subsection {* Generalized notion of additivity. *}
@@ -1847,7 +1847,7 @@
lemma monoidal_monoid[intro]:
shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
- unfolding monoidal_def neutral_monoid by(auto simp add: group_simps)
+ unfolding monoidal_def neutral_monoid by(auto simp add: algebra_simps)
lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
@@ -2380,8 +2380,8 @@
have lem2:"\<And>s1 s2 i1 i2. norm(s2 - s1) \<le> e/2 \<Longrightarrow> norm(s1 - i1) < e / 4 \<Longrightarrow> norm(s2 - i2) < e / 4 \<Longrightarrow>norm(i1 - i2) < e"
proof- case goal1 have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
- using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps)
- also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+ using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:algebra_simps)
+ also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
finally show ?case .
qed
show ?case unfolding vector_dist_norm apply(rule lem2) defer
@@ -2398,7 +2398,7 @@
also have "\<dots> = 2 / real M" unfolding real_divide_def by auto
finally show "norm (g n x - g m x) \<le> 2 / real M"
using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
- by(auto simp add:group_simps simp add:norm_minus_commute)
+ by(auto simp add:algebra_simps simp add:norm_minus_commute)
qed qed qed
from this[unfolded convergent_eq_cauchy[THEN sym]] guess s .. note s=this
@@ -2412,8 +2412,8 @@
have lem:"\<And>sf sg i. norm(sf - sg) \<le> e / 3 \<Longrightarrow> norm(i - s) < e / 3 \<Longrightarrow> norm(sg - i) < e / 3 \<Longrightarrow> norm(sf - s) < e"
proof- case goal1 have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
using norm_triangle_ineq[of "sf - sg" "sg - s"]
- using norm_triangle_ineq[of "sg - i" " i - s"] by(auto simp add:group_simps)
- also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+ using norm_triangle_ineq[of "sg - i" " i - s"] by(auto simp add:algebra_simps)
+ also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
finally show ?case .
qed
show ?case apply(rule_tac x=g' in exI) apply(rule,rule g')
@@ -2955,7 +2955,7 @@
have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
- unfolding scaleR.diff_left by(auto simp add:group_simps)
+ unfolding scaleR.diff_left by(auto simp add:algebra_simps)
also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
@@ -3097,7 +3097,7 @@
proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
apply(rule continuous_on_o_dest_vec1 assms)+ unfolding not_less using assms(2) goal1 by auto
- hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+ hence *:"?I a y - ?I a x = ?I x y" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
using False unfolding not_less using assms(2) goal1 by auto
have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
@@ -3111,7 +3111,7 @@
qed(insert e,auto)
next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
apply(rule continuous_on_o_dest_vec1 assms)+ unfolding not_less using assms(2) goal1 by auto
- hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+ hence *:"?I a x - ?I a y = ?I y x" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
using True using assms(2) goal1 by auto
have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto
@@ -3193,7 +3193,7 @@
apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI)
using X(2) assms(3)[rule_format,of x] by auto
qed note ** = d(2)[OF this] have *:"inj_on (\<lambda>(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp
- have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel
+ have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding algebra_simps add_left_cancel
unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv
apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto
also have "... = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym]
@@ -3331,7 +3331,7 @@
lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
apply(subst(asm)(2) norm_minus_cancel[THEN sym])
- apply(drule norm_triangle_le) by(auto simp add:group_simps)
+ apply(drule norm_triangle_le) by(auto simp add:algebra_simps)
lemma fundamental_theorem_of_calculus_interior:
assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
@@ -3640,11 +3640,11 @@
proof safe show "0 < ?d" using d(1) assms(3) unfolding Cart_simps by auto
fix t::"_^1" assume as:"c \<le> t" "t$1 < c$1 + ?d"
have *:"integral{a..c} f = integral{a..b} f - integral{c..b} f"
- "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding group_simps
+ "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding algebra_simps
apply(rule_tac[!] integral_combine) using assms as unfolding Cart_simps by auto
have "(- c)$1 - d < (- t)$1 \<and> - t \<le> - c" using as by auto note d(2)[rule_format,OF this]
thus "norm (integral {a..c} f - integral {a..t} f) < e" unfolding *
- unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:group_simps) qed qed
+ unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:algebra_simps) qed qed
declare dest_vec1_eq[simp del] not_less[simp] not_le[simp]
@@ -3714,7 +3714,7 @@
apply safe apply(rule conv) using assms(4,7) by auto
have *:"\<And>t xa. (1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x \<Longrightarrow> t = xa"
proof- case goal1 hence "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c"
- unfolding scaleR_simps by(auto simp add:group_simps)
+ unfolding scaleR_simps by(auto simp add:algebra_simps)
thus ?case using `x\<noteq>c` by auto qed
have as2:"finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}" using assms(2)
apply(rule finite_surj[where f="\<lambda>z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"])
@@ -4389,7 +4389,7 @@
have *:"\<And>ir ip i cr cp. norm((cp + cr) - i) < e \<Longrightarrow> norm(cr - ir) < k \<Longrightarrow>
ip + ir = i \<Longrightarrow> norm(cp - ip) \<le> e + k"
proof- case goal1 thus ?case using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]
- unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:group_simps) qed
+ unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:algebra_simps) qed
have "?x = norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
unfolding split_def setsum_subtractf ..
@@ -4500,7 +4500,7 @@
norm(c - d) < e / 4 \<longrightarrow> norm(a - d) < e"
proof safe case goal1 thus ?case using norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
norm_triangle_lt[of "a - b + (b - c)" "c - d" e] unfolding norm_minus_cancel
- by(auto simp add:group_simps) qed
+ by(auto simp add:algebra_simps) qed
show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - i) < e" apply(rule *[rule_format,where
b="\<Sum>(x, k)\<in>p. content k *\<^sub>R f (m x) x" and c="\<Sum>(x, k)\<in>p. integral k (f (m x))"])
proof safe case goal1
@@ -5151,7 +5151,7 @@
assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
shows "(\<lambda>x. f(x) - g(x)) absolutely_integrable_on s"
using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
- unfolding group_simps .
+ unfolding algebra_simps .
lemma absolutely_integrable_linear: fixes f::"real^'m \<Rightarrow> real^'n" and h::"real^'n \<Rightarrow> real^'p"
assumes "f absolutely_integrable_on s" "bounded_linear h"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 26 09:45:22 2010 -0700
@@ -4499,7 +4499,7 @@
let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
{ fix x y assume "x \<in> s" "y \<in> s"
- hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps) }
+ hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: field_simps) }
note * = this
{ fix x y assume "x\<in>s" "y\<in>s" hence "s \<noteq> {}" by auto
have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s`
@@ -5811,7 +5811,7 @@
shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
proof
assume h: "m *s x + c = y"
- hence "m *s x = y - c" by (simp add: ring_simps)
+ hence "m *s x = y - c" by (simp add: field_simps)
hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
then show "x = inverse m *s y + - (inverse m *s c)"
using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
@@ -5913,11 +5913,11 @@
also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
using cf_z[of "m + k"] and c by auto
also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
- using Suc by (auto simp add: ring_simps)
+ using Suc by (auto simp add: field_simps)
also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
- unfolding power_add by (auto simp add: ring_simps)
+ unfolding power_add by (auto simp add: field_simps)
also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
- using c by (auto simp add: ring_simps)
+ using c by (auto simp add: field_simps)
finally show ?case by auto
qed
} note cf_z2 = this
@@ -6074,7 +6074,7 @@
apply(erule_tac x="Na+Nb+n" in allE) apply simp
using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
"-b" "- f (r (Na + Nb + n)) y"]
- unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+ unfolding ** by (auto simp add: algebra_simps dist_commute)
moreover
have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
--- a/src/HOL/NSA/HyperDef.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/NSA/HyperDef.thy Mon Apr 26 09:45:22 2010 -0700
@@ -140,12 +140,12 @@
lemma of_hypreal_inverse [simp]:
"\<And>x. of_hypreal (inverse x) =
- inverse (of_hypreal x :: 'a::{real_div_algebra,division_by_zero} star)"
+ inverse (of_hypreal x :: 'a::{real_div_algebra,division_ring_inverse_zero} star)"
by transfer (rule of_real_inverse)
lemma of_hypreal_divide [simp]:
"\<And>x y. of_hypreal (x / y) =
- (of_hypreal x / of_hypreal y :: 'a::{real_field,division_by_zero} star)"
+ (of_hypreal x / of_hypreal y :: 'a::{real_field,division_ring_inverse_zero} star)"
by transfer (rule of_real_divide)
lemma of_hypreal_eq_iff [simp]:
@@ -454,7 +454,7 @@
by transfer (rule field_power_not_zero)
lemma hyperpow_inverse:
- "\<And>r n. r \<noteq> (0::'a::{division_by_zero,field} star)
+ "\<And>r n. r \<noteq> (0::'a::{division_ring_inverse_zero,field} star)
\<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
by transfer (rule power_inverse)
--- a/src/HOL/NSA/NSA.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/NSA/NSA.thy Mon Apr 26 09:45:22 2010 -0700
@@ -145,12 +145,12 @@
by transfer (rule nonzero_norm_inverse)
lemma hnorm_inverse:
- "\<And>a::'a::{real_normed_div_algebra,division_by_zero} star.
+ "\<And>a::'a::{real_normed_div_algebra,division_ring_inverse_zero} star.
hnorm (inverse a) = inverse (hnorm a)"
by transfer (rule norm_inverse)
lemma hnorm_divide:
- "\<And>a b::'a::{real_normed_field,division_by_zero} star.
+ "\<And>a b::'a::{real_normed_field,division_ring_inverse_zero} star.
hnorm (a / b) = hnorm a / hnorm b"
by transfer (rule norm_divide)
--- a/src/HOL/NSA/StarDef.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/NSA/StarDef.thy Mon Apr 26 09:45:22 2010 -0700
@@ -902,7 +902,7 @@
apply (transfer, rule divide_inverse)
done
-instance star :: (division_by_zero) division_by_zero
+instance star :: (division_ring_inverse_zero) division_ring_inverse_zero
by (intro_classes, transfer, rule inverse_zero)
instance star :: (ordered_semiring) ordered_semiring
--- a/src/HOL/Nominal/nominal_inductive.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Nominal/nominal_inductive.ML Mon Apr 26 09:45:22 2010 -0700
@@ -543,7 +543,7 @@
in
ctxt'' |>
- Proof.theorem_i NONE (fn thss => fn ctxt =>
+ Proof.theorem NONE (fn thss => fn ctxt =>
let
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
--- a/src/HOL/Nominal/nominal_inductive2.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Nominal/nominal_inductive2.ML Mon Apr 26 09:45:22 2010 -0700
@@ -445,7 +445,7 @@
in
ctxt'' |>
- Proof.theorem_i NONE (fn thss => fn ctxt =>
+ Proof.theorem NONE (fn thss => fn ctxt =>
let
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
--- a/src/HOL/Nominal/nominal_primrec.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Nominal/nominal_primrec.ML Mon Apr 26 09:45:22 2010 -0700
@@ -363,7 +363,7 @@
in
lthy' |>
Variable.add_fixes (map fst lsrs) |> snd |>
- Proof.theorem_i NONE
+ Proof.theorem NONE
(fn thss => fn goal_ctxt =>
let
val simps = ProofContext.export goal_ctxt lthy' (flat thss);
--- a/src/HOL/Number_Theory/Binomial.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Number_Theory/Binomial.thy Mon Apr 26 09:45:22 2010 -0700
@@ -208,7 +208,7 @@
have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
fact (k + 1) * fact (n - k) * (n choose (k + 1)) +
fact (k + 1) * fact (n - k) * (n choose k)"
- by (subst choose_reduce_nat, auto simp add: ring_simps)
+ by (subst choose_reduce_nat, auto simp add: field_simps)
also note one
also note two
also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)"
@@ -279,7 +279,7 @@
also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
(SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le
- power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
+ power_Suc field_simps One_nat_def del:setsum_cl_ivl_Suc)
also have "... = a^(n+1) + b^(n+1) +
(SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
(SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
@@ -287,10 +287,10 @@
also have
"... = a^(n+1) + b^(n+1) +
(SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
- by (auto simp add: ring_simps setsum_addf [symmetric]
+ by (auto simp add: field_simps setsum_addf [symmetric]
choose_reduce_nat)
also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
- using decomp by (simp add: ring_simps)
+ using decomp by (simp add: field_simps)
finally show "?P (n + 1)" by simp
qed
--- a/src/HOL/Number_Theory/Cong.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Number_Theory/Cong.thy Mon Apr 26 09:45:22 2010 -0700
@@ -350,7 +350,7 @@
apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
(* any way around this? *)
apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
- apply (auto simp add: ring_simps)
+ apply (auto simp add: field_simps)
done
lemma cong_mult_rcancel_int:
@@ -416,7 +416,7 @@
done
lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
- apply (auto simp add: cong_altdef_int dvd_def ring_simps)
+ apply (auto simp add: cong_altdef_int dvd_def field_simps)
apply (rule_tac [!] x = "-k" in exI, auto)
done
@@ -428,14 +428,14 @@
apply (unfold dvd_def, auto)
apply (rule_tac x = k in exI)
apply (rule_tac x = 0 in exI)
- apply (auto simp add: ring_simps)
+ apply (auto simp add: field_simps)
apply (subst (asm) cong_sym_eq_nat)
apply (subst (asm) cong_altdef_nat)
apply force
apply (unfold dvd_def, auto)
apply (rule_tac x = 0 in exI)
apply (rule_tac x = k in exI)
- apply (auto simp add: ring_simps)
+ apply (auto simp add: field_simps)
apply (unfold cong_nat_def)
apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
apply (erule ssubst)back
@@ -533,7 +533,7 @@
apply (auto simp add: cong_iff_lin_nat dvd_def)
apply (rule_tac x="k1 * k" in exI)
apply (rule_tac x="k2 * k" in exI)
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
done
lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
@@ -559,7 +559,7 @@
lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
apply (simp add: cong_altdef_int)
apply (subst dvd_minus_iff [symmetric])
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
done
lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
@@ -603,7 +603,7 @@
apply (unfold dvd_def)
apply auto [1]
apply (rule_tac x = k in exI)
- apply (auto simp add: ring_simps) [1]
+ apply (auto simp add: field_simps) [1]
apply (subst cong_altdef_nat)
apply (auto simp add: dvd_def)
done
@@ -611,7 +611,7 @@
lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
apply (subst cong_altdef_nat)
apply assumption
- apply (unfold dvd_def, auto simp add: ring_simps)
+ apply (unfold dvd_def, auto simp add: field_simps)
apply (rule_tac x = k in exI)
apply auto
done
--- a/src/HOL/Number_Theory/Fib.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Number_Theory/Fib.thy Mon Apr 26 09:45:22 2010 -0700
@@ -143,9 +143,9 @@
apply (induct n rule: fib_induct_nat)
apply auto
apply (subst fib_reduce_nat)
- apply (auto simp add: ring_simps)
+ apply (auto simp add: field_simps)
apply (subst (1 3 5) fib_reduce_nat)
- apply (auto simp add: ring_simps Suc_eq_plus1)
+ apply (auto simp add: field_simps Suc_eq_plus1)
(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
apply (erule ssubst) back back
@@ -184,7 +184,7 @@
lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) -
(fib (int n + 1))^2 = (-1)^(n + 1)"
apply (induct n)
- apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
+ apply (auto simp add: field_simps power2_eq_square fib_reduce_int
power_add)
done
@@ -222,7 +222,7 @@
apply (subst (2) fib_reduce_nat)
apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
apply (subst add_commute, auto)
- apply (subst gcd_commute_nat, auto simp add: ring_simps)
+ apply (subst gcd_commute_nat, auto simp add: field_simps)
done
lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
@@ -305,7 +305,7 @@
theorem fib_mult_eq_setsum_nat:
"fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
apply (induct n)
- apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
+ apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat field_simps)
done
theorem fib_mult_eq_setsum'_nat:
--- a/src/HOL/Number_Theory/Residues.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Number_Theory/Residues.thy Mon Apr 26 09:45:22 2010 -0700
@@ -69,7 +69,7 @@
apply (subst mod_add_eq [symmetric])
apply (subst mult_commute)
apply (subst zmod_zmult1_eq [symmetric])
- apply (simp add: ring_simps)
+ apply (simp add: field_simps)
done
end
--- a/src/HOL/Power.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Power.thy Mon Apr 26 09:45:22 2010 -0700
@@ -400,7 +400,7 @@
text{*Perhaps these should be simprules.*}
lemma power_inverse:
- fixes a :: "'a::{division_ring,division_by_zero,power}"
+ fixes a :: "'a::{division_ring,division_ring_inverse_zero,power}"
shows "inverse (a ^ n) = (inverse a) ^ n"
apply (cases "a = 0")
apply (simp add: power_0_left)
@@ -408,11 +408,11 @@
done (* TODO: reorient or rename to inverse_power *)
lemma power_one_over:
- "1 / (a::'a::{field,division_by_zero, power}) ^ n = (1 / a) ^ n"
+ "1 / (a::'a::{field,division_ring_inverse_zero, power}) ^ n = (1 / a) ^ n"
by (simp add: divide_inverse) (rule power_inverse)
lemma power_divide:
- "(a / b) ^ n = (a::'a::{field,division_by_zero}) ^ n / b ^ n"
+ "(a / b) ^ n = (a::'a::{field,division_ring_inverse_zero}) ^ n / b ^ n"
apply (cases "b = 0")
apply (simp add: power_0_left)
apply (rule nonzero_power_divide)
--- a/src/HOL/Quotient_Examples/FSet.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Quotient_Examples/FSet.thy Mon Apr 26 09:45:22 2010 -0700
@@ -387,7 +387,7 @@
apply (simp add: memb_def[symmetric] memb_finter_raw)
by (auto simp add: memb_def)
-instantiation fset :: (type) "{bot,distrib_lattice}"
+instantiation fset :: (type) "{bounded_lattice_bot,distrib_lattice}"
begin
quotient_definition
@@ -922,14 +922,13 @@
text {* funion *}
-lemma funion_simps[simp]:
- shows "{||} |\<union>| S = S"
- and "finsert x S |\<union>| T = finsert x (S |\<union>| T)"
- by (lifting append.simps)
+lemmas [simp] =
+ sup_bot_left[where 'a="'a fset"]
+ sup_bot_right[where 'a="'a fset"]
-lemma funion_empty[simp]:
- shows "S |\<union>| {||} = S"
- by (lifting append_Nil2)
+lemma funion_finsert[simp]:
+ shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)"
+ by (lifting append.simps(2))
lemma singleton_union_left:
"{|a|} |\<union>| S = finsert a S"
--- a/src/HOL/Rat.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Rat.thy Mon Apr 26 09:45:22 2010 -0700
@@ -444,7 +444,7 @@
end
-instance rat :: division_by_zero proof
+instance rat :: division_ring_inverse_zero proof
qed (simp add: rat_number_expand, simp add: rat_number_collapse)
@@ -818,7 +818,7 @@
done
lemma of_rat_inverse:
- "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
+ "(of_rat (inverse a)::'a::{field_char_0,division_ring_inverse_zero}) =
inverse (of_rat a)"
by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
@@ -827,7 +827,7 @@
by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
lemma of_rat_divide:
- "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
+ "(of_rat (a / b)::'a::{field_char_0,division_ring_inverse_zero})
= of_rat a / of_rat b"
by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
@@ -968,7 +968,7 @@
done
lemma Rats_inverse [simp]:
- fixes a :: "'a::{field_char_0,division_by_zero}"
+ fixes a :: "'a::{field_char_0,division_ring_inverse_zero}"
shows "a \<in> Rats \<Longrightarrow> inverse a \<in> Rats"
apply (auto simp add: Rats_def)
apply (rule range_eqI)
@@ -984,7 +984,7 @@
done
lemma Rats_divide [simp]:
- fixes a b :: "'a::{field_char_0,division_by_zero}"
+ fixes a b :: "'a::{field_char_0,division_ring_inverse_zero}"
shows "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
apply (auto simp add: Rats_def)
apply (rule range_eqI)
--- a/src/HOL/RealDef.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/RealDef.thy Mon Apr 26 09:45:22 2010 -0700
@@ -279,7 +279,7 @@
lemma INVERSE_ZERO: "inverse 0 = (0::real)"
by (simp add: real_inverse_def)
-instance real :: division_by_zero
+instance real :: division_ring_inverse_zero
proof
show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
qed
--- a/src/HOL/RealVector.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/RealVector.thy Mon Apr 26 09:45:22 2010 -0700
@@ -207,7 +207,7 @@
by (rule inverse_unique, simp)
lemma inverse_scaleR_distrib:
- fixes x :: "'a::{real_div_algebra,division_by_zero}"
+ fixes x :: "'a::{real_div_algebra,division_ring_inverse_zero}"
shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
apply (case_tac "a = 0", simp)
apply (case_tac "x = 0", simp)
@@ -250,7 +250,7 @@
lemma of_real_inverse [simp]:
"of_real (inverse x) =
- inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
+ inverse (of_real x :: 'a::{real_div_algebra,division_ring_inverse_zero})"
by (simp add: of_real_def inverse_scaleR_distrib)
lemma nonzero_of_real_divide:
@@ -260,7 +260,7 @@
lemma of_real_divide [simp]:
"of_real (x / y) =
- (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
+ (of_real x / of_real y :: 'a::{real_field,division_ring_inverse_zero})"
by (simp add: divide_inverse)
lemma of_real_power [simp]:
@@ -370,7 +370,7 @@
done
lemma Reals_inverse [simp]:
- fixes a :: "'a::{real_div_algebra,division_by_zero}"
+ fixes a :: "'a::{real_div_algebra,division_ring_inverse_zero}"
shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
apply (auto simp add: Reals_def)
apply (rule range_eqI)
@@ -386,7 +386,7 @@
done
lemma Reals_divide [simp]:
- fixes a b :: "'a::{real_field,division_by_zero}"
+ fixes a b :: "'a::{real_field,division_ring_inverse_zero}"
shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
apply (auto simp add: Reals_def)
apply (rule range_eqI)
@@ -726,7 +726,7 @@
done
lemma norm_inverse:
- fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
+ fixes a :: "'a::{real_normed_div_algebra,division_ring_inverse_zero}"
shows "norm (inverse a) = inverse (norm a)"
apply (case_tac "a = 0", simp)
apply (erule nonzero_norm_inverse)
@@ -738,7 +738,7 @@
by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
lemma norm_divide:
- fixes a b :: "'a::{real_normed_field,division_by_zero}"
+ fixes a b :: "'a::{real_normed_field,division_ring_inverse_zero}"
shows "norm (a / b) = norm a / norm b"
by (simp add: divide_inverse norm_mult norm_inverse)
--- a/src/HOL/Rings.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Rings.thy Mon Apr 26 09:45:22 2010 -0700
@@ -14,8 +14,8 @@
begin
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"
+ assumes left_distrib[algebra_simps, field_simps]: "(a + b) * c = a * c + b * c"
+ assumes right_distrib[algebra_simps, field_simps]: "a * (b + c) = a * b + a * c"
begin
text{*For the @{text combine_numerals} simproc*}
@@ -230,18 +230,15 @@
lemma minus_mult_commute: "- a * b = a * - b"
by simp
-lemma right_diff_distrib[algebra_simps]: "a * (b - c) = a * b - a * c"
+lemma right_diff_distrib[algebra_simps, field_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"
+lemma left_diff_distrib[algebra_simps, field_simps]: "(a - b) * c = a * c - b * c"
by (simp add: left_distrib diff_minus)
lemmas ring_distribs[no_atp] =
right_distrib left_distrib left_diff_distrib right_diff_distrib
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
lemma eq_add_iff1:
"a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
by (simp add: algebra_simps)
@@ -536,7 +533,7 @@
lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
by (simp add: diff_minus add_divide_distrib)
-lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
+lemma nonzero_eq_divide_eq [field_simps]: "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
@@ -544,7 +541,7 @@
finally show ?thesis .
qed
-lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
+lemma nonzero_divide_eq_eq [field_simps]: "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
@@ -560,7 +557,7 @@
end
-class division_by_zero = division_ring +
+class division_ring_inverse_zero = division_ring +
assumes inverse_zero [simp]: "inverse 0 = 0"
begin
@@ -861,9 +858,6 @@
subclass ordered_ab_group_add ..
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
lemma less_add_iff1:
"a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
by (simp add: algebra_simps)
@@ -1068,9 +1062,6 @@
end
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
lemmas mult_sign_intros =
mult_nonneg_nonneg mult_nonneg_nonpos
mult_nonpos_nonneg mult_nonpos_nonpos
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML Mon Apr 26 09:45:22 2010 -0700
@@ -1137,7 +1137,8 @@
handle THM _ => NONE
in
val z3_simpset = HOL_ss addsimps @{thms array_rules}
- addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps}
+ addsimps @{thms ring_distribs} addsimps @{thms field_simps}
+ addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}]
addsimps @{thms arith_special} addsimps @{thms less_bin_simps}
addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps}
addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps}
--- a/src/HOL/SMT/Z3.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/SMT/Z3.thy Mon Apr 26 09:45:22 2010 -0700
@@ -19,7 +19,7 @@
lemmas [z3_rewrite] =
refl eq_commute conj_commute disj_commute simp_thms nnf_simps
- ring_distribs field_eq_simps if_True if_False
+ ring_distribs field_simps times_divide_eq_right times_divide_eq_left if_True if_False
lemma [z3_rewrite]: "(P \<noteq> Q) = (Q = (\<not>P))" by fast
--- a/src/HOL/Series.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Series.thy Mon Apr 26 09:45:22 2010 -0700
@@ -381,7 +381,7 @@
shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
by (rule geometric_sums [THEN sums_summable])
-lemma half: "0 < 1 / (2::'a::{number_ring,division_by_zero,linordered_field})"
+lemma half: "0 < 1 / (2::'a::{number_ring,division_ring_inverse_zero,linordered_field})"
by arith
lemma power_half_series: "(\<lambda>n. (1/2::real)^Suc n) sums 1"
--- a/src/HOL/SetInterval.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/SetInterval.thy Mon Apr 26 09:45:22 2010 -0700
@@ -1095,7 +1095,7 @@
next
case (Suc n)
moreover with `y \<noteq> 0` have "(1 + y) ^ n = (y * inverse y) * (1 + y) ^ n" by simp
- ultimately show ?case by (simp add: field_eq_simps divide_inverse)
+ ultimately show ?case by (simp add: field_simps divide_inverse)
qed
ultimately show ?thesis by simp
qed
--- a/src/HOL/Tools/Datatype/datatype_data.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Mon Apr 26 09:45:22 2010 -0700
@@ -436,7 +436,7 @@
in
thy
|> ProofContext.init
- |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+ |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
end;
val rep_datatype = gen_rep_datatype Sign.cert_term;
--- a/src/HOL/Tools/Function/function.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/Function/function.ML Mon Apr 26 09:45:22 2010 -0700
@@ -120,7 +120,7 @@
end
in
lthy
- |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+ |> Proof.theorem NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
|> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
end
@@ -177,7 +177,7 @@
|> ProofContext.note_thmss ""
[((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
[([Goal.norm_result termination], [])])] |> snd
- |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+ |> Proof.theorem NONE afterqed [[(goal, [])]]
end
val termination_proof = gen_termination_proof Syntax.check_term
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Apr 26 09:45:22 2010 -0700
@@ -3059,7 +3059,7 @@
) options [const]))
end
in
- Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
+ Proof.theorem NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
end;
val code_pred = generic_code_pred (K I);
--- a/src/HOL/Tools/Quotient/quotient_typ.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML Mon Apr 26 09:45:22 2010 -0700
@@ -45,7 +45,7 @@
val goals' = map (rpair []) goals
fun after_qed' thms = after_qed (the_single thms)
in
- Proof.theorem_i NONE after_qed' [goals'] ctxt
+ Proof.theorem NONE after_qed' [goals'] ctxt
end
--- a/src/HOL/Tools/choice_specification.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/choice_specification.ML Mon Apr 26 09:45:22 2010 -0700
@@ -226,7 +226,7 @@
in
thy
|> ProofContext.init
- |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
+ |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
end;
--- a/src/HOL/Tools/numeral_simprocs.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/numeral_simprocs.ML Mon Apr 26 09:45:22 2010 -0700
@@ -332,8 +332,8 @@
val field_combine_numerals =
Arith_Data.prep_simproc @{theory}
("field_combine_numerals",
- ["(i::'a::{number_ring,field,division_by_zero}) + j",
- "(i::'a::{number_ring,field,division_by_zero}) - j"],
+ ["(i::'a::{number_ring,field,division_ring_inverse_zero}) + j",
+ "(i::'a::{number_ring,field,division_ring_inverse_zero}) - j"],
K FieldCombineNumerals.proc);
(** Constant folding for multiplication in semirings **)
@@ -442,9 +442,9 @@
"(l::'a::{semiring_div,number_ring}) div (m * n)"],
K DivCancelNumeralFactor.proc),
("divide_cancel_numeral_factor",
- ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
- "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
- "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+ ["((l::'a::{division_ring_inverse_zero,field,number_ring}) * m) / n",
+ "(l::'a::{division_ring_inverse_zero,field,number_ring}) / (m * n)",
+ "((number_of v)::'a::{division_ring_inverse_zero,field,number_ring}) / (number_of w)"],
K DivideCancelNumeralFactor.proc)];
val field_cancel_numeral_factors =
@@ -454,9 +454,9 @@
"(l::'a::{field,number_ring}) = m * n"],
K EqCancelNumeralFactor.proc),
("field_cancel_numeral_factor",
- ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
- "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
- "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+ ["((l::'a::{division_ring_inverse_zero,field,number_ring}) * m) / n",
+ "(l::'a::{division_ring_inverse_zero,field,number_ring}) / (m * n)",
+ "((number_of v)::'a::{division_ring_inverse_zero,field,number_ring}) / (number_of w)"],
K DivideCancelNumeralFactor.proc)]
@@ -598,8 +598,8 @@
["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"],
K DvdCancelFactor.proc),
("divide_cancel_factor",
- ["((l::'a::{division_by_zero,field}) * m) / n",
- "(l::'a::{division_by_zero,field}) / (m * n)"],
+ ["((l::'a::{division_ring_inverse_zero,field}) * m) / n",
+ "(l::'a::{division_ring_inverse_zero,field}) / (m * n)"],
K DivideCancelFactor.proc)];
end;
--- a/src/HOL/Tools/typedef.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/Tools/typedef.ML Mon Apr 26 09:45:22 2010 -0700
@@ -282,7 +282,7 @@
val ((goal, goal_pat, typedef_result), lthy') =
prepare_typedef prep_term def name (b, args, mx) set opt_morphs lthy;
fun after_qed [[th]] = snd o typedef_result th;
- in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] lthy' end;
+ in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end;
in
--- a/src/HOL/ex/Lagrange.thy Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOL/ex/Lagrange.thy Mon Apr 26 09:45:22 2010 -0700
@@ -34,7 +34,7 @@
sq (x1*y2 + x2*y1 + x3*y4 - x4*y3) +
sq (x1*y3 - x2*y4 + x3*y1 + x4*y2) +
sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
@@ -50,6 +50,6 @@
sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
end
--- a/src/HOLCF/Tools/pcpodef.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/HOLCF/Tools/pcpodef.ML Mon Apr 26 09:45:22 2010 -0700
@@ -328,7 +328,7 @@
prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy;
fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
| after_qed _ = raise Fail "cpodef_proof";
- in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
+ in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
fun gen_pcpodef_proof prep_term prep_constraint
((def, name), (b, raw_args, mx), set, opt_morphs) thy =
@@ -339,7 +339,7 @@
prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy;
fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
| after_qed _ = raise Fail "pcpodef_proof";
- in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
+ in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
in
--- a/src/Pure/Isar/calculation.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/calculation.ML Mon Apr 26 09:45:22 2010 -0700
@@ -13,11 +13,11 @@
val sym_add: attribute
val sym_del: attribute
val symmetric: attribute
- val also: (Facts.ref * Attrib.src list) list option -> bool -> Proof.state -> Proof.state Seq.seq
- val also_i: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
- val finally: (Facts.ref * Attrib.src list) list option -> bool ->
+ val also: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
+ val also_cmd: (Facts.ref * Attrib.src list) list option -> bool -> Proof.state -> Proof.state Seq.seq
+ val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
+ val finally_cmd: (Facts.ref * Attrib.src list) list option -> bool ->
Proof.state -> Proof.state Seq.seq
- val finally_i: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
val moreover: bool -> Proof.state -> Proof.state
val ultimately: bool -> Proof.state -> Proof.state
end;
@@ -148,10 +148,10 @@
state |> maintain_calculation final calc))
end;
-val also = calculate Proof.get_thmss false;
-val also_i = calculate (K I) false;
-val finally = calculate Proof.get_thmss true;
-val finally_i = calculate (K I) true;
+val also = calculate (K I) false;
+val also_cmd = calculate Proof.get_thmss_cmd false;
+val finally = calculate (K I) true;
+val finally_cmd = calculate Proof.get_thmss_cmd true;
(* moreover and ultimately *)
--- a/src/Pure/Isar/class_target.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/class_target.ML Mon Apr 26 09:45:22 2010 -0700
@@ -163,7 +163,7 @@
Symtab.empty
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
Symtab.map_default (class, []) (insert (op =) tyco)) arities)
- ((#arities o Sorts.rep_algebra) algebra);
+ (Sorts.arities_of algebra);
val the_arities = these o Symtab.lookup arities;
fun mk_arity class tyco =
let
@@ -370,7 +370,7 @@
in
thy
|> ProofContext.init
- |> Proof.theorem_i NONE after_qed [[(mk_prop thy classrel, [])]]
+ |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]]
end;
in
@@ -539,7 +539,7 @@
end;
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
- Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
+ Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
@@ -595,7 +595,7 @@
in
thy
|> ProofContext.init
- |> Proof.theorem_i NONE after_qed (map (fn t => [(t, [])]) arities)
+ |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities)
end;
--- a/src/Pure/Isar/element.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/element.ML Mon Apr 26 09:45:22 2010 -0700
@@ -283,10 +283,10 @@
in
fun witness_proof after_qed wit_propss =
- gen_witness_proof (Proof.theorem_i NONE) (fn wits => fn _ => after_qed wits)
+ gen_witness_proof (Proof.theorem NONE) (fn wits => fn _ => after_qed wits)
wit_propss [];
-val witness_proof_eqs = gen_witness_proof (Proof.theorem_i NONE);
+val witness_proof_eqs = gen_witness_proof (Proof.theorem NONE);
fun witness_local_proof after_qed cmd wit_propss goal_ctxt int =
gen_witness_proof (fn after_qed' => fn propss =>
--- a/src/Pure/Isar/isar_cmd.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/isar_cmd.ML Mon Apr 26 09:45:22 2010 -0700
@@ -219,10 +219,10 @@
fun goal opt_chain goal stmt int =
opt_chain #> goal NONE (K I) stmt int;
-val have = goal I Proof.have;
-val hence = goal Proof.chain Proof.have;
-val show = goal I Proof.show;
-val thus = goal Proof.chain Proof.show;
+val have = goal I Proof.have_cmd;
+val hence = goal Proof.chain Proof.have_cmd;
+val show = goal I Proof.show_cmd;
+val thus = goal Proof.chain Proof.show_cmd;
(* local endings *)
@@ -393,7 +393,7 @@
let
val thy = Toplevel.theory_of state;
val {classes = (space, algebra), ...} = Type.rep_tsig (Sign.tsig_of thy);
- val {classes, ...} = Sorts.rep_algebra algebra;
+ val classes = Sorts.classes_of algebra;
fun entry (c, (i, (_, cs))) =
(i, {name = Name_Space.extern space c, ID = c, parents = cs,
dir = "", unfold = true, path = ""});
@@ -403,7 +403,7 @@
in Present.display_graph gr end);
fun thm_deps args = Toplevel.unknown_theory o Toplevel.keep (fn state =>
- Thm_Deps.thm_deps (Proof.get_thmss (Toplevel.enter_proof_body state) args));
+ Thm_Deps.thm_deps (Proof.get_thmss_cmd (Toplevel.enter_proof_body state) args));
(* find unused theorems *)
@@ -437,12 +437,12 @@
local
fun string_of_stmts state args =
- Proof.get_thmss state args
+ Proof.get_thmss_cmd state args
|> map (Element.pretty_statement (Proof.context_of state) Thm.theoremK)
|> Pretty.chunks2 |> Pretty.string_of;
fun string_of_thms state args =
- Pretty.string_of (Display.pretty_thms (Proof.context_of state) (Proof.get_thmss state args));
+ Pretty.string_of (Display.pretty_thms (Proof.context_of state) (Proof.get_thmss_cmd state args));
fun string_of_prfs full state arg =
Pretty.string_of
@@ -460,7 +460,7 @@
end
| SOME args => Pretty.chunks
(map (Proof_Syntax.pretty_proof_of (Proof.context_of state) full)
- (Proof.get_thmss state args)));
+ (Proof.get_thmss_cmd state args)));
fun string_of_prop state s =
let
--- a/src/Pure/Isar/isar_syn.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/isar_syn.ML Mon Apr 26 09:45:22 2010 -0700
@@ -542,27 +542,27 @@
val _ =
OuterSyntax.command "from" "forward chaining from given facts"
(K.tag_proof K.prf_chain)
- (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss)));
+ (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd)));
val _ =
OuterSyntax.command "with" "forward chaining from given and current facts"
(K.tag_proof K.prf_chain)
- (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss)));
+ (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd)));
val _ =
OuterSyntax.command "note" "define facts"
(K.tag_proof K.prf_decl)
- (SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss)));
+ (SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));
val _ =
OuterSyntax.command "using" "augment goal facts"
(K.tag_proof K.prf_decl)
- (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using)));
+ (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd)));
val _ =
OuterSyntax.command "unfolding" "unfold definitions in goal and facts"
(K.tag_proof K.prf_decl)
- (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding)));
+ (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd)));
(* proof context *)
@@ -570,17 +570,17 @@
val _ =
OuterSyntax.command "fix" "fix local variables (Skolem constants)"
(K.tag_proof K.prf_asm)
- (P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix)));
+ (P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));
val _ =
OuterSyntax.command "assume" "assume propositions"
(K.tag_proof K.prf_asm)
- (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume)));
+ (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));
val _ =
OuterSyntax.command "presume" "assume propositions, to be established later"
(K.tag_proof K.prf_asm)
- (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume)));
+ (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));
val _ =
OuterSyntax.command "def" "local definition"
@@ -588,24 +588,24 @@
(P.and_list1
(SpecParse.opt_thm_name ":" --
((P.binding -- P.opt_mixfix) -- ((P.$$$ "\\<equiv>" || P.$$$ "==") |-- P.!!! P.termp)))
- >> (Toplevel.print oo (Toplevel.proof o Proof.def)));
+ >> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd)));
val _ =
OuterSyntax.command "obtain" "generalized existence"
(K.tag_proof K.prf_asm_goal)
(P.parname -- Scan.optional (P.fixes --| P.where_) [] -- SpecParse.statement
- >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain x y z)));
+ >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain_cmd x y z)));
val _ =
OuterSyntax.command "guess" "wild guessing (unstructured)"
(K.tag_proof K.prf_asm_goal)
- (Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess)));
+ (Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));
val _ =
OuterSyntax.command "let" "bind text variables"
(K.tag_proof K.prf_decl)
(P.and_list1 (P.and_list1 P.term -- (P.$$$ "=" |-- P.term))
- >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind)));
+ >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));
val case_spec =
(P.$$$ "(" |-- P.!!! (P.xname -- Scan.repeat1 (P.maybe P.name) --| P.$$$ ")") ||
@@ -614,7 +614,7 @@
val _ =
OuterSyntax.command "case" "invoke local context"
(K.tag_proof K.prf_asm)
- (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case)));
+ (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd)));
(* proof structure *)
@@ -711,12 +711,12 @@
val _ =
OuterSyntax.command "also" "combine calculation and current facts"
(K.tag_proof K.prf_decl)
- (calc_args >> (Toplevel.proofs' o Calculation.also));
+ (calc_args >> (Toplevel.proofs' o Calculation.also_cmd));
val _ =
OuterSyntax.command "finally" "combine calculation and current facts, exhibit result"
(K.tag_proof K.prf_chain)
- (calc_args >> (Toplevel.proofs' o Calculation.finally));
+ (calc_args >> (Toplevel.proofs' o Calculation.finally_cmd));
val _ =
OuterSyntax.command "moreover" "augment calculation by current facts"
--- a/src/Pure/Isar/obtain.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/obtain.ML Mon Apr 26 09:45:22 2010 -0700
@@ -39,14 +39,14 @@
signature OBTAIN =
sig
val thatN: string
- val obtain: string -> (binding * string option * mixfix) list ->
+ val obtain: string -> (binding * typ option * mixfix) list ->
+ (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
+ val obtain_cmd: string -> (binding * string option * mixfix) list ->
(Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
- val obtain_i: string -> (binding * typ option * mixfix) list ->
- (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
((string * cterm) list * thm list) * Proof.context
- val guess: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
- val guess_i: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+ val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+ val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
end;
structure Obtain: OBTAIN =
@@ -148,26 +148,26 @@
fun after_qed _ =
Proof.local_qed (NONE, false)
#> `Proof.the_fact #-> (fn rule =>
- Proof.fix_i vars
- #> Proof.assm_i (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
+ Proof.fix vars
+ #> Proof.assm (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
in
state
|> Proof.enter_forward
- |> Proof.have_i NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int
+ |> Proof.have NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int
|> Proof.proof (SOME Method.succeed_text) |> Seq.hd
- |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)]
- |> Proof.assume_i
+ |> Proof.fix [(Binding.name thesisN, NONE, NoSyn)]
+ |> Proof.assume
[((Binding.name that_name, [Context_Rules.intro_query NONE]), [(that_prop, [])])]
|> `Proof.the_facts
||> Proof.chain_facts chain_facts
- ||> Proof.show_i NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false
+ ||> Proof.show NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false
|-> Proof.refine_insert
end;
in
-val obtain = gen_obtain Attrib.attribute ProofContext.read_vars ProofContext.read_propp;
-val obtain_i = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp;
+val obtain = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp;
+val obtain_cmd = gen_obtain Attrib.attribute ProofContext.read_vars ProofContext.read_propp;
end;
@@ -290,8 +290,8 @@
in
state'
|> Proof.map_context (K ctxt')
- |> Proof.fix_i (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
- |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm_i
+ |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
+ |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm
(obtain_export fix_ctxt rule (map cert ts)) [(Thm.empty_binding, asms)])
|> Proof.bind_terms Auto_Bind.no_facts
end;
@@ -307,7 +307,7 @@
state
|> Proof.enter_forward
|> Proof.begin_block
- |> Proof.fix_i [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
+ |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
|> Proof.chain_facts chain_facts
|> Proof.local_goal print_result (K I) (apsnd (rpair I))
"guess" before_qed after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])]
@@ -316,8 +316,8 @@
in
-val guess = gen_guess ProofContext.read_vars;
-val guess_i = gen_guess ProofContext.cert_vars;
+val guess = gen_guess ProofContext.cert_vars;
+val guess_cmd = gen_guess ProofContext.read_vars;
end;
--- a/src/Pure/Isar/proof.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/proof.ML Mon Apr 26 09:45:22 2010 -0700
@@ -41,37 +41,35 @@
val raw_goal: state -> {context: context, facts: thm list, goal: thm}
val goal: state -> {context: context, facts: thm list, goal: thm}
val simple_goal: state -> {context: context, goal: thm}
- val match_bind: (string list * string) list -> state -> state
- val match_bind_i: (term list * term) list -> state -> state
- val let_bind: (string list * string) list -> state -> state
- val let_bind_i: (term list * term) list -> state -> state
- val fix: (binding * string option * mixfix) list -> state -> state
- val fix_i: (binding * typ option * mixfix) list -> state -> state
+ val let_bind: (term list * term) list -> state -> state
+ val let_bind_cmd: (string list * string) list -> state -> state
+ val fix: (binding * typ option * mixfix) list -> state -> state
+ val fix_cmd: (binding * string option * mixfix) list -> state -> state
val assm: Assumption.export ->
+ (Thm.binding * (term * term list) list) list -> state -> state
+ val assm_cmd: Assumption.export ->
(Attrib.binding * (string * string list) list) list -> state -> state
- val assm_i: Assumption.export ->
- (Thm.binding * (term * term list) list) list -> state -> state
- val assume: (Attrib.binding * (string * string list) list) list -> state -> state
- val assume_i: (Thm.binding * (term * term list) list) list -> state -> state
- val presume: (Attrib.binding * (string * string list) list) list -> state -> state
- val presume_i: (Thm.binding * (term * term list) list) list -> state -> state
- val def: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state
- val def_i: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state
+ val assume: (Thm.binding * (term * term list) list) list -> state -> state
+ val assume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
+ val presume: (Thm.binding * (term * term list) list) list -> state -> state
+ val presume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
+ val def: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state
+ val def_cmd: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state
val chain: state -> state
val chain_facts: thm list -> state -> state
- val get_thmss: state -> (Facts.ref * Attrib.src list) list -> thm list
- val note_thmss: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
- val note_thmss_i: (Thm.binding * (thm list * attribute list) list) list -> state -> state
- val from_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
- val from_thmss_i: ((thm list * attribute list) list) list -> state -> state
- val with_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
- val with_thmss_i: ((thm list * attribute list) list) list -> state -> state
- val using: ((Facts.ref * Attrib.src list) list) list -> state -> state
- val using_i: ((thm list * attribute list) list) list -> state -> state
- val unfolding: ((Facts.ref * Attrib.src list) list) list -> state -> state
- val unfolding_i: ((thm list * attribute list) list) list -> state -> state
- val invoke_case: string * string option list * Attrib.src list -> state -> state
- val invoke_case_i: string * string option list * attribute list -> state -> state
+ val get_thmss_cmd: state -> (Facts.ref * Attrib.src list) list -> thm list
+ val note_thmss: (Thm.binding * (thm list * attribute list) list) list -> state -> state
+ val note_thmss_cmd: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
+ val from_thmss: ((thm list * attribute list) list) list -> state -> state
+ val from_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+ val with_thmss: ((thm list * attribute list) list) list -> state -> state
+ val with_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+ val using: ((thm list * attribute list) list) list -> state -> state
+ val using_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+ val unfolding: ((thm list * attribute list) list) list -> state -> state
+ val unfolding_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+ val invoke_case: string * string option list * attribute list -> state -> state
+ val invoke_case_cmd: string * string option list * Attrib.src list -> state -> state
val begin_block: state -> state
val next_block: state -> state
val end_block: state -> state
@@ -87,9 +85,9 @@
((binding * 'a list) * 'b) list -> state -> state
val local_qed: Method.text option * bool -> state -> state
val theorem: Method.text option -> (thm list list -> context -> context) ->
+ (term * term list) list list -> context -> state
+ val theorem_cmd: Method.text option -> (thm list list -> context -> context) ->
(string * string list) list list -> context -> state
- val theorem_i: Method.text option -> (thm list list -> context -> context) ->
- (term * term list) list list -> context -> state
val global_qed: Method.text option * bool -> state -> context
val local_terminal_proof: Method.text * Method.text option -> state -> state
val local_default_proof: state -> state
@@ -102,13 +100,13 @@
val global_skip_proof: bool -> state -> context
val global_done_proof: state -> context
val have: Method.text option -> (thm list list -> state -> state) ->
+ (Thm.binding * (term * term list) list) list -> bool -> state -> state
+ val have_cmd: Method.text option -> (thm list list -> state -> state) ->
(Attrib.binding * (string * string list) list) list -> bool -> state -> state
- val have_i: Method.text option -> (thm list list -> state -> state) ->
- (Thm.binding * (term * term list) list) list -> bool -> state -> state
val show: Method.text option -> (thm list list -> state -> state) ->
+ (Thm.binding * (term * term list) list) list -> bool -> state -> state
+ val show_cmd: Method.text option -> (thm list list -> state -> state) ->
(Attrib.binding * (string * string list) list) list -> bool -> state -> state
- val show_i: Method.text option -> (thm list list -> state -> state) ->
- (Thm.binding * (term * term list) list) list -> bool -> state -> state
val schematic_goal: state -> bool
val is_relevant: state -> bool
val local_future_proof: (state -> ('a * state) Future.future) ->
@@ -523,22 +521,20 @@
(** context elements **)
-(* bindings *)
+(* let bindings *)
local
fun gen_bind bind args state =
state
|> assert_forward
- |> map_context (bind args #> snd)
+ |> map_context (bind true args #> snd)
|> put_facts NONE;
in
-val match_bind = gen_bind (ProofContext.match_bind false);
-val match_bind_i = gen_bind (ProofContext.match_bind_i false);
-val let_bind = gen_bind (ProofContext.match_bind true);
-val let_bind_i = gen_bind (ProofContext.match_bind_i true);
+val let_bind = gen_bind ProofContext.match_bind_i;
+val let_bind_cmd = gen_bind ProofContext.match_bind;
end;
@@ -554,8 +550,8 @@
in
-val fix = gen_fix (fn ctxt => fn args => fst (ProofContext.read_vars args ctxt));
-val fix_i = gen_fix (K I);
+val fix = gen_fix (K I);
+val fix_cmd = gen_fix (fn ctxt => fn args => fst (ProofContext.read_vars args ctxt));
end;
@@ -572,12 +568,12 @@
in
-val assm = gen_assume ProofContext.add_assms Attrib.attribute;
-val assm_i = gen_assume ProofContext.add_assms_i (K I);
-val assume = assm Assumption.assume_export;
-val assume_i = assm_i Assumption.assume_export;
-val presume = assm Assumption.presume_export;
-val presume_i = assm_i Assumption.presume_export;
+val assm = gen_assume ProofContext.add_assms_i (K I);
+val assm_cmd = gen_assume ProofContext.add_assms Attrib.attribute;
+val assume = assm Assumption.assume_export;
+val assume_cmd = assm_cmd Assumption.assume_export;
+val presume = assm Assumption.presume_export;
+val presume_cmd = assm_cmd Assumption.presume_export;
end;
@@ -605,8 +601,8 @@
in
-val def = gen_def Attrib.attribute ProofContext.read_vars ProofContext.match_bind;
-val def_i = gen_def (K I) ProofContext.cert_vars ProofContext.match_bind_i;
+val def = gen_def (K I) ProofContext.cert_vars ProofContext.match_bind_i;
+val def_cmd = gen_def Attrib.attribute ProofContext.read_vars ProofContext.match_bind;
end;
@@ -646,18 +642,18 @@
in
-val note_thmss = gen_thmss (K []) I #2 Attrib.attribute ProofContext.get_fact;
-val note_thmss_i = gen_thmss (K []) I #2 (K I) (K I);
+val note_thmss = gen_thmss (K []) I #2 (K I) (K I);
+val note_thmss_cmd = gen_thmss (K []) I #2 Attrib.attribute ProofContext.get_fact;
-val from_thmss = gen_thmss (K []) chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
-val from_thmss_i = gen_thmss (K []) chain #2 (K I) (K I) o no_binding;
+val from_thmss = gen_thmss (K []) chain #2 (K I) (K I) o no_binding;
+val from_thmss_cmd = gen_thmss (K []) chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
-val with_thmss = gen_thmss the_facts chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
-val with_thmss_i = gen_thmss the_facts chain #2 (K I) (K I) o no_binding;
+val with_thmss = gen_thmss the_facts chain #2 (K I) (K I) o no_binding;
+val with_thmss_cmd = gen_thmss the_facts chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
val local_results = gen_thmss (K []) I I (K I) (K I) o map (apsnd Thm.simple_fact);
-fun get_thmss state srcs = the_facts (note_thmss [((Binding.empty, []), srcs)] state);
+fun get_thmss_cmd state srcs = the_facts (note_thmss_cmd [((Binding.empty, []), srcs)] state);
end;
@@ -686,10 +682,10 @@
in
-val using = gen_using append_using (K (K I)) Attrib.attribute ProofContext.get_fact;
-val using_i = gen_using append_using (K (K I)) (K I) (K I);
-val unfolding = gen_using unfold_using unfold_goals Attrib.attribute ProofContext.get_fact;
-val unfolding_i = gen_using unfold_using unfold_goals (K I) (K I);
+val using = gen_using append_using (K (K I)) (K I) (K I);
+val using_cmd = gen_using append_using (K (K I)) Attrib.attribute ProofContext.get_fact;
+val unfolding = gen_using unfold_using unfold_goals (K I) (K I);
+val unfolding_cmd = gen_using unfold_using unfold_goals Attrib.attribute ProofContext.get_fact;
end;
@@ -709,15 +705,15 @@
val assumptions = asms |> map (fn (a, ts) => ((qualified_binding a, atts), map (rpair []) ts));
in
state'
- |> assume_i assumptions
+ |> assume assumptions
|> bind_terms Auto_Bind.no_facts
- |> `the_facts |-> (fn thms => note_thmss_i [((Binding.name name, []), [(thms, [])])])
+ |> `the_facts |-> (fn thms => note_thmss [((Binding.name name, []), [(thms, [])])])
end;
in
-val invoke_case = gen_invoke_case Attrib.attribute;
-val invoke_case_i = gen_invoke_case (K I);
+val invoke_case = gen_invoke_case (K I);
+val invoke_case_cmd = gen_invoke_case Attrib.attribute;
end;
@@ -790,15 +786,16 @@
local
-fun implicit_vars dest add props =
+val is_var =
+ can (dest_TVar o Logic.dest_type o Logic.dest_term) orf
+ can (dest_Var o Logic.dest_term);
+
+fun implicit_vars props =
let
- val (explicit_vars, props') = take_prefix (can dest) props |>> map dest;
- val vars = rev (subtract (op =) explicit_vars (fold add props []));
- val _ =
- if null vars then ()
- else warning ("Goal statement contains unbound schematic variable(s): " ^
- commas_quote (map (Term.string_of_vname o fst) vars));
- in (rev vars, props') end;
+ val (var_props, props') = take_prefix is_var props;
+ val explicit_vars = fold Term.add_vars var_props [];
+ val vars = filter_out (member (op =) explicit_vars) (fold Term.add_vars props []);
+ in map (Logic.mk_term o Var) vars end;
fun refine_terms n =
refine (Method.Basic (K (RAW_METHOD
@@ -823,11 +820,8 @@
|> map_context_result (fn ctxt => swap (prepp (ctxt, raw_propp)));
val props = flat propss;
- val (_, props') =
- implicit_vars (dest_TVar o Logic.dest_type o Logic.dest_term) Term.add_tvars props;
- val (vars, _) = implicit_vars (dest_Var o Logic.dest_term) Term.add_vars props';
-
- val propss' = map (Logic.mk_term o Var) vars :: propss;
+ val vars = implicit_vars props;
+ val propss' = vars :: propss;
val goal_propss = filter_out null propss';
val goal =
cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
@@ -903,8 +897,8 @@
init ctxt
|> generic_goal (prepp #> ProofContext.auto_fixes) "" before_qed (K I, after_qed) propp;
-val theorem = global_goal ProofContext.bind_propp_schematic;
-val theorem_i = global_goal ProofContext.bind_propp_schematic_i;
+val theorem = global_goal ProofContext.bind_propp_schematic_i;
+val theorem_cmd = global_goal ProofContext.bind_propp_schematic;
fun global_qeds txt =
end_proof true txt
@@ -976,10 +970,10 @@
in
-val have = gen_have Attrib.attribute ProofContext.bind_propp;
-val have_i = gen_have (K I) ProofContext.bind_propp_i;
-val show = gen_show Attrib.attribute ProofContext.bind_propp;
-val show_i = gen_show (K I) ProofContext.bind_propp_i;
+val have = gen_have (K I) ProofContext.bind_propp_i;
+val have_cmd = gen_have Attrib.attribute ProofContext.bind_propp;
+val show = gen_show (K I) ProofContext.bind_propp_i;
+val show_cmd = gen_show Attrib.attribute ProofContext.bind_propp;
end;
--- a/src/Pure/Isar/specification.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Isar/specification.ML Mon Apr 26 09:45:22 2010 -0700
@@ -403,7 +403,7 @@
goal_ctxt
|> ProofContext.note_thmss "" [((Binding.name Auto_Bind.assmsN, []), [(prems, [])])]
|> snd
- |> Proof.theorem_i before_qed after_qed' (map snd stmt)
+ |> Proof.theorem before_qed after_qed' (map snd stmt)
|> (case facts of NONE => I | SOME ths => Proof.refine_insert ths)
|> tap (fn state => not schematic andalso Proof.schematic_goal state andalso
error "Illegal schematic goal statement")
--- a/src/Pure/ML/ml_thms.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/ML/ml_thms.ML Mon Apr 26 09:45:22 2010 -0700
@@ -64,7 +64,7 @@
fun after_qed res goal_ctxt =
put_thms (i, map prep_result (ProofContext.export goal_ctxt ctxt (flat res))) goal_ctxt;
val ctxt' = ctxt
- |> Proof.theorem_i NONE after_qed propss
+ |> Proof.theorem NONE after_qed propss
|> Proof.global_terminal_proof methods;
val (a, background') = background
|> ML_Antiquote.variant "lemma" ||> put_thms (i, the_thms ctxt' i);
--- a/src/Pure/Thy/thy_output.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/Thy/thy_output.ML Mon Apr 26 09:45:22 2010 -0700
@@ -574,7 +574,7 @@
val prop_src =
(case Args.dest_src source of ((a, arg :: _), pos) => Args.src ((a, [arg]), pos));
val _ = context
- |> Proof.theorem_i NONE (K I) [[(prop, [])]]
+ |> Proof.theorem NONE (K I) [[(prop, [])]]
|> Proof.global_terminal_proof methods;
in output (maybe_pretty_source (pretty_term context) prop_src [prop]) end);
--- a/src/Pure/axclass.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/axclass.ML Mon Apr 26 09:45:22 2010 -0700
@@ -13,9 +13,8 @@
val add_arity: thm -> theory -> theory
val prove_classrel: class * class -> tactic -> theory -> theory
val prove_arity: string * sort list * sort -> tactic -> theory -> theory
- val get_info: theory -> class ->
- {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
- val class_intros: theory -> thm list
+ type info = {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
+ val get_info: theory -> class -> info
val class_of_param: theory -> string -> class option
val cert_classrel: theory -> class * class -> class * class
val read_classrel: theory -> xstring * xstring -> class * class
@@ -37,8 +36,6 @@
val param_of_inst: theory -> string * string -> string
val inst_of_param: theory -> string -> (string * string) option
val thynames_of_arity: theory -> class * string -> string list
- val introN: string
- val axiomsN: string
end;
structure AxClass: AX_CLASS =
@@ -46,6 +43,18 @@
(** theory data **)
+(* axclass info *)
+
+type info =
+ {def: thm,
+ intro: thm,
+ axioms: thm list,
+ params: (string * typ) list};
+
+fun make_axclass (def, intro, axioms, params): info =
+ {def = def, intro = intro, axioms = axioms, params = params};
+
+
(* class parameters (canonical order) *)
type param = string * class;
@@ -57,188 +66,256 @@
" for " ^ Pretty.string_of_sort pp [c] ^
(if c = c' then "" else " and " ^ Pretty.string_of_sort pp [c'])));
-fun merge_params _ ([], qs) = qs
- | merge_params pp (ps, qs) =
- fold_rev (fn q => if member (op =) ps q then I else add_param pp q) qs ps;
+
+(* setup data *)
+
+datatype data = Data of
+ {axclasses: info Symtab.table,
+ params: param list,
+ proven_classrels: (thm * proof) Symreltab.table,
+ proven_arities: ((class * sort list) * ((thm * string) * proof)) list Symtab.table,
+ (*arity theorems with theory name*)
+ inst_params:
+ (string * thm) Symtab.table Symtab.table *
+ (*constant name ~> type constructor ~> (constant name, equation)*)
+ (string * string) Symtab.table (*constant name ~> (constant name, type constructor)*),
+ diff_merge_classrels: (class * class) list};
+
+fun make_data
+ (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =
+ Data {axclasses = axclasses, params = params, proven_classrels = proven_classrels,
+ proven_arities = proven_arities, inst_params = inst_params,
+ diff_merge_classrels = diff_merge_classrels};
+
+structure Data = Theory_Data_PP
+(
+ type T = data;
+ val empty =
+ make_data (Symtab.empty, [], Symreltab.empty, Symtab.empty, (Symtab.empty, Symtab.empty), []);
+ val extend = I;
+ fun merge pp
+ (Data {axclasses = axclasses1, params = params1, proven_classrels = proven_classrels1,
+ proven_arities = proven_arities1, inst_params = inst_params1,
+ diff_merge_classrels = diff_merge_classrels1},
+ Data {axclasses = axclasses2, params = params2, proven_classrels = proven_classrels2,
+ proven_arities = proven_arities2, inst_params = inst_params2,
+ diff_merge_classrels = diff_merge_classrels2}) =
+ let
+ val axclasses' = Symtab.merge (K true) (axclasses1, axclasses2);
+ val params' =
+ if null params1 then params2
+ else fold_rev (fn q => if member (op =) params1 q then I else add_param pp q) params2 params1;
+
+ (*transitive closure of classrels and arity completion is done in Theory.at_begin hook*)
+ val proven_classrels' = Symreltab.join (K #1) (proven_classrels1, proven_classrels2);
+ val proven_arities' =
+ Symtab.join (K (Library.merge (eq_fst op =))) (proven_arities1, proven_arities2);
+
+ val classrels1 = Symreltab.keys proven_classrels1;
+ val classrels2 = Symreltab.keys proven_classrels2;
+ val diff_merge_classrels' =
+ subtract (op =) classrels1 classrels2 @
+ subtract (op =) classrels2 classrels1 @
+ diff_merge_classrels1 @ diff_merge_classrels2;
+
+ val inst_params' =
+ (Symtab.join (K (Symtab.merge (K true))) (#1 inst_params1, #1 inst_params2),
+ Symtab.merge (K true) (#2 inst_params1, #2 inst_params2));
+ in
+ make_data (axclasses', params', proven_classrels', proven_arities', inst_params',
+ diff_merge_classrels')
+ end;
+);
+
+fun map_data f =
+ Data.map (fn Data {axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels} =>
+ make_data (f (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels)));
+
+fun map_axclasses f =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+ (f axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_params f =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+ (axclasses, f params, proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_proven_classrels f =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+ (axclasses, params, f proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_proven_arities f =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+ (axclasses, params, proven_classrels, f proven_arities, inst_params, diff_merge_classrels));
+
+fun map_inst_params f =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+ (axclasses, params, proven_classrels, proven_arities, f inst_params, diff_merge_classrels));
+
+val clear_diff_merge_classrels =
+ map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, _) =>
+ (axclasses, params, proven_classrels, proven_arities, inst_params, []));
+
+val rep_data = Data.get #> (fn Data args => args);
+
+val axclasses_of = #axclasses o rep_data;
+val params_of = #params o rep_data;
+val proven_classrels_of = #proven_classrels o rep_data;
+val proven_arities_of = #proven_arities o rep_data;
+val inst_params_of = #inst_params o rep_data;
+val diff_merge_classrels_of = #diff_merge_classrels o rep_data;
-(* axclasses *)
-
-val introN = "intro";
-val superN = "super";
-val axiomsN = "axioms";
+(* axclasses with parameters *)
-datatype axclass = AxClass of
- {def: thm,
- intro: thm,
- axioms: thm list,
- params: (string * typ) list};
+fun get_info thy c =
+ (case Symtab.lookup (axclasses_of thy) c of
+ SOME info => info
+ | NONE => error ("No such axclass: " ^ quote c));
-type axclasses = axclass Symtab.table * param list;
+fun all_params_of thy S =
+ let val params = params_of thy;
+ in fold (fn (x, c) => if Sign.subsort thy (S, [c]) then cons x else I) params [] end;
-fun make_axclass ((def, intro, axioms), params) = AxClass
- {def = def, intro = intro, axioms = axioms, params = params};
-
-fun merge_axclasses pp ((tab1, params1), (tab2, params2)) : axclasses =
- (Symtab.merge (K true) (tab1, tab2), merge_params pp (params1, params2));
+fun class_of_param thy = AList.lookup (op =) (params_of thy);
-(* instances *)
+(* maintain instances *)
val classrel_prefix = "classrel_";
val arity_prefix = "arity_";
-type instances =
- ((class * class) * thm) list * (*classrel theorems*)
- ((class * sort list) * (thm * string)) list Symtab.table; (*arity theorems with theory name*)
-
-fun merge_instances ((classrel1, arities1): instances, (classrel2, arities2)) =
- (merge (eq_fst op =) (classrel1, classrel2),
- Symtab.join (K (merge (eq_fst op =))) (arities1, arities2));
-
-
-(* instance parameters *)
-
-type inst_params =
- (string * thm) Symtab.table Symtab.table
- (*constant name ~> type constructor ~> (constant name, equation)*)
- * (string * string) Symtab.table; (*constant name ~> (constant name, type constructor)*)
-
-fun merge_inst_params ((const_param1, param_const1), (const_param2, param_const2)) =
- (Symtab.join (K (Symtab.merge (K true))) (const_param1, const_param2),
- Symtab.merge (K true) (param_const1, param_const2));
-
-
-(* setup data *)
-
-structure AxClassData = Theory_Data_PP
-(
- type T = axclasses * (instances * inst_params);
- val empty = ((Symtab.empty, []), (([], Symtab.empty), (Symtab.empty, Symtab.empty)));
- val extend = I;
- fun merge pp ((axclasses1, (instances1, inst_params1)), (axclasses2, (instances2, inst_params2))) =
- (merge_axclasses pp (axclasses1, axclasses2),
- (merge_instances (instances1, instances2), merge_inst_params (inst_params1, inst_params2)));
-);
-
-
-(* maintain axclasses *)
-
-val get_axclasses = #1 o AxClassData.get;
-val map_axclasses = AxClassData.map o apfst;
-
-val lookup_def = Symtab.lookup o #1 o get_axclasses;
-
-fun get_info thy c =
- (case lookup_def thy c of
- SOME (AxClass info) => info
- | NONE => error ("No such axclass: " ^ quote c));
-
-fun class_intros thy =
- let
- fun add_intro c =
- (case lookup_def thy c of SOME (AxClass {intro, ...}) => cons intro | _ => I);
- val classes = Sign.all_classes thy;
- in map (Thm.class_triv thy) classes @ fold add_intro classes [] end;
-
-
-fun get_params thy pred =
- let val params = #2 (get_axclasses thy);
- in fold (fn (x, c) => if pred c then cons x else I) params [] end;
-
-fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));
-
-fun class_of_param thy = AList.lookup (op =) (#2 (get_axclasses thy));
-
-
-(* maintain instances *)
-
fun instance_name (a, c) = Long_Name.base_name c ^ "_" ^ Long_Name.base_name a;
-val get_instances = #1 o #2 o AxClassData.get;
-val map_instances = AxClassData.map o apsnd o apfst;
-
fun the_classrel thy (c1, c2) =
- (case AList.lookup (op =) (#1 (get_instances thy)) (c1, c2) of
- SOME th => Thm.transfer thy th
+ (case Symreltab.lookup (proven_classrels_of thy) (c1, c2) of
+ SOME classrel => classrel
| NONE => error ("Unproven class relation " ^
Syntax.string_of_classrel (ProofContext.init thy) [c1, c2]));
-fun put_classrel arg = map_instances (fn (classrel, arities) =>
- (insert (eq_fst op =) arg classrel, arities));
+fun the_classrel_thm thy = Thm.transfer thy o #1 o the_classrel thy;
+fun the_classrel_prf thy = #2 o the_classrel thy;
+
+fun put_trancl_classrel ((c1, c2), th) thy =
+ let
+ val cert = Thm.cterm_of thy;
+ val certT = Thm.ctyp_of thy;
+
+ val classes = Sorts.classes_of (Sign.classes_of thy);
+ val classrels = proven_classrels_of thy;
+
+ fun reflcl_classrel (c1', c2') =
+ if c1' = c2'
+ then Thm.trivial (cert (Logic.mk_of_class (TVar ((Name.aT, 0), []), c1')))
+ else the_classrel_thm thy (c1', c2');
+ fun gen_classrel (c1_pred, c2_succ) =
+ let
+ val th' = ((reflcl_classrel (c1_pred, c1) RS th) RS reflcl_classrel (c2, c2_succ))
+ |> Drule.instantiate' [SOME (certT (TVar ((Name.aT, 0), [])))] []
+ |> Thm.close_derivation;
+ val prf' = Thm.proof_of th';
+ in ((c1_pred, c2_succ), (th', prf')) end;
+
+ val new_classrels =
+ Library.map_product pair (c1 :: Graph.imm_preds classes c1) (c2 :: Graph.imm_succs classes c2)
+ |> filter_out (Symreltab.defined classrels)
+ |> map gen_classrel;
+ val needed = not (null new_classrels);
+ in
+ (needed,
+ if needed then map_proven_classrels (fold Symreltab.update new_classrels) thy
+ else thy)
+ end;
+
+fun complete_classrels thy =
+ let
+ val classrels = proven_classrels_of thy;
+ val diff_merge_classrels = diff_merge_classrels_of thy;
+ val (needed, thy') = (false, thy) |>
+ fold (fn c12 => fn (needed, thy) =>
+ put_trancl_classrel (c12, Symreltab.lookup classrels c12 |> the |> #1) thy
+ |>> (fn b => needed orelse b))
+ diff_merge_classrels;
+ in
+ if null diff_merge_classrels then NONE
+ else SOME (clear_diff_merge_classrels thy')
+ end;
fun the_arity thy a (c, Ss) =
- (case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of
- SOME (th, _) => Thm.transfer thy th
+ (case AList.lookup (op =) (Symtab.lookup_list (proven_arities_of thy) a) (c, Ss) of
+ SOME arity => arity
| NONE => error ("Unproven type arity " ^
Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));
+fun the_arity_thm thy a c_Ss = the_arity thy a c_Ss |> #1 |> #1 |> Thm.transfer thy;
+fun the_arity_prf thy a c_Ss = the_arity thy a c_Ss |> #2;
+
fun thynames_of_arity thy (c, a) =
- Symtab.lookup_list (#2 (get_instances thy)) a
- |> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)
+ Symtab.lookup_list (proven_arities_of thy) a
+ |> map_filter (fn ((c', _), ((_, name),_)) => if c = c' then SOME name else NONE)
|> rev;
-fun insert_arity_completions thy (t, ((c, Ss), (th, thy_name))) arities =
+fun insert_arity_completions thy (t, ((c, Ss), ((th, thy_name), _))) arities =
let
val algebra = Sign.classes_of thy;
val super_class_completions =
Sign.super_classes thy c
|> filter_out (fn c1 => exists (fn ((c2, Ss2), _) => c1 = c2
andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t));
- val completions = map (fn c1 => (Sorts.classrel_derivation algebra
- (fn (th, c2) => fn c3 => th RS the_classrel thy (c2, c3)) (th, c) c1
- |> Thm.close_derivation, c1)) super_class_completions;
- val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), (th1, thy_name))))
+ val names_and_Ss = Name.names Name.context Name.aT (map (K []) Ss);
+ val completions = super_class_completions |> map (fn c1 =>
+ let
+ val th1 = (th RS the_classrel_thm thy (c, c1))
+ |> Drule.instantiate' (map (SOME o ctyp_of thy o TVar o apfst (rpair 0)) names_and_Ss) []
+ |> Thm.close_derivation;
+ val prf1 = Thm.proof_of th1;
+ in (((th1, thy_name), prf1), c1) end);
+ val arities' = fold (fn (th_thy_prf1, c1) => Symtab.cons_list (t, ((c1, Ss), th_thy_prf1)))
completions arities;
in (null completions, arities') end;
fun put_arity ((t, Ss, c), th) thy =
let
- val arity' = (t, ((c, Ss), (th, Context.theory_name thy)));
+ val arity' = (t, ((c, Ss), ((th, Context.theory_name thy), Thm.proof_of th)));
in
thy
- |> map_instances (fn (classrel, arities) => (classrel,
- arities
- |> Symtab.insert_list (eq_fst op =) arity'
- |> insert_arity_completions thy arity'
- |> snd))
+ |> map_proven_arities
+ (Symtab.insert_list (eq_fst op =) arity' #>
+ insert_arity_completions thy arity' #> snd)
end;
fun complete_arities thy =
let
- val arities = snd (get_instances thy);
+ val arities = proven_arities_of thy;
val (finished, arities') = arities
|> fold_map (insert_arity_completions thy) (Symtab.dest_list arities);
in
- if forall I finished then NONE
- else SOME (thy |> map_instances (fn (classrel, _) => (classrel, arities')))
+ if forall I finished
+ then NONE
+ else SOME (map_proven_arities (K arities') thy)
end;
-val _ = Context.>> (Context.map_theory (Theory.at_begin complete_arities));
+val _ = Context.>> (Context.map_theory
+ (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities));
(* maintain instance parameters *)
-val get_inst_params = #2 o #2 o AxClassData.get;
-val map_inst_params = AxClassData.map o apsnd o apsnd;
+fun get_inst_param thy (c, tyco) =
+ (case Symtab.lookup (the_default Symtab.empty (Symtab.lookup (#1 (inst_params_of thy)) c)) tyco of
+ SOME c' => c'
+ | NONE => error ("No instance parameter for constant " ^ quote c ^ " on type " ^ quote tyco));
-fun get_inst_param thy (c, tyco) =
- case Symtab.lookup ((the_default Symtab.empty o Symtab.lookup (fst (get_inst_params thy))) c) tyco
- of SOME c' => c'
- | NONE => error ("No instance parameter for constant " ^ quote c
- ^ " on type constructor " ^ quote tyco);
-
-fun add_inst_param (c, tyco) inst = (map_inst_params o apfst
- o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
+fun add_inst_param (c, tyco) inst =
+ (map_inst_params o apfst o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
#> (map_inst_params o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
-val inst_of_param = Symtab.lookup o snd o get_inst_params;
+val inst_of_param = Symtab.lookup o #2 o inst_params_of;
val param_of_inst = fst oo get_inst_param;
-fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
- (get_inst_params thy) [];
+fun inst_thms thy =
+ Symtab.fold (Symtab.fold (cons o #2 o #2) o #2) (#1 (inst_params_of thy)) [];
fun get_inst_tyco consts = try (fst o dest_Type o the_single o Consts.typargs consts);
@@ -248,18 +325,20 @@
fun unoverload_conv thy = MetaSimplifier.rewrite true (inst_thms thy);
fun overload_conv thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
-fun lookup_inst_param consts params (c, T) = case get_inst_tyco consts (c, T)
- of SOME tyco => AList.lookup (op =) params (c, tyco)
- | NONE => NONE;
+fun lookup_inst_param consts params (c, T) =
+ (case get_inst_tyco consts (c, T) of
+ SOME tyco => AList.lookup (op =) params (c, tyco)
+ | NONE => NONE);
fun unoverload_const thy (c_ty as (c, _)) =
- if is_some (class_of_param thy c)
- then case get_inst_tyco (Sign.consts_of thy) c_ty
- of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
- | NONE => c
+ if is_some (class_of_param thy c) then
+ (case get_inst_tyco (Sign.consts_of thy) c_ty of
+ SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
+ | NONE => c)
else c;
+
(** instances **)
(* class relations *)
@@ -331,6 +410,8 @@
(* primitive rules *)
+val shyps_topped = forall null o #shyps o Thm.rep_thm;
+
fun add_classrel raw_th thy =
let
val th = Thm.strip_shyps (Thm.transfer thy raw_th);
@@ -338,10 +419,14 @@
fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
val rel = Logic.dest_classrel prop handle TERM _ => err ();
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
+ val th' = th
+ |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [c1])))] []
+ |> Thm.unconstrain_allTs;
+ val _ = shyps_topped th' orelse raise Fail "add_classrel: nontop shyps after unconstrain";
in
thy
|> Sign.primitive_classrel (c1, c2)
- |> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th))
+ |> (snd oo put_trancl_classrel) ((c1, c2), th')
|> perhaps complete_arities
end;
@@ -351,17 +436,22 @@
val prop = Thm.plain_prop_of th;
fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();
- val T = Type (t, map TFree (Name.names Name.context Name.aT Ss));
+ val names = Name.names Name.context Name.aT Ss;
+ val T = Type (t, map TFree names);
val missing_params = Sign.complete_sort thy [c]
|> maps (these o Option.map #params o try (get_info thy))
|> filter_out (fn (const, _) => can (get_inst_param thy) (const, t))
|> (map o apsnd o map_atyps) (K T);
val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();
+ val th' = th
+ |> Drule.instantiate' (map (SOME o ctyp_of thy o TVar o apfst (rpair 0)) names) []
+ |> Thm.unconstrain_allTs;
+ val _ = shyps_topped th' orelse raise Fail "add_arity: nontop shyps after unconstrain";
in
thy
|> fold (snd oo declare_overloaded) missing_params
|> Sign.primitive_arity (t, Ss, [c])
- |> put_arity ((t, Ss, c), Thm.close_derivation (Drule.unconstrainTs th))
+ |> put_arity ((t, Ss, c), th')
end;
@@ -483,25 +573,24 @@
def_thy
|> Sign.qualified_path true bconst
|> PureThy.note_thmss ""
- [((Binding.name introN, []), [([Drule.export_without_context raw_intro], [])]),
- ((Binding.name superN, []), [(map Drule.export_without_context raw_classrel, [])]),
- ((Binding.name axiomsN, []),
+ [((Binding.name "intro", []), [([Drule.export_without_context raw_intro], [])]),
+ ((Binding.name "super", []), [(map Drule.export_without_context raw_classrel, [])]),
+ ((Binding.name "axioms", []),
[(map (fn th => Drule.export_without_context (class_triv RS th)) raw_axioms, [])])]
||> Sign.restore_naming def_thy;
(* result *)
- val axclass = make_axclass ((def, intro, axioms), params);
+ val axclass = make_axclass (def, intro, axioms, params);
val result_thy =
facts_thy
- |> fold put_classrel (map (pair class) super ~~ classrel)
+ |> fold (snd oo put_trancl_classrel) (map (pair class) super ~~ classrel)
|> Sign.qualified_path false bconst
|> PureThy.note_thmss "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms)) |> snd
|> Sign.restore_naming facts_thy
- |> map_axclasses (fn (axclasses, parameters) =>
- (Symtab.update (class, axclass) axclasses,
- fold (fn (x, _) => add_param pp (x, class)) params parameters));
+ |> map_axclasses (Symtab.update (class, axclass))
+ |> map_params (fold (fn (x, _) => add_param pp (x, class)) params);
in (class, result_thy) end;
--- a/src/Pure/display.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/display.ML Mon Apr 26 09:45:22 2010 -0700
@@ -182,7 +182,8 @@
val extern_const = Name_Space.extern (#1 constants);
val {classes, default, types, ...} = Type.rep_tsig tsig;
val (class_space, class_algebra) = classes;
- val {classes, arities} = Sorts.rep_algebra class_algebra;
+ val classes = Sorts.classes_of class_algebra;
+ val arities = Sorts.arities_of class_algebra;
val clsses = Name_Space.dest_table (class_space, Symtab.make (Graph.dest classes));
val tdecls = Name_Space.dest_table types;
--- a/src/Pure/drule.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/drule.ML Mon Apr 26 09:45:22 2010 -0700
@@ -106,7 +106,6 @@
val dummy_thm: thm
val sort_constraintI: thm
val sort_constraint_eq: thm
- val unconstrainTs: thm -> thm
val with_subgoal: int -> (thm -> thm) -> thm -> thm
val comp_no_flatten: thm * int -> int -> thm -> thm
val rename_bvars: (string * string) list -> thm -> thm
@@ -204,12 +203,6 @@
(** Standardization of rules **)
-(* type classes and sorts *)
-
-fun unconstrainTs th =
- fold (Thm.unconstrainT o Thm.ctyp_of (Thm.theory_of_thm th) o TVar)
- (Thm.fold_terms Term.add_tvars th []) th;
-
(*Generalization over a list of variables*)
val forall_intr_list = fold_rev forall_intr;
--- a/src/Pure/sorts.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/sorts.ML Mon Apr 26 09:45:22 2010 -0700
@@ -25,9 +25,8 @@
val insert_term: term -> sort OrdList.T -> sort OrdList.T
val insert_terms: term list -> sort OrdList.T -> sort OrdList.T
type algebra
- val rep_algebra: algebra ->
- {classes: serial Graph.T,
- arities: (class * (class * sort list)) list Symtab.table}
+ val classes_of: algebra -> serial Graph.T
+ val arities_of: algebra -> (class * (class * sort list)) list Symtab.table
val all_classes: algebra -> class list
val super_classes: algebra -> class -> class list
val class_less: algebra -> class * class -> bool
@@ -116,10 +115,8 @@
{classes: serial Graph.T,
arities: (class * (class * sort list)) list Symtab.table};
-fun rep_algebra (Algebra args) = args;
-
-val classes_of = #classes o rep_algebra;
-val arities_of = #arities o rep_algebra;
+fun classes_of (Algebra {classes, ...}) = classes;
+fun arities_of (Algebra {arities, ...}) = arities;
fun make_algebra (classes, arities) =
Algebra {classes = classes, arities = arities};
--- a/src/Pure/thm.ML Mon Apr 26 09:37:46 2010 -0700
+++ b/src/Pure/thm.ML Mon Apr 26 09:45:22 2010 -0700
@@ -92,8 +92,6 @@
val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm
val instantiate_cterm: (ctyp * ctyp) list * (cterm * cterm) list -> cterm -> cterm
val trivial: cterm -> thm
- val of_class: ctyp * class -> thm
- val unconstrainT: ctyp -> thm -> thm
val dest_state: thm * int -> (term * term) list * term list * term * term
val lift_rule: cterm -> thm -> thm
val incr_indexes: int -> thm -> thm
@@ -139,6 +137,9 @@
val adjust_maxidx_thm: int -> thm -> thm
val varifyT_global: thm -> thm
val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
+ val of_class: ctyp * class -> thm
+ val unconstrainT: ctyp -> thm -> thm
+ val unconstrain_allTs: thm -> thm
val freezeT: thm -> thm
val assumption: int -> thm -> thm Seq.seq
val eq_assumption: int -> thm -> thm
@@ -1240,6 +1241,11 @@
prop = Logic.list_implies (constraints, unconstrain prop)})
end;
+fun unconstrain_allTs th =
+ fold (unconstrainT o ctyp_of (theory_of_thm th) o TVar)
+ (fold_terms Term.add_tvars th []) th;
+
+
(* Replace all TFrees not fixed or in the hyps by new TVars *)
fun varifyT_global' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
let