--- a/src/HOL/ATP_Linkup.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/ATP_Linkup.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header{* The Isabelle-ATP Linkup *}
theory ATP_Linkup
-imports Record Presburger SAT Recdef Extraction Relation_Power Hilbert_Choice
+imports Record Hilbert_Choice
uses
"Tools/polyhash.ML"
"Tools/res_clause.ML"
@@ -88,6 +88,9 @@
apply (simp add: COMBC_def)
done
+
+subsection {* Setup of Vampire, E prover and SPASS *}
+
use "Tools/res_axioms.ML" --{*requires the combinators declared above*}
setup ResAxioms.setup
@@ -96,9 +99,6 @@
use "Tools/watcher.ML"
use "Tools/res_atp.ML"
-
-subsection {* Setup for Vampire, E prover and SPASS *}
-
use "Tools/res_atp_provers.ML"
oracle vampire_oracle ("string * int") = {* ResAtpProvers.vampire_o *}
--- a/src/HOL/Complex/Complex_Main.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Complex/Complex_Main.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header{*Comprehensive Complex Theory*}
theory Complex_Main
-imports Fundamental_Theorem_Algebra CLim "../Hyperreal/Hyperreal"
+imports Main Fundamental_Theorem_Algebra CLim "../Hyperreal/Hyperreal"
begin
end
--- a/src/HOL/Complex/ex/MIR.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Complex/ex/MIR.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,9 +5,9 @@
header {* Quatifier elimination for R(0,1,+,floor,<) *}
theory MIR
- imports Real GCD Code_Integer
- uses ("mireif.ML") ("mirtac.ML")
- begin
+imports List Real Code_Integer
+uses ("mireif.ML") ("mirtac.ML")
+begin
declare real_of_int_floor_cancel [simp del]
--- a/src/HOL/Complex/ex/ReflectedFerrack.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Complex/ex/ReflectedFerrack.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,8 +5,8 @@
header {* Quatifier elimination for R(0,1,+,<) *}
theory ReflectedFerrack
- imports GCD Real Efficient_Nat
- uses ("linreif.ML") ("linrtac.ML")
+imports Main GCD Real Efficient_Nat
+uses ("linreif.ML") ("linrtac.ML")
begin
--- a/src/HOL/Dense_Linear_Order.thy Thu Jun 26 10:06:54 2008 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,553 +0,0 @@
-(*
- ID: $Id$
- Author: Amine Chaieb, TU Muenchen
-*)
-
-header {* Dense linear order without endpoints
- and a quantifier elimination procedure in Ferrante and Rackoff style *}
-
-theory Dense_Linear_Order
-imports Finite_Set
-uses
- "Tools/Qelim/qelim.ML"
- "Tools/Qelim/langford_data.ML"
- "Tools/Qelim/ferrante_rackoff_data.ML"
- ("Tools/Qelim/langford.ML")
- ("Tools/Qelim/ferrante_rackoff.ML")
-begin
-
-setup Langford_Data.setup
-setup Ferrante_Rackoff_Data.setup
-
-context linorder
-begin
-
-lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
-
-lemma gather_simps:
- shows
- "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
- and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
- "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
- and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))" by auto
-
-lemma
- gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)"
- by simp
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
-lemma minf_lt: "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
-lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow> (t < x \<longleftrightarrow> False)"
- by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
- by (auto simp add: less_le not_less not_le)
-lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
-lemma pinf_gt: "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
-lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow> (x < t \<longleftrightarrow> False)"
- by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
- by (auto simp add: less_le not_less not_le)
-lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)"
- by (auto simp add: le_less)
-lemma nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x) ;
- \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
- \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x) ;
- \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
- \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto
-
-lemma npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x < t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less)
-lemma npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x \<le> t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists> u\<in> U. x \<le> u )" by auto
-lemma npi_P: "\<forall> x. ~P \<and> P \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
- \<Longrightarrow> \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-lemma npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
- \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto
-
-lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
-proof(clarsimp)
- fix x l u y assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
- and xu: "x<u" and px: "x < t" and ly: "l<y" and yu:"y < u"
- from tU noU ly yu have tny: "t\<noteq>y" by auto
- {assume H: "t < y"
- from less_trans[OF lx px] less_trans[OF H yu]
- have "l < t \<and> t < u" by simp
- with tU noU have "False" by auto}
- hence "\<not> t < y" by auto hence "y \<le> t" by (simp add: not_less)
- thus "y < t" using tny by (simp add: less_le)
-qed
-
-lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
-proof(clarsimp)
- fix x l u y
- assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
- and px: "t < x" and ly: "l<y" and yu:"y < u"
- from tU noU ly yu have tny: "t\<noteq>y" by auto
- {assume H: "y< t"
- from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
- with tU noU have "False" by auto}
- hence "\<not> y<t" by auto hence "t \<le> y" by (auto simp add: not_less)
- thus "t < y" using tny by (simp add:less_le)
-qed
-
-lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
-proof(clarsimp)
- fix x l u y
- assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
- and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
- from tU noU ly yu have tny: "t\<noteq>y" by auto
- {assume H: "t < y"
- from less_le_trans[OF lx px] less_trans[OF H yu]
- have "l < t \<and> t < u" by simp
- with tU noU have "False" by auto}
- hence "\<not> t < y" by auto thus "y \<le> t" by (simp add: not_less)
-qed
-
-lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
-proof(clarsimp)
- fix x l u y
- assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
- and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
- from tU noU ly yu have tny: "t\<noteq>y" by auto
- {assume H: "y< t"
- from less_trans[OF ly H] le_less_trans[OF px xu]
- have "l < t \<and> t < u" by simp
- with tU noU have "False" by auto}
- hence "\<not> y<t" by auto thus "t \<le> y" by (simp add: not_less)
-qed
-lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)" by auto
-lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)" by auto
-lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)" by auto
-
-lemma lin_dense_conj:
- "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
- \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
- \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<and> P2 x)
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
- by blast
-lemma lin_dense_disj:
- "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
- \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
- \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<or> P2 x)
- \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
- by blast
-
-lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
- \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
-by auto
-
-lemma finite_set_intervals:
- assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
- and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
- shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
-proof-
- let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
- let ?xM = "{y. y\<in> S \<and> x \<le> y}"
- let ?a = "Max ?Mx"
- let ?b = "Min ?xM"
- have MxS: "?Mx \<subseteq> S" by blast
- hence fMx: "finite ?Mx" using fS finite_subset by auto
- from lx linS have linMx: "l \<in> ?Mx" by blast
- hence Mxne: "?Mx \<noteq> {}" by blast
- have xMS: "?xM \<subseteq> S" by blast
- hence fxM: "finite ?xM" using fS finite_subset by auto
- from xu uinS have linxM: "u \<in> ?xM" by blast
- hence xMne: "?xM \<noteq> {}" by blast
- have ax:"?a \<le> x" using Mxne fMx by auto
- have xb:"x \<le> ?b" using xMne fxM by auto
- have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
- have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
- have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
- proof(clarsimp)
- fix y assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
- from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
- moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
- moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
- ultimately show "False" by blast
- qed
- from ainS binS noy ax xb px show ?thesis by blast
-qed
-
-lemma finite_set_intervals2:
- assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
- and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
- shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
-proof-
- from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
- obtain a and b where
- as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
- and axb: "a \<le> x \<and> x \<le> b \<and> P x" by auto
- from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
- thus ?thesis using px as bs noS by blast
-qed
-
-end
-
-section {* The classical QE after Langford for dense linear orders *}
-
-context dense_linear_order
-begin
-
-lemma dlo_qe_bnds:
- assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
- shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> (\<forall> l \<in> L. \<forall>u \<in> U. l < u)"
-proof (simp only: atomize_eq, rule iffI)
- assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
- then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
- {fix l u assume l: "l \<in> L" and u: "u \<in> U"
- have "l < x" using xL l by blast
- also have "x < u" using xU u by blast
- finally (less_trans) have "l < u" .}
- thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
-next
- assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
- let ?ML = "Max L"
- let ?MU = "Min U"
- from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
- from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
- from th1 th2 H have "?ML < ?MU" by auto
- with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
- from th3 th1' have "\<forall>l \<in> L. l < w" by auto
- moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
- ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
-qed
-
-lemma dlo_qe_noub:
- assumes ne: "L \<noteq> {}" and fL: "finite L"
- shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
- from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
- from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
- with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
- thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
-qed
-
-lemma dlo_qe_nolb:
- assumes ne: "U \<noteq> {}" and fU: "finite U"
- shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
- from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
- from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
- with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
- thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
-qed
-
-lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x"
- using gt_ex[of t] by auto
-
-lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq
- le_less neq_iff linear less_not_permute
-
-lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
-lemma atoms:
- includes meta_term_syntax
- shows "TERM (less :: 'a \<Rightarrow> _)"
- and "TERM (less_eq :: 'a \<Rightarrow> _)"
- and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
-declare dlo_simps[langfordsimp]
-
-end
-
-(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
-lemma dnf:
- "(P & (Q | R)) = ((P&Q) | (P&R))"
- "((Q | R) & P) = ((Q&P) | (R&P))"
- by blast+
-
-lemmas weak_dnf_simps = simp_thms dnf
-
-lemma nnf_simps:
- "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
- "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
- by blast+
-
-lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
-
-lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
-
-use "Tools/Qelim/langford.ML"
-method_setup dlo = {*
- Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
-*} "Langford's algorithm for quantifier elimination in dense linear orders"
-
-
-section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
-
-text {* Linear order without upper bounds *}
-
-locale linorder_stupid_syntax = linorder
-begin
-notation
- less_eq ("op \<sqsubseteq>") and
- less_eq ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
- less ("op \<sqsubset>") and
- less ("(_/ \<sqsubset> _)" [51, 51] 50)
-
-end
-
-locale linorder_no_ub = linorder_stupid_syntax +
- assumes gt_ex: "\<exists>y. less x y"
-begin
-lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
-lemma pinf_conj:
- assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
- shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
- from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
- from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
- from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
- {fix x assume H: "z \<sqsubset> x"
- from less_trans[OF zz1 H] less_trans[OF zz2 H]
- have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" using z1 zz1 z2 zz2 by auto
- }
- thus ?thesis by blast
-qed
-
-lemma pinf_disj:
- assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
- shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
- from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
- from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
- from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
- {fix x assume H: "z \<sqsubset> x"
- from less_trans[OF zz1 H] less_trans[OF zz2 H]
- have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" using z1 zz1 z2 zz2 by auto
- }
- thus ?thesis by blast
-qed
-
-lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
- from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
- from gt_ex obtain x where x: "z \<sqsubset> x" by blast
- from z x p1 show ?thesis by blast
-qed
-
-end
-
-text {* Linear order without upper bounds *}
-
-locale linorder_no_lb = linorder_stupid_syntax +
- assumes lt_ex: "\<exists>y. less y x"
-begin
-lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
-
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
-lemma minf_conj:
- assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
- shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
- from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
- from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
- from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
- {fix x assume H: "x \<sqsubset> z"
- from less_trans[OF H zz1] less_trans[OF H zz2]
- have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" using z1 zz1 z2 zz2 by auto
- }
- thus ?thesis by blast
-qed
-
-lemma minf_disj:
- assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
- and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
- shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
- from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
- from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
- from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
- {fix x assume H: "x \<sqsubset> z"
- from less_trans[OF H zz1] less_trans[OF H zz2]
- have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" using z1 zz1 z2 zz2 by auto
- }
- thus ?thesis by blast
-qed
-
-lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
-proof-
- from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
- from lt_ex obtain x where x: "x \<sqsubset> z" by blast
- from z x p1 show ?thesis by blast
-qed
-
-end
-
-
-locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
- fixes between
- assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
- and between_same: "between x x = x"
-
-interpretation constr_dense_linear_order < dense_linear_order
- apply unfold_locales
- using gt_ex lt_ex between_less
- by (auto, rule_tac x="between x y" in exI, simp)
-
-context constr_dense_linear_order
-begin
-
-lemma rinf_U:
- assumes fU: "finite U"
- and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
- \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
- and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
- and nmi: "\<not> MP" and npi: "\<not> PP" and ex: "\<exists> x. P x"
- shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
-proof-
- from ex obtain x where px: "P x" by blast
- from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
- then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
- from uU have Une: "U \<noteq> {}" by auto
- term "linorder.Min less_eq"
- let ?l = "linorder.Min less_eq U"
- let ?u = "linorder.Max less_eq U"
- have linM: "?l \<in> U" using fU Une by simp
- have uinM: "?u \<in> U" using fU Une by simp
- have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
- have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
- have th:"?l \<sqsubseteq> u" using uU Une lM by auto
- from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
- have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
- from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
- from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
- have "(\<exists> s\<in> U. P s) \<or>
- (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
- moreover { fix u assume um: "u\<in>U" and pu: "P u"
- have "between u u = u" by (simp add: between_same)
- with um pu have "P (between u u)" by simp
- with um have ?thesis by blast}
- moreover{
- assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
- then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
- and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
- by blast
- from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
- let ?u = "between t1 t2"
- from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
- from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
- with t1M t2M have ?thesis by blast}
- ultimately show ?thesis by blast
- qed
-
-theorem fr_eq:
- assumes fU: "finite U"
- and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
- \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
- and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
- and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
- and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)" and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
- shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
- (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
-proof-
- {
- assume px: "\<exists> x. P x"
- have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
- moreover {assume "MP \<or> PP" hence "?D" by blast}
- moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
- from npmibnd[OF nmibnd npibnd]
- have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
- from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
- ultimately have "?D" by blast}
- moreover
- { assume "?D"
- moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
- moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
- moreover {assume f:"?F" hence "?E" by blast}
- ultimately have "?E" by blast}
- ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
-qed
-
-lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
-lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
-
-lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
-lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
-lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P
-
-lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
- by (rule constr_dense_linear_order_axioms)
-lemma atoms:
- includes meta_term_syntax
- shows "TERM (less :: 'a \<Rightarrow> _)"
- and "TERM (less_eq :: 'a \<Rightarrow> _)"
- and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
- nmi: nmi_thms npi: npi_thms lindense:
- lin_dense_thms qe: fr_eq atoms: atoms]
-
-declaration {*
-let
-fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
-fun generic_whatis phi =
- let
- val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
- fun h x t =
- case term_of t of
- Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
- else Ferrante_Rackoff_Data.Nox
- | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
- else Ferrante_Rackoff_Data.Nox
- | b$y$z => if Term.could_unify (b, lt) then
- if term_of x aconv y then Ferrante_Rackoff_Data.Lt
- else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
- else Ferrante_Rackoff_Data.Nox
- else if Term.could_unify (b, le) then
- if term_of x aconv y then Ferrante_Rackoff_Data.Le
- else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
- else Ferrante_Rackoff_Data.Nox
- else Ferrante_Rackoff_Data.Nox
- | _ => Ferrante_Rackoff_Data.Nox
- in h end
- fun ss phi = HOL_ss addsimps (simps phi)
-in
- Ferrante_Rackoff_Data.funs @{thm "ferrack_axiom"}
- {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
-end
-*}
-
-end
-
-use "Tools/Qelim/ferrante_rackoff.ML"
-
-method_setup ferrack = {*
- Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
-*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
-
-end
--- a/src/HOL/Hyperreal/Poly.thy Thu Jun 26 10:06:54 2008 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1054 +0,0 @@
-(* Title: Poly.thy
- ID: $Id$
- Author: Jacques D. Fleuriot
- Copyright: 2000 University of Edinburgh
-
- Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
-*)
-
-header{*Univariate Real Polynomials*}
-
-theory Poly
-imports Deriv
-begin
-
-text{*Application of polynomial as a real function.*}
-
-consts poly :: "real list => real => real"
-primrec
- poly_Nil: "poly [] x = 0"
- poly_Cons: "poly (h#t) x = h + x * poly t x"
-
-
-subsection{*Arithmetic Operations on Polynomials*}
-
-text{*addition*}
-consts padd :: "[real list, real list] => real list" (infixl "+++" 65)
-primrec
- padd_Nil: "[] +++ l2 = l2"
- padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
- else (h + hd l2)#(t +++ tl l2))"
-
-text{*Multiplication by a constant*}
-consts cmult :: "[real, real list] => real list" (infixl "%*" 70)
-primrec
- cmult_Nil: "c %* [] = []"
- cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
-
-text{*Multiplication by a polynomial*}
-consts pmult :: "[real list, real list] => real list" (infixl "***" 70)
-primrec
- pmult_Nil: "[] *** l2 = []"
- pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
- else (h %* l2) +++ ((0) # (t *** l2)))"
-
-text{*Repeated multiplication by a polynomial*}
-consts mulexp :: "[nat, real list, real list] => real list"
-primrec
- mulexp_zero: "mulexp 0 p q = q"
- mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q"
-
-text{*Exponential*}
-consts pexp :: "[real list, nat] => real list" (infixl "%^" 80)
-primrec
- pexp_0: "p %^ 0 = [1]"
- pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
-
-text{*Quotient related value of dividing a polynomial by x + a*}
-(* Useful for divisor properties in inductive proofs *)
-consts "pquot" :: "[real list, real] => real list"
-primrec
- pquot_Nil: "pquot [] a= []"
- pquot_Cons: "pquot (h#t) a = (if t = [] then [h]
- else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
-
-text{*Differentiation of polynomials (needs an auxiliary function).*}
-consts pderiv_aux :: "nat => real list => real list"
-primrec
- pderiv_aux_Nil: "pderiv_aux n [] = []"
- pderiv_aux_Cons: "pderiv_aux n (h#t) =
- (real n * h)#(pderiv_aux (Suc n) t)"
-
-text{*normalization of polynomials (remove extra 0 coeff)*}
-consts pnormalize :: "real list => real list"
-primrec
- pnormalize_Nil: "pnormalize [] = []"
- pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = [])
- then (if (h = 0) then [] else [h])
- else (h#(pnormalize p)))"
-
-definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
-definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
-text{*Other definitions*}
-
-definition
- poly_minus :: "real list => real list" ("-- _" [80] 80) where
- "-- p = (- 1) %* p"
-
-definition
- pderiv :: "real list => real list" where
- "pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))"
-
-definition
- divides :: "[real list,real list] => bool" (infixl "divides" 70) where
- "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
-
-definition
- order :: "real => real list => nat" where
- --{*order of a polynomial*}
- "order a p = (SOME n. ([-a, 1] %^ n) divides p &
- ~ (([-a, 1] %^ (Suc n)) divides p))"
-
-definition
- degree :: "real list => nat" where
- --{*degree of a polynomial*}
- "degree p = length (pnormalize p) - 1"
-
-definition
- rsquarefree :: "real list => bool" where
- --{*squarefree polynomials --- NB with respect to real roots only.*}
- "rsquarefree p = (poly p \<noteq> poly [] &
- (\<forall>a. (order a p = 0) | (order a p = 1)))"
-
-
-
-lemma padd_Nil2: "p +++ [] = p"
-by (induct p) auto
-declare padd_Nil2 [simp]
-
-lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
-by auto
-
-lemma pminus_Nil: "-- [] = []"
-by (simp add: poly_minus_def)
-declare pminus_Nil [simp]
-
-lemma pmult_singleton: "[h1] *** p1 = h1 %* p1"
-by simp
-
-lemma poly_ident_mult: "1 %* t = t"
-by (induct "t", auto)
-declare poly_ident_mult [simp]
-
-lemma poly_simple_add_Cons: "[a] +++ ((0)#t) = (a#t)"
-by simp
-declare poly_simple_add_Cons [simp]
-
-text{*Handy general properties*}
-
-lemma padd_commut: "b +++ a = a +++ b"
-apply (subgoal_tac "\<forall>a. b +++ a = a +++ b")
-apply (induct_tac [2] "b", auto)
-apply (rule padd_Cons [THEN ssubst])
-apply (case_tac "aa", auto)
-done
-
-lemma padd_assoc [rule_format]: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
-apply (induct "a", simp, clarify)
-apply (case_tac b, simp_all)
-done
-
-lemma poly_cmult_distr [rule_format]:
- "\<forall>q. a %* ( p +++ q) = (a %* p +++ a %* q)"
-apply (induct "p", simp, clarify)
-apply (case_tac "q")
-apply (simp_all add: right_distrib)
-done
-
-lemma pmult_by_x: "[0, 1] *** t = ((0)#t)"
-apply (induct "t", simp)
-apply (auto simp add: poly_ident_mult padd_commut)
-done
-declare pmult_by_x [simp]
-
-
-text{*properties of evaluation of polynomials.*}
-
-lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
-apply (subgoal_tac "\<forall>p2. poly (p1 +++ p2) x = poly (p1) x + poly (p2) x")
-apply (induct_tac [2] "p1", auto)
-apply (case_tac "p2")
-apply (auto simp add: right_distrib)
-done
-
-lemma poly_cmult: "poly (c %* p) x = c * poly p x"
-apply (induct "p")
-apply (case_tac [2] "x=0")
-apply (auto simp add: right_distrib mult_ac)
-done
-
-lemma poly_minus: "poly (-- p) x = - (poly p x)"
-apply (simp add: poly_minus_def)
-apply (auto simp add: poly_cmult)
-done
-
-lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
-apply (subgoal_tac "\<forall>p2. poly (p1 *** p2) x = poly p1 x * poly p2 x")
-apply (simp (no_asm_simp))
-apply (induct "p1")
-apply (auto simp add: poly_cmult)
-apply (case_tac p1)
-apply (auto simp add: poly_cmult poly_add left_distrib right_distrib mult_ac)
-done
-
-lemma poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
-apply (induct "n")
-apply (auto simp add: poly_cmult poly_mult)
-done
-
-text{*More Polynomial Evaluation Lemmas*}
-
-lemma poly_add_rzero: "poly (a +++ []) x = poly a x"
-by simp
-declare poly_add_rzero [simp]
-
-lemma poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
-by (simp add: poly_mult real_mult_assoc)
-
-lemma poly_mult_Nil2: "poly (p *** []) x = 0"
-by (induct "p", auto)
-declare poly_mult_Nil2 [simp]
-
-lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
-apply (induct "n")
-apply (auto simp add: poly_mult real_mult_assoc)
-done
-
-text{*The derivative*}
-
-lemma pderiv_Nil: "pderiv [] = []"
-
-apply (simp add: pderiv_def)
-done
-declare pderiv_Nil [simp]
-
-lemma pderiv_singleton: "pderiv [c] = []"
-by (simp add: pderiv_def)
-declare pderiv_singleton [simp]
-
-lemma pderiv_Cons: "pderiv (h#t) = pderiv_aux 1 t"
-by (simp add: pderiv_def)
-
-lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c"
-by (simp add: DERIV_cmult mult_commute [of _ c])
-
-lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"
-by (rule lemma_DERIV_subst, rule DERIV_pow, simp)
-declare DERIV_pow2 [simp] DERIV_pow [simp]
-
-lemma lemma_DERIV_poly1: "\<forall>n. DERIV (%x. (x ^ (Suc n) * poly p x)) x :>
- x ^ n * poly (pderiv_aux (Suc n) p) x "
-apply (induct "p")
-apply (auto intro!: DERIV_add DERIV_cmult2
- simp add: pderiv_def right_distrib real_mult_assoc [symmetric]
- simp del: realpow_Suc)
-apply (subst mult_commute)
-apply (simp del: realpow_Suc)
-apply (simp add: mult_commute realpow_Suc [symmetric] del: realpow_Suc)
-done
-
-lemma lemma_DERIV_poly: "DERIV (%x. (x ^ (Suc n) * poly p x)) x :>
- x ^ n * poly (pderiv_aux (Suc n) p) x "
-by (simp add: lemma_DERIV_poly1 del: realpow_Suc)
-
-lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: real) x :> D"
-by (rule lemma_DERIV_subst, rule DERIV_add, auto)
-
-lemma poly_DERIV: "DERIV (%x. poly p x) x :> poly (pderiv p) x"
-apply (induct "p")
-apply (auto simp add: pderiv_Cons)
-apply (rule DERIV_add_const)
-apply (rule lemma_DERIV_subst)
-apply (rule lemma_DERIV_poly [where n=0, simplified], simp)
-done
-declare poly_DERIV [simp]
-
-
-text{* Consequences of the derivative theorem above*}
-
-lemma poly_differentiable: "(%x. poly p x) differentiable x"
-
-apply (simp add: differentiable_def)
-apply (blast intro: poly_DERIV)
-done
-declare poly_differentiable [simp]
-
-lemma poly_isCont: "isCont (%x. poly p x) x"
-by (rule poly_DERIV [THEN DERIV_isCont])
-declare poly_isCont [simp]
-
-lemma poly_IVT_pos: "[| a < b; poly p a < 0; 0 < poly p b |]
- ==> \<exists>x. a < x & x < b & (poly p x = 0)"
-apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl)
-apply (auto simp add: order_le_less)
-done
-
-lemma poly_IVT_neg: "[| a < b; 0 < poly p a; poly p b < 0 |]
- ==> \<exists>x. a < x & x < b & (poly p x = 0)"
-apply (insert poly_IVT_pos [where p = "-- p" ])
-apply (simp add: poly_minus neg_less_0_iff_less)
-done
-
-lemma poly_MVT: "a < b ==>
- \<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)"
-apply (drule_tac f = "poly p" in MVT, auto)
-apply (rule_tac x = z in exI)
-apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique])
-done
-
-text{*Lemmas for Derivatives*}
-
-lemma lemma_poly_pderiv_aux_add: "\<forall>p2 n. poly (pderiv_aux n (p1 +++ p2)) x =
- poly (pderiv_aux n p1 +++ pderiv_aux n p2) x"
-apply (induct "p1", simp, clarify)
-apply (case_tac "p2")
-apply (auto simp add: right_distrib)
-done
-
-lemma poly_pderiv_aux_add: "poly (pderiv_aux n (p1 +++ p2)) x =
- poly (pderiv_aux n p1 +++ pderiv_aux n p2) x"
-apply (simp add: lemma_poly_pderiv_aux_add)
-done
-
-lemma lemma_poly_pderiv_aux_cmult: "\<forall>n. poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x"
-apply (induct "p")
-apply (auto simp add: poly_cmult mult_ac)
-done
-
-lemma poly_pderiv_aux_cmult: "poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x"
-by (simp add: lemma_poly_pderiv_aux_cmult)
-
-lemma poly_pderiv_aux_minus:
- "poly (pderiv_aux n (-- p)) x = poly (-- pderiv_aux n p) x"
-apply (simp add: poly_minus_def poly_pderiv_aux_cmult)
-done
-
-lemma lemma_poly_pderiv_aux_mult1: "\<forall>n. poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x"
-apply (induct "p")
-apply (auto simp add: real_of_nat_Suc left_distrib)
-done
-
-lemma lemma_poly_pderiv_aux_mult: "poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x"
-by (simp add: lemma_poly_pderiv_aux_mult1)
-
-lemma lemma_poly_pderiv_add: "\<forall>q. poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x"
-apply (induct "p", simp, clarify)
-apply (case_tac "q")
-apply (auto simp add: poly_pderiv_aux_add poly_add pderiv_def)
-done
-
-lemma poly_pderiv_add: "poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x"
-by (simp add: lemma_poly_pderiv_add)
-
-lemma poly_pderiv_cmult: "poly (pderiv (c %* p)) x = poly (c %* (pderiv p)) x"
-apply (induct "p")
-apply (auto simp add: poly_pderiv_aux_cmult poly_cmult pderiv_def)
-done
-
-lemma poly_pderiv_minus: "poly (pderiv (--p)) x = poly (--(pderiv p)) x"
-by (simp add: poly_minus_def poly_pderiv_cmult)
-
-lemma lemma_poly_mult_pderiv:
- "poly (pderiv (h#t)) x = poly ((0 # (pderiv t)) +++ t) x"
-apply (simp add: pderiv_def)
-apply (induct "t")
-apply (auto simp add: poly_add lemma_poly_pderiv_aux_mult)
-done
-
-lemma poly_pderiv_mult: "\<forall>q. poly (pderiv (p *** q)) x =
- poly (p *** (pderiv q) +++ q *** (pderiv p)) x"
-apply (induct "p")
-apply (auto simp add: poly_add poly_cmult poly_pderiv_cmult poly_pderiv_add poly_mult)
-apply (rule lemma_poly_mult_pderiv [THEN ssubst])
-apply (rule lemma_poly_mult_pderiv [THEN ssubst])
-apply (rule poly_add [THEN ssubst])
-apply (rule poly_add [THEN ssubst])
-apply (simp (no_asm_simp) add: poly_mult right_distrib add_ac mult_ac)
-done
-
-lemma poly_pderiv_exp: "poly (pderiv (p %^ (Suc n))) x =
- poly ((real (Suc n)) %* (p %^ n) *** pderiv p) x"
-apply (induct "n")
-apply (auto simp add: poly_add poly_pderiv_cmult poly_cmult poly_pderiv_mult
- real_of_nat_zero poly_mult real_of_nat_Suc
- right_distrib left_distrib mult_ac)
-done
-
-lemma poly_pderiv_exp_prime: "poly (pderiv ([-a, 1] %^ (Suc n))) x =
- poly (real (Suc n) %* ([-a, 1] %^ n)) x"
-apply (simp add: poly_pderiv_exp poly_mult del: pexp_Suc)
-apply (simp add: poly_cmult pderiv_def)
-done
-
-subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
- @{term "p(x)"} *}
-
-lemma lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-apply (induct "t", safe)
-apply (rule_tac x = "[]" in exI)
-apply (rule_tac x = h in exI, simp)
-apply (drule_tac x = aa in spec, safe)
-apply (rule_tac x = "r#q" in exI)
-apply (rule_tac x = "a*r + h" in exI)
-apply (case_tac "q", auto)
-done
-
-lemma poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
-by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto)
-
-
-lemma poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
-apply (auto simp add: poly_add poly_cmult right_distrib)
-apply (case_tac "p", simp)
-apply (cut_tac h = aa and t = list and a = a in poly_linear_rem, safe)
-apply (case_tac "q", auto)
-apply (drule_tac x = "[]" in spec, simp)
-apply (auto simp add: poly_add poly_cmult add_assoc)
-apply (drule_tac x = "aa#lista" in spec, auto)
-done
-
-lemma lemma_poly_length_mult: "\<forall>h k a. length (k %* p +++ (h # (a %* p))) = Suc (length p)"
-by (induct "p", auto)
-declare lemma_poly_length_mult [simp]
-
-lemma lemma_poly_length_mult2: "\<forall>h k. length (k %* p +++ (h # p)) = Suc (length p)"
-by (induct "p", auto)
-declare lemma_poly_length_mult2 [simp]
-
-lemma poly_length_mult: "length([-a,1] *** q) = Suc (length q)"
-by auto
-declare poly_length_mult [simp]
-
-
-subsection{*Polynomial length*}
-
-lemma poly_cmult_length: "length (a %* p) = length p"
-by (induct "p", auto)
-declare poly_cmult_length [simp]
-
-lemma poly_add_length [rule_format]:
- "\<forall>p2. length (p1 +++ p2) =
- (if (length p1 < length p2) then length p2 else length p1)"
-apply (induct "p1", simp_all)
-apply arith
-done
-
-lemma poly_root_mult_length: "length([a,b] *** p) = Suc (length p)"
-by (simp add: poly_cmult_length poly_add_length)
-declare poly_root_mult_length [simp]
-
-lemma poly_mult_not_eq_poly_Nil: "(poly (p *** q) x \<noteq> poly [] x) =
- (poly p x \<noteq> poly [] x & poly q x \<noteq> poly [] x)"
-apply (auto simp add: poly_mult)
-done
-declare poly_mult_not_eq_poly_Nil [simp]
-
-lemma poly_mult_eq_zero_disj: "(poly (p *** q) x = 0) = (poly p x = 0 | poly q x = 0)"
-by (auto simp add: poly_mult)
-
-text{*Normalisation Properties*}
-
-lemma poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
-by (induct "p", auto)
-
-text{*A nontrivial polynomial of degree n has no more than n roots*}
-
-lemma poly_roots_index_lemma [rule_format]:
- "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
- --> (\<exists>i. \<forall>x. (poly p x = (0::real)) --> (\<exists>m. (m \<le> n & x = i m)))"
-apply (induct "n", safe)
-apply (rule ccontr)
-apply (subgoal_tac "\<exists>a. poly p a = 0", safe)
-apply (drule poly_linear_divides [THEN iffD1], safe)
-apply (drule_tac x = q in spec)
-apply (drule_tac x = x in spec)
-apply (simp del: poly_Nil pmult_Cons)
-apply (erule exE)
-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 simp
-apply (drule_tac x = xa in spec, safe)
-apply (drule_tac x = m in spec, simp, blast)
-done
-lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard]
-
-lemma poly_roots_index_length: "poly p x \<noteq> poly [] x ==>
- \<exists>i. \<forall>x. (poly p x = 0) --> (\<exists>n. n \<le> length p & x = i n)"
-by (blast intro: poly_roots_index_lemma2)
-
-lemma poly_roots_finite_lemma: "poly p x \<noteq> poly [] x ==>
- \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)"
-apply (drule poly_roots_index_length, safe)
-apply (rule_tac x = "Suc (length p)" in exI)
-apply (rule_tac x = i in exI)
-apply (simp add: less_Suc_eq_le)
-done
-
-(* annoying proof *)
-lemma real_finite_lemma [rule_format (no_asm)]:
- "\<forall>P. (\<forall>x. P x --> (\<exists>n. n < N & x = (j::nat=>real) n))
- --> (\<exists>a. \<forall>x. P x --> x < a)"
-apply (induct "N", simp, safe)
-apply (drule_tac x = "%z. P z & (z \<noteq> j N)" in spec)
-apply (auto simp add: less_Suc_eq)
-apply (rename_tac N P a)
-apply (rule_tac x = "abs a + abs (j N) + 1" in exI)
-apply safe
-apply (drule_tac x = x in spec, safe)
-apply (drule_tac x = "j n" in spec)
-apply arith
-apply arith
-done
-
-lemma poly_roots_finite: "(poly p \<noteq> poly []) =
- (\<exists>N j. \<forall>x. poly p x = 0 --> (\<exists>n. (n::nat) < N & x = j n))"
-apply safe
-apply (erule contrapos_np, rule ext)
-apply (rule ccontr)
-apply (clarify dest!: poly_roots_finite_lemma)
-apply (clarify dest!: real_finite_lemma)
-apply (drule_tac x = a in fun_cong, auto)
-done
-
-text{*Entirety and Cancellation for polynomials*}
-
-lemma poly_entire_lemma: "[| poly p \<noteq> poly [] ; poly q \<noteq> poly [] |]
- ==> poly (p *** q) \<noteq> poly []"
-apply (auto simp add: poly_roots_finite)
-apply (rule_tac x = "N + Na" in exI)
-apply (rule_tac x = "%n. if n < N then j n else ja (n - N)" in exI)
-apply (auto simp add: poly_mult_eq_zero_disj, force)
-done
-
-lemma poly_entire: "(poly (p *** q) = poly []) = ((poly p = poly []) | (poly q = poly []))"
-apply (auto intro: ext dest: fun_cong simp add: poly_entire_lemma poly_mult)
-apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst])
-done
-
-lemma poly_entire_neg: "(poly (p *** q) \<noteq> poly []) = ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))"
-by (simp add: poly_entire)
-
-lemma fun_eq: " (f = g) = (\<forall>x. f x = g x)"
-by (auto intro!: ext)
-
-lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
-by (auto simp add: 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)
-
-lemma poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)"
-apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst])
-apply (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
-done
-
-lemma real_mult_zero_disj_iff: "(x * y = 0) = (x = (0::real) | y = 0)"
-by simp
-
-lemma poly_exp_eq_zero:
- "(poly (p %^ n) = poly []) = (poly p = poly [] & n \<noteq> 0)"
-apply (simp only: fun_eq add: all_simps [symmetric])
-apply (rule arg_cong [where f = All])
-apply (rule ext)
-apply (induct_tac "n")
-apply (auto simp add: poly_mult real_mult_zero_disj_iff)
-done
-declare poly_exp_eq_zero [simp]
-
-lemma poly_prime_eq_zero: "poly [a,1] \<noteq> poly []"
-apply (simp add: fun_eq)
-apply (rule_tac x = "1 - a" in exI, simp)
-done
-declare poly_prime_eq_zero [simp]
-
-lemma poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \<noteq> poly [])"
-by auto
-declare poly_exp_prime_eq_zero [simp]
-
-text{*A more constructive notion of polynomials being trivial*}
-
-lemma poly_zero_lemma: "poly (h # t) = poly [] ==> h = 0 & poly t = poly []"
-apply (simp add: fun_eq)
-apply (case_tac "h = 0")
-apply (drule_tac [2] x = 0 in spec, auto)
-apply (case_tac "poly t = poly []", simp)
-apply (auto simp add: poly_roots_finite real_mult_zero_disj_iff)
-apply (drule real_finite_lemma, safe)
-apply (drule_tac x = "abs a + 1" in spec)+
-apply arith
-done
-
-
-lemma poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
-apply (induct "p", simp)
-apply (rule iffI)
-apply (drule poly_zero_lemma, auto)
-done
-
-declare real_mult_zero_disj_iff [simp]
-
-lemma pderiv_aux_iszero [rule_format, simp]:
- "\<forall>n. list_all (%c. c = 0) (pderiv_aux (Suc n) p) = list_all (%c. c = 0) p"
-by (induct "p", auto)
-
-lemma pderiv_aux_iszero_num: "(number_of n :: nat) \<noteq> 0
- ==> (list_all (%c. c = 0) (pderiv_aux (number_of n) p) =
- list_all (%c. c = 0) p)"
-apply(drule not0_implies_Suc, clarify)
-apply (rule_tac n1 = "m" in pderiv_aux_iszero [THEN subst])
-apply (simp (no_asm_simp) del: pderiv_aux_iszero)
-done
-
-lemma pderiv_iszero [rule_format]:
- "poly (pderiv p) = poly [] --> (\<exists>h. poly p = poly [h])"
-apply (simp add: poly_zero)
-apply (induct "p", force)
-apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons)
-apply (auto simp add: poly_zero [symmetric])
-done
-
-lemma pderiv_zero_obj: "poly p = poly [] --> (poly (pderiv p) = poly [])"
-apply (simp add: poly_zero)
-apply (induct "p", force)
-apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons)
-done
-
-lemma pderiv_zero: "poly p = poly [] ==> (poly (pderiv p) = poly [])"
-by (blast elim: pderiv_zero_obj [THEN impE])
-declare pderiv_zero [simp]
-
-lemma poly_pderiv_welldef: "poly p = poly q ==> (poly (pderiv p) = poly (pderiv q))"
-apply (cut_tac p = "p +++ --q" in pderiv_zero_obj)
-apply (simp add: fun_eq poly_add poly_minus poly_pderiv_add poly_pderiv_minus del: pderiv_zero)
-done
-
-text{*Basics of divisibility.*}
-
-lemma poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)"
-apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric])
-apply (drule_tac x = "-a" in spec)
-apply (auto simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
-apply (rule_tac x = "qa *** q" in exI)
-apply (rule_tac [2] x = "p *** qa" in exI)
-apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
-done
-
-lemma poly_divides_refl: "p divides p"
-apply (simp add: divides_def)
-apply (rule_tac x = "[1]" in exI)
-apply (auto simp add: poly_mult fun_eq)
-done
-declare poly_divides_refl [simp]
-
-lemma poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r"
-apply (simp add: divides_def, safe)
-apply (rule_tac x = "qa *** qaa" in exI)
-apply (auto simp add: poly_mult fun_eq real_mult_assoc)
-done
-
-lemma poly_divides_exp: "m \<le> n ==> (p %^ m) divides (p %^ n)"
-apply (auto simp add: le_iff_add)
-apply (induct_tac k)
-apply (rule_tac [2] poly_divides_trans)
-apply (auto simp add: divides_def)
-apply (rule_tac x = p in exI)
-apply (auto simp add: poly_mult fun_eq mult_ac)
-done
-
-lemma poly_exp_divides: "[| (p %^ n) divides q; m\<le>n |] ==> (p %^ m) divides q"
-by (blast intro: poly_divides_exp poly_divides_trans)
-
-lemma poly_divides_add:
- "[| p divides q; p divides r |] ==> p divides (q +++ r)"
-apply (simp add: divides_def, auto)
-apply (rule_tac x = "qa +++ qaa" in exI)
-apply (auto simp add: poly_add fun_eq poly_mult right_distrib)
-done
-
-lemma poly_divides_diff:
- "[| p divides q; p divides (q +++ r) |] ==> p divides r"
-apply (simp add: divides_def, auto)
-apply (rule_tac x = "qaa +++ -- qa" in exI)
-apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac)
-done
-
-lemma poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q"
-apply (erule poly_divides_diff)
-apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
-done
-
-lemma poly_divides_zero: "poly p = poly [] ==> q divides p"
-apply (simp add: divides_def)
-apply (auto simp add: fun_eq poly_mult)
-done
-
-lemma poly_divides_zero2: "q divides []"
-apply (simp add: divides_def)
-apply (rule_tac x = "[]" in exI)
-apply (auto simp add: fun_eq)
-done
-declare poly_divides_zero2 [simp]
-
-text{*At last, we can consider the order of a root.*}
-
-
-lemma poly_order_exists_lemma [rule_format]:
- "\<forall>p. length p = d --> poly p \<noteq> poly []
- --> (\<exists>n q. p = mulexp n [-a, 1] q & poly q a \<noteq> 0)"
-apply (induct "d")
-apply (simp add: fun_eq, safe)
-apply (case_tac "poly p a = 0")
-apply (drule_tac poly_linear_divides [THEN iffD1], safe)
-apply (drule_tac x = q in spec)
-apply (drule_tac poly_entire_neg [THEN iffD1], safe, force, blast)
-apply (rule_tac x = "Suc n" in exI)
-apply (rule_tac x = qa in exI)
-apply (simp del: pmult_Cons)
-apply (rule_tac x = 0 in exI, force)
-done
-
-(* FIXME: Tidy up *)
-lemma poly_order_exists:
- "[| length p = d; poly p \<noteq> poly [] |]
- ==> \<exists>n. ([-a, 1] %^ n) divides p &
- ~(([-a, 1] %^ (Suc n)) divides p)"
-apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)
-apply (rule_tac x = n in exI, safe)
-apply (unfold divides_def)
-apply (rule_tac x = q in exI)
-apply (induct_tac "n", simp)
-apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac)
-apply safe
-apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)")
-apply simp
-apply (induct_tac "n")
-apply (simp del: pmult_Cons pexp_Suc)
-apply (erule_tac Q = "poly q a = 0" in contrapos_np)
-apply (simp add: poly_add poly_cmult)
-apply (rule pexp_Suc [THEN ssubst])
-apply (rule ccontr)
-apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
-done
-
-lemma poly_one_divides: "[1] divides p"
-by (simp add: divides_def, auto)
-declare poly_one_divides [simp]
-
-lemma poly_order: "poly p \<noteq> poly []
- ==> EX! n. ([-a, 1] %^ n) divides p &
- ~(([-a, 1] %^ (Suc n)) divides p)"
-apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
-apply (metis Suc_leI less_linear poly_exp_divides)
-done
-
-text{*Order*}
-
-lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n"
-by (blast intro: someI2)
-
-lemma order:
- "(([-a, 1] %^ n) divides p &
- ~(([-a, 1] %^ (Suc n)) divides p)) =
- ((n = order a p) & ~(poly p = poly []))"
-apply (unfold order_def)
-apply (rule iffI)
-apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
-apply (blast intro!: poly_order [THEN [2] some1_equalityD])
-done
-
-lemma order2: "[| poly p \<noteq> poly [] |]
- ==> ([-a, 1] %^ (order a p)) divides p &
- ~(([-a, 1] %^ (Suc(order a p))) divides p)"
-by (simp add: order del: pexp_Suc)
-
-lemma order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
- ~(([-a, 1] %^ (Suc n)) divides p)
- |] ==> (n = order a p)"
-by (insert order [of a n p], auto)
-
-lemma order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
- ~(([-a, 1] %^ (Suc n)) divides p))
- ==> (n = order a p)"
-by (blast intro: order_unique)
-
-lemma order_poly: "poly p = poly q ==> order a p = order a q"
-by (auto simp add: fun_eq divides_def poly_mult order_def)
-
-lemma pexp_one: "p %^ (Suc 0) = p"
-apply (induct "p")
-apply (auto simp add: numeral_1_eq_1)
-done
-declare pexp_one [simp]
-
-lemma lemma_order_root [rule_format]:
- "\<forall>p a. n > 0 & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p
- --> poly p a = 0"
-apply (induct "n", blast)
-apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
-done
-
-lemma order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \<noteq> 0)"
-apply (case_tac "poly p = poly []", auto)
-apply (simp add: poly_linear_divides del: pmult_Cons, safe)
-apply (drule_tac [!] a = a in order2)
-apply (rule ccontr)
-apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
-apply (blast intro: lemma_order_root)
-done
-
-lemma order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
-apply (case_tac "poly p = poly []", auto)
-apply (simp add: divides_def fun_eq poly_mult)
-apply (rule_tac x = "[]" in exI)
-apply (auto dest!: order2 [where a=a]
- intro: poly_exp_divides simp del: pexp_Suc)
-done
-
-lemma order_decomp:
- "poly p \<noteq> poly []
- ==> \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) &
- ~([-a, 1] divides q)"
-apply (unfold divides_def)
-apply (drule order2 [where a = a])
-apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
-apply (rule_tac x = q in exI, safe)
-apply (drule_tac x = qa in spec)
-apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
-done
-
-text{*Important composition properties of orders.*}
-
-lemma order_mult: "poly (p *** q) \<noteq> poly []
- ==> order a (p *** q) = order a p + order a q"
-apply (cut_tac a = a and p = "p***q" and n = "order a p + order a q" in order)
-apply (auto simp add: poly_entire simp del: pmult_Cons)
-apply (drule_tac a = a in order2)+
-apply safe
-apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
-apply (rule_tac x = "qa *** qaa" in exI)
-apply (simp add: poly_mult mult_ac del: pmult_Cons)
-apply (drule_tac a = a in order_decomp)+
-apply safe
-apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
-apply (simp add: poly_primes del: pmult_Cons)
-apply (auto simp add: divides_def simp del: pmult_Cons)
-apply (rule_tac x = qb in exI)
-apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
-apply (drule poly_mult_left_cancel [THEN iffD1], force)
-apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
-done
-
-(* FIXME: too too long! *)
-lemma lemma_order_pderiv [rule_format]:
- "\<forall>p q a. n > 0 &
- poly (pderiv p) \<noteq> poly [] &
- poly p = poly ([- a, 1] %^ n *** q) & ~ [- a, 1] divides q
- --> n = Suc (order a (pderiv p))"
-apply (induct "n", safe)
-apply (rule order_unique_lemma, rule conjI, assumption)
-apply (subgoal_tac "\<forall>r. r divides (pderiv p) = r divides (pderiv ([-a, 1] %^ Suc n *** q))")
-apply (drule_tac [2] poly_pderiv_welldef)
- prefer 2 apply (simp add: divides_def del: pmult_Cons pexp_Suc)
-apply (simp del: pmult_Cons pexp_Suc)
-apply (rule conjI)
-apply (simp add: divides_def fun_eq del: pmult_Cons pexp_Suc)
-apply (rule_tac x = "[-a, 1] *** (pderiv q) +++ real (Suc n) %* q" in exI)
-apply (simp add: poly_pderiv_mult poly_pderiv_exp_prime poly_add poly_mult poly_cmult right_distrib mult_ac del: pmult_Cons pexp_Suc)
-apply (simp add: poly_mult right_distrib left_distrib mult_ac del: pmult_Cons)
-apply (erule_tac V = "\<forall>r. r divides pderiv p = r divides pderiv ([- a, 1] %^ Suc n *** q)" in thin_rl)
-apply (unfold divides_def)
-apply (simp (no_asm) add: poly_pderiv_mult poly_pderiv_exp_prime fun_eq poly_add poly_mult del: pmult_Cons pexp_Suc)
-apply (rule contrapos_np, assumption)
-apply (rotate_tac 3, erule contrapos_np)
-apply (simp del: pmult_Cons pexp_Suc, safe)
-apply (rule_tac x = "inverse (real (Suc n)) %* (qa +++ -- (pderiv q))" in exI)
-apply (subgoal_tac "poly ([-a, 1] %^ n *** q) = poly ([-a, 1] %^ n *** ([-a, 1] *** (inverse (real (Suc n)) %* (qa +++ -- (pderiv q))))) ")
-apply (drule poly_mult_left_cancel [THEN iffD1], simp)
-apply (simp add: fun_eq poly_mult poly_add poly_cmult poly_minus del: pmult_Cons mult_cancel_left, safe)
-apply (rule_tac c1 = "real (Suc n)" in real_mult_left_cancel [THEN iffD1])
-apply (simp (no_asm))
-apply (subgoal_tac "real (Suc n) * (poly ([- a, 1] %^ n) xa * poly q xa) =
- (poly qa xa + - poly (pderiv q) xa) *
- (poly ([- a, 1] %^ n) xa *
- ((- a + xa) * (inverse (real (Suc n)) * real (Suc n))))")
-apply (simp only: mult_ac)
-apply (rotate_tac 2)
-apply (drule_tac x = xa in spec)
-apply (simp add: left_distrib mult_ac del: pmult_Cons)
-done
-
-lemma order_pderiv: "[| poly (pderiv p) \<noteq> poly []; order a p \<noteq> 0 |]
- ==> (order a p = Suc (order a (pderiv p)))"
-apply (case_tac "poly p = poly []")
-apply (auto dest: pderiv_zero)
-apply (drule_tac a = a and p = p in order_decomp)
-apply (metis lemma_order_pderiv length_0_conv length_greater_0_conv)
-done
-
-text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *)
-(* `a la Harrison*}
-
-lemma poly_squarefree_decomp_order: "[| poly (pderiv p) \<noteq> poly [];
- poly p = poly (q *** d);
- poly (pderiv p) = poly (e *** d);
- poly d = poly (r *** p +++ s *** pderiv p)
- |] ==> order a q = (if order a p = 0 then 0 else 1)"
-apply (subgoal_tac "order a p = order a q + order a d")
-apply (rule_tac [2] s = "order a (q *** d)" in trans)
-prefer 2 apply (blast intro: order_poly)
-apply (rule_tac [2] order_mult)
- prefer 2 apply force
-apply (case_tac "order a p = 0", simp)
-apply (subgoal_tac "order a (pderiv p) = order a e + order a d")
-apply (rule_tac [2] s = "order a (e *** d)" in trans)
-prefer 2 apply (blast intro: order_poly)
-apply (rule_tac [2] order_mult)
- prefer 2 apply force
-apply (case_tac "poly p = poly []")
-apply (drule_tac p = p in pderiv_zero, simp)
-apply (drule order_pderiv, assumption)
-apply (subgoal_tac "order a (pderiv p) \<le> order a d")
-apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides d")
- prefer 2 apply (simp add: poly_entire order_divides)
-apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides p & ([-a, 1] %^ (order a (pderiv p))) divides (pderiv p) ")
- prefer 3 apply (simp add: order_divides)
- prefer 2 apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
-apply (rule_tac x = "r *** qa +++ s *** qaa" in exI)
-apply (simp add: fun_eq poly_add poly_mult left_distrib right_distrib mult_ac del: pexp_Suc pmult_Cons, auto)
-done
-
-
-lemma poly_squarefree_decomp_order2: "[| poly (pderiv p) \<noteq> poly [];
- poly p = poly (q *** d);
- poly (pderiv p) = poly (e *** d);
- poly d = poly (r *** p +++ s *** pderiv p)
- |] ==> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-apply (blast intro: poly_squarefree_decomp_order)
-done
-
-lemma order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order a p \<noteq> 0)"
-by (rule order_root [THEN ssubst], auto)
-
-lemma order_pderiv2: "[| poly (pderiv p) \<noteq> poly []; order a p \<noteq> 0 |]
- ==> (order a (pderiv p) = n) = (order a p = Suc n)"
-by (metis Suc_Suc_eq order_pderiv)
-
-lemma rsquarefree_roots:
- "rsquarefree p = (\<forall>a. ~(poly p a = 0 & poly (pderiv p) a = 0))"
-apply (simp add: rsquarefree_def)
-apply (case_tac "poly p = poly []", simp, simp)
-apply (case_tac "poly (pderiv p) = poly []")
-apply simp
-apply (drule pderiv_iszero, clarify)
-apply (subgoal_tac "\<forall>a. order a p = order a [h]")
-apply (simp add: fun_eq)
-apply (rule allI)
-apply (cut_tac p = "[h]" and a = a in order_root)
-apply (simp add: fun_eq)
-apply (blast intro: order_poly)
-apply (metis One_nat_def order_pderiv2 order_root rsquarefree_def)
-done
-
-lemma pmult_one: "[1] *** p = p"
-by auto
-declare pmult_one [simp]
-
-lemma poly_Nil_zero: "poly [] = poly [0]"
-by (simp add: fun_eq)
-
-lemma rsquarefree_decomp:
- "[| rsquarefree p; poly p a = 0 |]
- ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0"
-apply (simp add: rsquarefree_def, safe)
-apply (frule_tac a = a in order_decomp)
-apply (drule_tac x = a in spec)
-apply (drule_tac a = a in order_root2 [symmetric])
-apply (auto simp del: pmult_Cons)
-apply (rule_tac x = q in exI, safe)
-apply (simp add: poly_mult fun_eq)
-apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
-apply (simp add: divides_def del: pmult_Cons, safe)
-apply (drule_tac x = "[]" in spec)
-apply (auto simp add: fun_eq)
-done
-
-lemma poly_squarefree_decomp: "[| poly (pderiv p) \<noteq> poly [];
- poly p = poly (q *** d);
- poly (pderiv p) = poly (e *** d);
- poly d = poly (r *** p +++ s *** pderiv p)
- |] ==> rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
-apply (frule poly_squarefree_decomp_order2, assumption+)
-apply (case_tac "poly p = poly []")
-apply (blast dest: pderiv_zero)
-apply (simp (no_asm) add: rsquarefree_def order_root del: pmult_Cons)
-apply (simp add: poly_entire del: pmult_Cons)
-done
-
-
-text{*Normalization of a polynomial.*}
-
-lemma poly_normalize: "poly (pnormalize p) = poly p"
-apply (induct "p")
-apply (auto simp add: fun_eq)
-done
-declare poly_normalize [simp]
-
-
-text{*The degree of a polynomial.*}
-
-lemma lemma_degree_zero:
- "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
-by (induct "p", auto)
-
-lemma degree_zero: "(poly p = poly []) \<Longrightarrow> (degree p = 0)"
-apply (simp add: degree_def)
-apply (case_tac "pnormalize p = []")
-apply (auto simp add: poly_zero lemma_degree_zero )
-done
-
-lemma pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
-lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
-lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
- unfolding pnormal_def by simp
-lemma pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
- unfolding pnormal_def
- apply (cases "pnormalize p = []", auto)
- by (cases "c = 0", auto)
-lemma pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
- apply (induct p, auto simp add: pnormal_def)
- apply (case_tac "pnormalize p = []", auto)
- by (case_tac "a=0", auto)
-lemma pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
- unfolding pnormal_def length_greater_0_conv by blast
-lemma pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
- apply (induct p, auto)
- apply (case_tac "p = []", auto)
- apply (simp add: pnormal_def)
- by (rule pnormal_cons, auto)
-lemma pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
- using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
-
-text{*Tidier versions of finiteness of roots.*}
-
-lemma poly_roots_finite_set: "poly p \<noteq> poly [] ==> finite {x. poly p x = 0}"
-apply (auto simp add: poly_roots_finite)
-apply (rule_tac B = "{x::real. \<exists>n. (n::nat) < N & (x = j n) }" in finite_subset)
-apply (induct_tac [2] "N", auto)
-apply (subgoal_tac "{x::real. \<exists>na. na < Suc n & (x = j na) } = { (j n) } Un {x. \<exists>na. na < n & (x = j na) }")
-apply (auto simp add: less_Suc_eq)
-done
-
-text{*bound for polynomial.*}
-
-lemma poly_mono: "abs(x) \<le> k ==> abs(poly p x) \<le> poly (map abs p) k"
-apply (induct "p", auto)
-apply (rule_tac j = "abs a + abs (x * poly p x)" in real_le_trans)
-apply (rule abs_triangle_ineq)
-apply (auto intro!: mult_mono simp add: abs_mult)
-done
-
-lemma poly_Sing: "poly [c] x = c" by simp
-end
--- a/src/HOL/IsaMakefile Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/IsaMakefile Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
default: HOL
generate: HOL-Complex-Generate-HOL HOL-Complex-Generate-HOLLight
-images: HOL HOL-Algebra HOL-Complex HOL-Complex-Matrix HOL-Nominal \
+images: HOL-Plain HOL HOL-Algebra HOL-Complex HOL-Complex-Matrix HOL-Nominal \
HOL-Word TLA HOL4
#Note: keep targets sorted (except for HOL-Library and HOL-ex)
@@ -64,86 +64,180 @@
## HOL
-HOL: Pure $(OUT)/HOL
+HOL: HOL-Plain $(OUT)/HOL
+
+HOL-Plain: Pure $(OUT)/HOL-Plain
Pure:
@cd $(SRC)/Pure; $(ISATOOL) make Pure
-$(OUT)/HOL: $(OUT)/Pure $(SRC)/Provers/Arith/abel_cancel.ML \
- $(SRC)/Provers/Arith/assoc_fold.ML \
- $(SRC)/Provers/Arith/cancel_div_mod.ML \
- $(SRC)/Provers/Arith/cancel_numeral_factor.ML \
- $(SRC)/Provers/Arith/cancel_numerals.ML \
- $(SRC)/Provers/Arith/cancel_sums.ML \
- $(SRC)/Provers/Arith/combine_numerals.ML \
- $(SRC)/Provers/Arith/extract_common_term.ML \
- $(SRC)/Provers/Arith/fast_lin_arith.ML \
- $(SRC)/Tools/IsaPlanner/isand.ML $(SRC)/Tools/IsaPlanner/rw_inst.ML \
- $(SRC)/Tools/IsaPlanner/rw_tools.ML \
- $(SRC)/Tools/IsaPlanner/zipper.ML $(SRC)/Tools/induct.ML \
- $(SRC)/Tools/induct_tacs.ML $(SRC)/Provers/blast.ML \
- $(SRC)/Provers/clasimp.ML $(SRC)/Provers/classical.ML \
- $(SRC)/Provers/eqsubst.ML $(SRC)/Provers/hypsubst.ML \
- $(SRC)/Provers/order.ML $(SRC)/Provers/project_rule.ML \
- $(SRC)/Provers/quantifier1.ML $(SRC)/Provers/quasi.ML \
- $(SRC)/Provers/splitter.ML $(SRC)/Provers/trancl.ML \
- $(SRC)/Tools/Metis/metis.ML $(SRC)/Tools/code/code_funcgr.ML \
- $(SRC)/Tools/code/code_name.ML $(SRC)/Tools/code/code_target.ML \
- $(SRC)/Tools/code/code_thingol.ML $(SRC)/Tools/nbe.ML \
- $(SRC)/Tools/atomize_elim.ML $(SRC)/Tools/random_word.ML \
- $(SRC)/Tools/rat.ML Tools/TFL/casesplit.ML ATP_Linkup.thy \
- Arith_Tools.thy Code_Setup.thy Datatype.thy Divides.thy \
- Equiv_Relations.thy Extraction.thy Finite_Set.thy Fun.thy FunDef.thy \
- HOL.thy Hilbert_Choice.thy Inductive.thy Int.thy IntDiv.thy \
- Lattices.thy List.thy Main.thy Map.thy Nat.thy NatBin.thy \
- OrderedGroup.thy Orderings.thy Power.thy Predicate.thy \
- Product_Type.thy ROOT.ML Recdef.thy Record.thy Refute.thy \
- Relation.thy Relation_Power.thy Ring_and_Field.thy SAT.thy Set.thy \
- SetInterval.thy Sum_Type.thy Groebner_Basis.thy Tools/watcher.ML \
- Tools/Groebner_Basis/groebner.ML Tools/Groebner_Basis/misc.ML \
- Tools/Groebner_Basis/normalizer.ML \
- Tools/Groebner_Basis/normalizer_data.ML Tools/Qelim/cooper.ML \
- Tools/Qelim/langford_data.ML Tools/Qelim/langford.ML \
- Tools/Qelim/cooper_data.ML Tools/Qelim/ferrante_rackoff.ML \
- Tools/Qelim/ferrante_rackoff_data.ML Tools/Qelim/generated_cooper.ML \
- Tools/Qelim/presburger.ML Tools/Qelim/qelim.ML Tools/TFL/dcterm.ML \
- Tools/TFL/post.ML Tools/TFL/rules.ML Tools/TFL/tfl.ML \
- Tools/TFL/thms.ML Tools/TFL/thry.ML Tools/TFL/usyntax.ML \
- Tools/TFL/utils.ML Tools/cnf_funcs.ML Tools/datatype_abs_proofs.ML \
- Tools/datatype_aux.ML Tools/datatype_case.ML \
- Tools/datatype_codegen.ML Tools/datatype_package.ML \
- Tools/datatype_prop.ML Tools/datatype_realizer.ML \
- Tools/datatype_rep_proofs.ML Tools/dseq.ML \
- Tools/function_package/auto_term.ML \
- Tools/function_package/context_tree.ML \
- Tools/function_package/fundef_common.ML \
- Tools/function_package/fundef_core.ML \
- Tools/function_package/fundef_datatype.ML \
- Tools/function_package/fundef_lib.ML \
- Tools/function_package/fundef_package.ML \
- Tools/function_package/inductive_wrap.ML \
- Tools/function_package/lexicographic_order.ML \
- Tools/function_package/measure_functions.ML \
- Tools/function_package/mutual.ML \
- Tools/function_package/pattern_split.ML \
- Tools/function_package/size.ML Tools/inductive_codegen.ML \
- Tools/inductive_package.ML Tools/inductive_realizer.ML \
- Tools/inductive_set_package.ML Tools/lin_arith.ML Tools/meson.ML \
- Tools/metis_tools.ML Tools/numeral.ML Tools/numeral_syntax.ML \
- Tools/old_primrec_package.ML Tools/polyhash.ML \
- Tools/primrec_package.ML Tools/prop_logic.ML Tools/recdef_package.ML \
- Tools/recfun_codegen.ML Tools/record_package.ML Tools/refute.ML \
- Tools/refute_isar.ML Tools/res_atp.ML Tools/res_atp_methods.ML \
- Tools/res_atp_provers.ML Tools/res_axioms.ML Tools/res_clause.ML \
- Tools/res_hol_clause.ML Tools/res_reconstruct.ML \
- Tools/rewrite_hol_proof.ML Tools/sat_funcs.ML Tools/sat_solver.ML \
- Tools/specification_package.ML Tools/split_rule.ML \
- Tools/string_syntax.ML Tools/typecopy_package.ML \
- Tools/typedef_codegen.ML Tools/typedef_package.ML \
- Transitive_Closure.thy Typedef.thy Wellfounded.thy arith_data.ML \
- document/root.tex hologic.ML int_arith1.ML int_factor_simprocs.ML \
- nat_simprocs.ML simpdata.ML
- @$(ISATOOL) usedir $(HOL_USEDIR_OPTIONS) -b -g true $(OUT)/Pure HOL
+$(OUT)/HOL-Plain: $(OUT)/Pure \
+ plain.ML \
+ Code_Setup.thy \
+ Datatype.thy \
+ Divides.thy \
+ Extraction.thy \
+ Finite_Set.thy \
+ Fun.thy \
+ FunDef.thy \
+ HOL.thy \
+ Inductive.thy \
+ Lattices.thy \
+ Nat.thy \
+ OrderedGroup.thy \
+ Orderings.thy \
+ Plain.thy \
+ Power.thy \
+ Predicate.thy \
+ Product_Type.thy \
+ Record.thy \
+ Refute.thy \
+ Relation.thy \
+ Ring_and_Field.thy \
+ SAT.thy \
+ Set.thy \
+ Sum_Type.thy \
+ Tools/cnf_funcs.ML \
+ Tools/datatype_abs_proofs.ML \
+ Tools/datatype_aux.ML \
+ Tools/datatype_case.ML \
+ Tools/datatype_codegen.ML \
+ Tools/datatype_package.ML \
+ Tools/datatype_prop.ML \
+ Tools/datatype_realizer.ML \
+ Tools/datatype_rep_proofs.ML \
+ Tools/dseq.ML \
+ Tools/function_package/auto_term.ML \
+ Tools/function_package/context_tree.ML \
+ Tools/function_package/fundef_common.ML \
+ Tools/function_package/fundef_core.ML \
+ Tools/function_package/fundef_datatype.ML \
+ Tools/function_package/fundef_lib.ML \
+ Tools/function_package/fundef_package.ML \
+ Tools/function_package/induction_scheme.ML \
+ Tools/function_package/inductive_wrap.ML \
+ Tools/function_package/lexicographic_order.ML \
+ Tools/function_package/measure_functions.ML \
+ Tools/function_package/mutual.ML \
+ Tools/function_package/pattern_split.ML \
+ Tools/function_package/size.ML \
+ Tools/function_package/sum_tree.ML \
+ Tools/inductive_codegen.ML \
+ Tools/inductive_package.ML \
+ Tools/inductive_realizer.ML \
+ Tools/inductive_set_package.ML \
+ Tools/lin_arith.ML \
+ Tools/old_primrec_package.ML \
+ Tools/primrec_package.ML \
+ Tools/prop_logic.ML \
+ Tools/recfun_codegen.ML \
+ Tools/record_package.ML \
+ Tools/refute.ML \
+ Tools/refute_isar.ML \
+ Tools/rewrite_hol_proof.ML \
+ Tools/sat_funcs.ML \
+ Tools/sat_solver.ML \
+ Tools/split_rule.ML \
+ Tools/typecopy_package.ML \
+ Tools/typedef_codegen.ML \
+ Tools/typedef_package.ML \
+ Transitive_Closure.thy \
+ Typedef.thy \
+ Wellfounded.thy \
+ arith_data.ML \
+ hologic.ML \
+ simpdata.ML \
+ $(SRC)/Provers/Arith/abel_cancel.ML \
+ $(SRC)/Provers/Arith/cancel_div_mod.ML \
+ $(SRC)/Provers/Arith/cancel_sums.ML \
+ $(SRC)/Provers/Arith/fast_lin_arith.ML \
+ $(SRC)/Provers/blast.ML \
+ $(SRC)/Provers/clasimp.ML \
+ $(SRC)/Provers/classical.ML \
+ $(SRC)/Provers/eqsubst.ML \
+ $(SRC)/Provers/hypsubst.ML \
+ $(SRC)/Provers/order.ML \
+ $(SRC)/Provers/project_rule.ML \
+ $(SRC)/Provers/quantifier1.ML \
+ $(SRC)/Provers/splitter.ML \
+ $(SRC)/Provers/trancl.ML \
+ $(SRC)/Tools/IsaPlanner/isand.ML \
+ $(SRC)/Tools/IsaPlanner/rw_inst.ML \
+ $(SRC)/Tools/IsaPlanner/rw_tools.ML \
+ $(SRC)/Tools/IsaPlanner/zipper.ML \
+ $(SRC)/Tools/atomize_elim.ML \
+ $(SRC)/Tools/code/code_funcgr.ML \
+ $(SRC)/Tools/code/code_name.ML \
+ $(SRC)/Tools/code/code_target.ML \
+ $(SRC)/Tools/code/code_thingol.ML \
+ $(SRC)/Tools/induct.ML \
+ $(SRC)/Tools/induct_tacs.ML \
+ $(SRC)/Tools/nbe.ML \
+ $(SRC)/Tools/random_word.ML \
+ $(SRC)/Tools/rat.ML
+ @$(ISATOOL) usedir $(HOL_USEDIR_OPTIONS) -b -f plain.ML -g true $(OUT)/Pure HOL-Plain
+
+$(OUT)/HOL: $(OUT)/HOL-Plain \
+ ROOT.ML \
+ ATP_Linkup.thy \
+ Arith_Tools.thy \
+ Equiv_Relations.thy \
+ Groebner_Basis.thy \
+ Hilbert_Choice.thy \
+ Int.thy \
+ IntDiv.thy \
+ List.thy \
+ Main.thy \
+ Map.thy \
+ NatBin.thy \
+ Presburger.thy \
+ Recdef.thy \
+ Relation_Power.thy \
+ SetInterval.thy \
+ Tools/Groebner_Basis/groebner.ML \
+ Tools/Groebner_Basis/misc.ML \
+ Tools/Groebner_Basis/normalizer.ML \
+ Tools/Groebner_Basis/normalizer_data.ML \
+ Tools/Qelim/cooper.ML \
+ Tools/Qelim/cooper_data.ML \
+ Tools/Qelim/generated_cooper.ML \
+ Tools/Qelim/presburger.ML \
+ Tools/Qelim/qelim.ML \
+ Tools/TFL/casesplit.ML \
+ Tools/TFL/dcterm.ML \
+ Tools/TFL/post.ML \
+ Tools/TFL/rules.ML \
+ Tools/TFL/tfl.ML \
+ Tools/TFL/thms.ML \
+ Tools/TFL/thry.ML \
+ Tools/TFL/usyntax.ML \
+ Tools/TFL/utils.ML \
+ Tools/meson.ML \
+ Tools/metis_tools.ML \
+ Tools/numeral.ML \
+ Tools/numeral_syntax.ML \
+ Tools/polyhash.ML \
+ Tools/recdef_package.ML \
+ Tools/res_atp.ML \
+ Tools/res_atp_methods.ML \
+ Tools/res_atp_provers.ML \
+ Tools/res_axioms.ML \
+ Tools/res_clause.ML \
+ Tools/res_hol_clause.ML \
+ Tools/res_reconstruct.ML \
+ Tools/specification_package.ML \
+ Tools/string_syntax.ML \
+ Tools/watcher.ML \
+ int_arith1.ML \
+ int_factor_simprocs.ML \
+ nat_simprocs.ML \
+ $(SRC)/Provers/Arith/assoc_fold.ML \
+ $(SRC)/Provers/Arith/cancel_numeral_factor.ML \
+ $(SRC)/Provers/Arith/cancel_numerals.ML \
+ $(SRC)/Provers/Arith/combine_numerals.ML \
+ $(SRC)/Provers/Arith/extract_common_term.ML \
+ $(SRC)/Tools/Metis/metis.ML
+ @$(ISATOOL) usedir $(HOL_USEDIR_OPTIONS) -b -g true $(OUT)/HOL-Plain HOL
## HOL-Complex-HahnBanach
@@ -167,27 +261,66 @@
HOL-Complex: HOL $(OUT)/HOL-Complex
-$(OUT)/HOL-Complex: $(OUT)/HOL Complex/ROOT.ML \
- Library/Zorn.thy Library/Order_Relation.thy Real/ContNotDenum.thy \
- Real/float_arith.ML Real/Lubs.thy Real/PReal.thy \
- Real/RComplete.thy Real/Rational.thy Real/Real.thy Real/RealDef.thy \
- Real/RealPow.thy Real/RealVector.thy Real/rat_arith.ML \
- Real/real_arith.ML Hyperreal/StarDef.thy Hyperreal/Fact.thy \
- Hyperreal/HLog.thy Hyperreal/Filter.thy Hyperreal/HSeries.thy \
- Hyperreal/transfer.ML Hyperreal/HLim.thy Hyperreal/HSEQ.thy \
- Hyperreal/HTranscendental.thy Hyperreal/HDeriv.thy \
- Hyperreal/HyperDef.thy Hyperreal/HyperNat.thy \
- Hyperreal/Hyperreal.thy Hyperreal/Integration.thy Hyperreal/Lim.thy \
- Hyperreal/Log.thy Hyperreal/Ln.thy Hyperreal/MacLaurin.thy \
- Hyperreal/NatStar.thy Hyperreal/NSA.thy Hyperreal/NthRoot.thy \
- Library/Univ_Poly.thy Hyperreal/SEQ.thy Hyperreal/Series.thy \
- Hyperreal/Star.thy Hyperreal/Taylor.thy Hyperreal/FrechetDeriv.thy \
- Hyperreal/Deriv.thy Hyperreal/Transcendental.thy \
- Hyperreal/hypreal_arith.ML Complex/Complex_Main.thy \
- Complex/Fundamental_Theorem_Algebra.thy Complex/CLim.thy \
- Complex/CStar.thy Complex/Complex.thy Complex/NSCA.thy \
- Complex/NSComplex.thy Complex/document/root.tex \
- Library/Infinite_Set.thy Library/Parity.thy
+$(OUT)/HOL-Complex: $(OUT)/HOL \
+ Complex/ROOT.ML \
+ Complex/CLim.thy \
+ Complex/CStar.thy \
+ Complex/Complex.thy \
+ Complex/Complex_Main.thy \
+ Complex/Fundamental_Theorem_Algebra.thy \
+ Complex/NSCA.thy \
+ Complex/NSComplex.thy \
+ Hyperreal/Deriv.thy \
+ Hyperreal/Fact.thy \
+ Hyperreal/Filter.thy \
+ Hyperreal/HDeriv.thy \
+ Hyperreal/HLim.thy \
+ Hyperreal/HLog.thy \
+ Hyperreal/HSEQ.thy \
+ Hyperreal/HSeries.thy \
+ Hyperreal/HTranscendental.thy \
+ Hyperreal/HyperDef.thy \
+ Hyperreal/HyperNat.thy \
+ Hyperreal/Hyperreal.thy \
+ Hyperreal/Integration.thy \
+ Hyperreal/Lim.thy \
+ Hyperreal/Ln.thy \
+ Hyperreal/Log.thy \
+ Hyperreal/MacLaurin.thy \
+ Hyperreal/NSA.thy \
+ Hyperreal/NatStar.thy \
+ Hyperreal/NthRoot.thy \
+ Hyperreal/SEQ.thy \
+ Hyperreal/Series.thy \
+ Hyperreal/Star.thy \
+ Hyperreal/StarDef.thy \
+ Hyperreal/Taylor.thy \
+ Hyperreal/Transcendental.thy \
+ Hyperreal/hypreal_arith.ML \
+ Hyperreal/transfer.ML \
+ Library/Abstract_Rat.thy \
+ Library/Dense_Linear_Order.thy \
+ Library/GCD.thy \
+ Library/Infinite_Set.thy \
+ Library/Order_Relation.thy \
+ Library/Parity.thy \
+ Library/Univ_Poly.thy \
+ Library/Zorn.thy \
+ Real/ContNotDenum.thy \
+ Real/Lubs.thy \
+ Real/PReal.thy \
+ Real/RComplete.thy \
+ Real/Rational.thy \
+ Real/Real.thy \
+ Real/RealDef.thy \
+ Real/RealPow.thy \
+ Real/RealVector.thy \
+ Real/rat_arith.ML \
+ Real/real_arith.ML \
+ Tools/Qelim/ferrante_rackoff.ML \
+ Tools/Qelim/ferrante_rackoff_data.ML \
+ Tools/Qelim/langford.ML \
+ Tools/Qelim/langford_data.ML
@cd Complex; $(ISATOOL) usedir -b -g true $(OUT)/HOL HOL-Complex
@@ -211,26 +344,26 @@
$(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy \
Library/BigO.thy Library/Ramsey.thy Library/Efficient_Nat.thy \
- Library/Executable_Set.thy Library/Infinite_Set.thy \
- Library/Dense_Linear_Order.thy Library/FuncSet.thy \
+ Library/Executable_Set.thy \
+ Library/FuncSet.thy \
Library/Library.thy Library/List_Prefix.thy Library/State_Monad.thy \
Library/Multiset.thy Library/NatPair.thy Library/Permutation.thy \
Library/Primes.thy Library/Pocklington.thy Library/Quotient.thy \
Library/Quicksort.thy Library/Nat_Infinity.thy Library/Word.thy \
Library/README.html Library/Continuity.thy \
- Library/Nested_Environment.thy Library/Zorn.thy \
+ Library/Nested_Environment.thy \
Library/Library/ROOT.ML Library/Library/document/root.tex \
Library/Library/document/root.bib Library/While_Combinator.thy \
Library/Product_ord.thy Library/Char_nat.thy Library/Char_ord.thy \
Library/Option_ord.thy Library/Sublist_Order.thy \
Library/List_lexord.thy Library/Commutative_Ring.thy \
Library/comm_ring.ML Library/Coinductive_List.thy \
- Library/AssocList.thy Library/Parity.thy Library/GCD.thy \
+ Library/AssocList.thy \
Library/Binomial.thy Library/Eval.thy Library/Eval_Witness.thy \
Library/Code_Index.thy Library/Code_Char.thy \
Library/Code_Char_chr.thy Library/Code_Integer.thy \
- Library/Code_Message.thy Library/Abstract_Rat.thy \
- Library/Univ_Poly.thy Library/Numeral_Type.thy \
+ Library/Code_Message.thy \
+ Library/Numeral_Type.thy \
Library/Boolean_Algebra.thy Library/Countable.thy Library/RType.thy \
Library/Heap.thy Library/Heap_Monad.thy Library/Array.thy \
Library/Ref.thy Library/Imperative_HOL.thy Library/RBT.thy \
--- a/src/HOL/Library/Abstract_Rat.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Abstract_Rat.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Abstract rational numbers *}
theory Abstract_Rat
-imports GCD
+imports Plain GCD
begin
types Num = "int \<times> int"
--- a/src/HOL/Library/AssocList.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/AssocList.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Map operations implemented on association lists*}
theory AssocList
-imports Map
+imports Plain Map
begin
text {*
--- a/src/HOL/Library/BigO.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/BigO.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Big O notation *}
theory BigO
-imports Main SetsAndFunctions
+imports Plain Presburger SetsAndFunctions
begin
text {*
--- a/src/HOL/Library/Binomial.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Binomial.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Binomial Coefficients *}
theory Binomial
-imports ATP_Linkup
+imports Plain SetInterval
begin
text {* This development is based on the work of Andy Gordon and
--- a/src/HOL/Library/Boolean_Algebra.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Boolean_Algebra.thy Thu Jun 26 10:07:01 2008 +0200
@@ -8,7 +8,7 @@
header {* Boolean Algebras *}
theory Boolean_Algebra
-imports ATP_Linkup
+imports Plain
begin
locale boolean =
--- a/src/HOL/Library/Char_nat.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Char_nat.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Mapping between characters and natural numbers *}
theory Char_nat
-imports List
+imports Plain List
begin
text {* Conversions between nibbles and natural numbers in [0..15]. *}
--- a/src/HOL/Library/Char_ord.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Char_ord.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Order on characters *}
theory Char_ord
-imports Product_ord Char_nat
+imports Plain Product_ord Char_nat
begin
instantiation nibble :: linorder
--- a/src/HOL/Library/Code_Char.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Code_Char.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Code generation of pretty characters (and strings) *}
theory Code_Char
-imports List
+imports Plain List
begin
declare char.recs [code func del] char.cases [code func del]
--- a/src/HOL/Library/Code_Char_chr.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Code_Char_chr.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Code generation of pretty characters with character codes *}
theory Code_Char_chr
-imports Char_nat Code_Char Code_Integer
+imports Plain Char_nat Code_Char Code_Integer
begin
definition
--- a/src/HOL/Library/Code_Index.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Code_Index.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,7 @@
header {* Type of indices *}
theory Code_Index
-imports ATP_Linkup
+imports Plain Presburger
begin
text {*
--- a/src/HOL/Library/Code_Integer.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Code_Integer.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Pretty integer literals for code generation *}
theory Code_Integer
-imports ATP_Linkup
+imports Plain Presburger
begin
text {*
--- a/src/HOL/Library/Code_Message.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Code_Message.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,7 @@
header {* Monolithic strings (message strings) for code generation *}
theory Code_Message
-imports List
+imports Plain List
begin
subsection {* Datatype of messages *}
--- a/src/HOL/Library/Coinductive_List.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Coinductive_List.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Potentially infinite lists as greatest fixed-point *}
theory Coinductive_List
-imports List
+imports Plain List
begin
subsection {* List constructors over the datatype universe *}
--- a/src/HOL/Library/Commutative_Ring.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Commutative_Ring.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Proving equalities in commutative rings *}
theory Commutative_Ring
-imports List Parity
+imports Plain List Parity
uses ("comm_ring.ML")
begin
--- a/src/HOL/Library/Continuity.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Continuity.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Continuity and iterations (of set transformers) *}
theory Continuity
-imports ATP_Linkup
+imports Plain Relation_Power
begin
subsection {* Continuity for complete lattices *}
--- a/src/HOL/Library/Countable.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Countable.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Encoding (almost) everything into natural numbers *}
theory Countable
-imports Finite_Set List Hilbert_Choice
+imports Plain List Hilbert_Choice
begin
subsection {* The class of countable types *}
--- a/src/HOL/Library/Dense_Linear_Order.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Dense_Linear_Order.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
and a quantifier elimination procedure in Ferrante and Rackoff style *}
theory Dense_Linear_Order
-imports Arith_Tools
+imports Plain Presburger
uses
"~~/src/HOL/Tools/Qelim/qelim.ML"
"~~/src/HOL/Tools/Qelim/langford_data.ML"
--- a/src/HOL/Library/Efficient_Nat.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Implementation of natural numbers by target-language integers *}
theory Efficient_Nat
-imports Code_Integer Code_Index
+imports Plain Code_Integer Code_Index
begin
text {*
--- a/src/HOL/Library/Enum.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Enum.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Finite types as explicit enumerations *}
theory Enum
-imports Main
+imports Plain Map
begin
subsection {* Class @{text enum} *}
--- a/src/HOL/Library/Eval.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Eval.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,6 +7,7 @@
theory Eval
imports
+ Plain
RType
Code_Index (* this theory is just imported for a term_of setup *)
begin
--- a/src/HOL/Library/Eval_Witness.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Eval_Witness.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Evaluation Oracle with ML witnesses *}
theory Eval_Witness
-imports List
+imports Plain List
begin
text {*
--- a/src/HOL/Library/Executable_Set.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Executable_Set.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Implementation of finite sets by lists *}
theory Executable_Set
-imports List
+imports Plain List
begin
subsection {* Definitional rewrites *}
--- a/src/HOL/Library/FuncSet.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/FuncSet.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Pi and Function Sets *}
theory FuncSet
-imports Main
+imports Plain Hilbert_Choice
begin
definition
--- a/src/HOL/Library/GCD.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/GCD.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* The Greatest Common Divisor *}
theory GCD
-imports ATP_Linkup
+imports Plain Presburger
begin
text {*
--- a/src/HOL/Library/Heap.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Heap.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* A polymorphic heap based on cantor encodings *}
theory Heap
-imports Main Countable RType
+imports Plain List Countable RType
begin
subsection {* Representable types *}
--- a/src/HOL/Library/Infinite_Set.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Infinite_Set.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,9 +6,10 @@
header {* Infinite Sets and Related Concepts *}
theory Infinite_Set
-imports ATP_Linkup
+imports Plain SetInterval Hilbert_Choice
begin
+
subsection "Infinite Sets"
text {*
--- a/src/HOL/Library/LaTeXsugar.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/LaTeXsugar.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
(*<*)
theory LaTeXsugar
-imports List
+imports Plain List
begin
(* LOGIC *)
--- a/src/HOL/Library/Library.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Library.thy Thu Jun 26 10:07:01 2008 +0200
@@ -2,7 +2,6 @@
(*<*)
theory Library
imports
- Abstract_Rat
AssocList
BigO
Binomial
@@ -24,9 +23,7 @@
Executable_Set
"../Real/Float"
FuncSet
- GCD
Imperative_HOL
- Infinite_Set
ListVector
Multiset
NatPair
@@ -35,8 +32,6 @@
Numeral_Type
OptionalSugar
Option_ord
- Order_Relation
- Parity
Permutation
Primes
Quicksort
@@ -44,10 +39,8 @@
Ramsey
RBT
State_Monad
- Univ_Poly
While_Combinator
Word
- Zorn
begin
end
(*>*)
--- a/src/HOL/Library/ListVector.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/ListVector.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,7 @@
header "Lists as vectors"
theory ListVector
-imports Main
+imports Plain List
begin
text{* \noindent
--- a/src/HOL/Library/List_Prefix.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/List_Prefix.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* List prefixes and postfixes *}
theory List_Prefix
-imports List
+imports Plain List
begin
subsection {* Prefix order on lists *}
--- a/src/HOL/Library/List_lexord.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/List_lexord.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Lexicographic order on lists *}
theory List_lexord
-imports List
+imports Plain List
begin
instantiation list :: (ord) ord
--- a/src/HOL/Library/Multiset.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Multiset.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Multisets *}
theory Multiset
-imports List
+imports Plain List
begin
subsection {* The type of multisets *}
--- a/src/HOL/Library/NatPair.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/NatPair.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Pairs of Natural Numbers *}
theory NatPair
-imports ATP_Linkup
+imports Plain Presburger
begin
text{*
--- a/src/HOL/Library/Nat_Infinity.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Nat_Infinity.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Natural numbers with infinity *}
theory Nat_Infinity
-imports ATP_Linkup
+imports Plain Presburger
begin
subsection {* Type definition *}
--- a/src/HOL/Library/Nested_Environment.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Nested_Environment.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Nested environments *}
theory Nested_Environment
-imports List
+imports Plain List
begin
text {*
--- a/src/HOL/Library/Numeral_Type.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Numeral_Type.thy Thu Jun 26 10:07:01 2008 +0200
@@ -8,7 +8,7 @@
header "Numeral Syntax for Types"
theory Numeral_Type
- imports ATP_Linkup
+imports Plain Presburger
begin
subsection {* Preliminary lemmas *}
--- a/src/HOL/Library/Option_ord.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Option_ord.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Canonical order on option type *}
theory Option_ord
-imports ATP_Linkup
+imports Plain
begin
instantiation option :: (order) order
--- a/src/HOL/Library/Order_Relation.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Order_Relation.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,7 @@
header {* Orders as Relations *}
theory Order_Relation
-imports ATP_Linkup Hilbert_Choice
+imports Plain Hilbert_Choice ATP_Linkup
begin
text{* This prelude could be moved to theory Relation: *}
--- a/src/HOL/Library/Parity.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Parity.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Even and Odd for int and nat *}
theory Parity
-imports ATP_Linkup
+imports Plain Presburger
begin
class even_odd = type +
--- a/src/HOL/Library/Permutation.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Permutation.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,7 @@
header {* Permutations *}
theory Permutation
-imports Multiset
+imports Plain Multiset
begin
inductive
--- a/src/HOL/Library/Pocklington.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Pocklington.thy Thu Jun 26 10:07:01 2008 +0200
@@ -1,5 +1,5 @@
(* Title: HOL/Library/Pocklington.thy
- ID: $Id:
+ ID: $Id$
Author: Amine Chaieb
*)
@@ -7,7 +7,7 @@
theory Pocklington
-imports List Primes
+imports Plain List Primes
begin
definition modeq:: "nat => nat => nat => bool" ("(1[_ = _] '(mod _'))")
@@ -596,9 +596,7 @@
and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 * y1) (x2 * y2)"
and fS: "finite S" and Rfg: "\<forall>x\<in>S. R (h x) (g x)"
shows "R (fold (op *) h e S) (fold (op *) g e S)"
- using prems
- by -(rule finite_subset_induct,auto)
-
+ using fS by (rule finite_subset_induct) (insert assms, auto)
lemma nproduct_mod:
assumes fS: "finite S" and n0: "n \<noteq> 0"
@@ -621,11 +619,8 @@
lemma coprime_nproduct:
assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
shows "coprime n (setprod a S)"
- using fS Sn
-unfolding setprod_def
-apply -
-apply (rule finite_subset_induct)
-by (auto simp add: coprime_mul)
+ using fS unfolding setprod_def by (rule finite_subset_induct)
+ (insert Sn, auto simp add: coprime_mul)
lemma (in comm_monoid_mult)
fold_eq_general:
--- a/src/HOL/Library/Primes.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Primes.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Primality on nat *}
theory Primes
-imports GCD Parity
+imports Plain ATP_Linkup GCD Parity
begin
definition
--- a/src/HOL/Library/Product_ord.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Product_ord.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Order on product types *}
theory Product_ord
-imports ATP_Linkup
+imports Plain
begin
instantiation "*" :: (ord, ord) ord
--- a/src/HOL/Library/Quicksort.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Quicksort.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header{*Quicksort*}
theory Quicksort
-imports Multiset
+imports Plain Multiset
begin
context linorder
--- a/src/HOL/Library/Quotient.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Quotient.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Quotient types *}
theory Quotient
-imports ATP_Linkup Hilbert_Choice
+imports Plain Hilbert_Choice
begin
text {*
--- a/src/HOL/Library/RBT.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/RBT.thy Thu Jun 26 10:07:01 2008 +0200
@@ -8,7 +8,7 @@
(*<*)
theory RBT
-imports Main AssocList
+imports Plain AssocList
begin
datatype color = R | B
--- a/src/HOL/Library/RType.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/RType.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Reflecting Pure types into HOL *}
theory RType
-imports Main Code_Message Code_Index (* import all 'special code' types *)
+imports Plain List Code_Message Code_Index (* import all 'special code' types *)
begin
datatype "rtype" = RType message_string "rtype list"
--- a/src/HOL/Library/Ramsey.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Ramsey.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header "Ramsey's Theorem"
theory Ramsey
-imports ATP_Linkup Infinite_Set
+imports Plain Presburger Infinite_Set
begin
subsection {* Preliminaries *}
--- a/src/HOL/Library/SetsAndFunctions.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/SetsAndFunctions.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Operations on sets and functions *}
theory SetsAndFunctions
-imports ATP_Linkup
+imports Plain
begin
text {*
--- a/src/HOL/Library/State_Monad.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/State_Monad.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Combinators syntax for generic, open state monads (single threaded monads) *}
theory State_Monad
-imports List
+imports Plain List
begin
subsection {* Motivation *}
--- a/src/HOL/Library/Sublist_Order.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Sublist_Order.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Sublist Ordering *}
theory Sublist_Order
-imports Main
+imports Plain List
begin
text {*
--- a/src/HOL/Library/Univ_Poly.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Univ_Poly.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header{*Univariate Polynomials*}
theory Univ_Poly
-imports Main
+imports Plain "../List"
begin
text{* Application of polynomial as a function. *}
--- a/src/HOL/Library/While_Combinator.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/While_Combinator.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* A general ``while'' combinator *}
theory While_Combinator
-imports Main
+imports Plain Presburger
begin
text {*
--- a/src/HOL/Library/Word.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Library/Word.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Binary Words *}
theory Word
-imports List
+imports Main
begin
subsection {* Auxilary Lemmas *}
@@ -430,13 +430,13 @@
"bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2"
proof -
have "\<forall>l2. bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2"
- proof (induct l1,simp_all)
+ proof (induct l1, simp_all)
fix x xs
assume ind: "\<forall>l2. bv_to_nat (xs @ l2) = bv_to_nat xs * 2 ^ length l2 + bv_to_nat l2"
- show "\<forall>l2. bv_to_nat xs * 2 ^ length l2 + bitval x * 2 ^ (length xs + length l2) = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
+ show "\<forall>l2::bit list. bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
proof
fix l2
- show "bv_to_nat xs * 2 ^ length l2 + bitval x * 2 ^ (length xs + length l2) = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
+ show "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2"
proof -
have "(2::nat) ^ (length xs + length l2) = 2 ^ length xs * 2 ^ length l2"
by (induct "length xs",simp_all)
--- a/src/HOL/List.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/List.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* The datatype of finite lists *}
theory List
-imports ATP_Linkup
+imports Plain Relation_Power Presburger Recdef ATP_Linkup
uses "Tools/string_syntax.ML"
begin
--- a/src/HOL/MetisExamples/Abstraction.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/MetisExamples/Abstraction.thy Thu Jun 26 10:07:01 2008 +0200
@@ -5,7 +5,8 @@
Testing the metis method
*)
-theory Abstraction imports FuncSet
+theory Abstraction
+imports Main FuncSet
begin
(*For Christoph Benzmueller*)
@@ -180,9 +181,10 @@
CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
f \<in> pset cl \<inter> cl"
by auto
+
(*??no longer terminates, with combinators
by (metis Collect_mem_eq Int_def SigmaD2 UnCI Un_absorb1)
- --{*@{text Int_def} is redundant}
+ --{*@{text Int_def} is redundant*}
*)
ML{*ResAtp.problem_name := "Abstraction__CLF_eq_Collect_Int"*}
--- a/src/HOL/MetisExamples/Tarski.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/MetisExamples/Tarski.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,9 @@
header {* The Full Theorem of Tarski *}
-theory Tarski imports FuncSet begin
+theory Tarski
+imports Main FuncSet
+begin
(*Many of these higher-order problems appear to be impossible using the
current linkup. They often seem to need either higher-order unification
@@ -961,8 +963,10 @@
done
ML_command{*ResAtp.problem_name:="Tarski__intY1_func"*} (*ALL THEOREMS*)
-lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
-by (metis intY1_f_closed restrict_in_funcset)
+lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
+apply (rule restrict_in_funcset)
+apply (metis intY1_f_closed restrict_in_funcset)
+done
ML_command{*ResAtp.problem_name:="Tarski__intY1_mono"*} (*ALL THEOREMS*)
lemma (in Tarski) intY1_mono:
--- a/src/HOL/NumberTheory/Factorization.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/NumberTheory/Factorization.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,9 @@
header {* Fundamental Theorem of Arithmetic (unique factorization into primes) *}
-theory Factorization imports Primes Permutation begin
+theory Factorization
+imports Main Primes Permutation
+begin
subsection {* Definitions *}
--- a/src/HOL/NumberTheory/IntPrimes.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/NumberTheory/IntPrimes.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header {* Divisibility and prime numbers (on integers) *}
theory IntPrimes
-imports Primes
+imports Main Primes
begin
text {*
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Plain.thy Thu Jun 26 10:07:01 2008 +0200
@@ -0,0 +1,15 @@
+(* Title: HOL/Plain.thy
+ ID: $Id$
+
+Contains fundamental HOL tools and packages. Does not include Hilbert_Choice.
+*)
+
+header {* Plain HOL *}
+
+theory Plain
+imports Datatype FunDef Record SAT Extraction
+begin
+
+ML {* path_add "~~/src/HOL/Library" *}
+
+end
--- a/src/HOL/ROOT.ML Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/ROOT.ML Thu Jun 26 10:07:01 2008 +0200
@@ -1,7 +1,7 @@
(* Title: HOL/ROOT.ML
ID: $Id$
-Classical Higher-order Logic.
+Classical Higher-order Logic -- batteries included.
*)
use_thy "Main";
--- a/src/HOL/Real/ContNotDenum.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Real/ContNotDenum.thy Thu Jun 26 10:07:01 2008 +0200
@@ -6,7 +6,7 @@
header {* Non-denumerability of the Continuum. *}
theory ContNotDenum
-imports RComplete
+imports RComplete "../Hilbert_Choice"
begin
subsection {* Abstract *}
--- a/src/HOL/Real/Lubs.thy Thu Jun 26 10:06:54 2008 +0200
+++ b/src/HOL/Real/Lubs.thy Thu Jun 26 10:07:01 2008 +0200
@@ -7,7 +7,7 @@
header{*Definitions of Upper Bounds and Least Upper Bounds*}
theory Lubs
-imports Main
+imports Plain
begin
text{*Thanks to suggestions by James Margetson*}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/plain.ML Thu Jun 26 10:07:01 2008 +0200
@@ -0,0 +1,7 @@
+(* Title: HOL/plain.ML
+ ID: $Id$
+
+Classical Higher-order Logic -- plain Tool bootstrap.
+*)
+
+use_thy "Plain";