| author | wenzelm | 
| Wed, 03 May 2017 23:04:25 +0200 | |
| changeset 65703 | cead65c19f2e | 
| parent 63205 | 97b721666890 | 
| child 67398 | 5eb932e604a2 | 
| permissions | -rw-r--r-- | 
| 30439 | 1 | (* Title : HOL/Decision_Procs/Dense_Linear_Order.thy | 
| 29655 
ac31940cfb69
Plain, Main form meeting points in import hierarchy
 haftmann parents: 
29509diff
changeset | 2 | Author : Amine Chaieb, TU Muenchen | 
| 
ac31940cfb69
Plain, Main form meeting points in import hierarchy
 haftmann parents: 
29509diff
changeset | 3 | *) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 4 | |
| 60533 | 5 | section \<open>Dense linear order without endpoints | 
| 6 | and a quantifier elimination procedure in Ferrante and Rackoff style\<close> | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 7 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 8 | theory Dense_Linear_Order | 
| 30738 | 9 | imports Main | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 10 | begin | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 11 | |
| 48891 | 12 | ML_file "langford_data.ML" | 
| 13 | ML_file "ferrante_rackoff_data.ML" | |
| 14 | ||
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 15 | context linorder | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 16 | begin | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 17 | |
| 49519 | 18 | lemma less_not_permute[no_atp]: "\<not> (x < y \<and> y < x)" | 
| 19 | by (simp add: not_less linear) | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 20 | |
| 60539 | 21 | lemma gather_simps[no_atp]: | 
| 49519 | 22 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> | 
| 23 | (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)" | |
| 24 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> | |
| 25 | (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)" | |
| 26 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> | |
| 27 | (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))" | |
| 28 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> | |
| 29 | (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))" | |
| 30 | by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 31 | |
| 60539 | 32 | lemma gather_start [no_atp]: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)"
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 33 | by simp | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 34 | |
| 61586 | 35 | text\<open>Theorems for \<open>\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)\<close>\<close> | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 36 | lemma minf_lt[no_atp]: "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 37 | lemma minf_gt[no_atp]: "\<exists>z . \<forall>x. x < z \<longrightarrow> (t < x \<longleftrightarrow> False)" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 38 | by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 39 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 40 | lemma minf_le[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le) | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 41 | lemma minf_ge[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 42 | by (auto simp add: less_le not_less not_le) | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 43 | lemma minf_eq[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 44 | lemma minf_neq[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 45 | lemma minf_P[no_atp]: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 46 | |
| 61586 | 47 | text\<open>Theorems for \<open>\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)\<close>\<close> | 
| 49519 | 48 | lemma pinf_gt[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto | 
| 49 | lemma pinf_lt[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x < t \<longleftrightarrow> False)" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 50 | by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 51 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 52 | lemma pinf_ge[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le) | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 53 | lemma pinf_le[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 54 | by (auto simp add: less_le not_less not_le) | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 55 | lemma pinf_eq[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 56 | lemma pinf_neq[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 57 | lemma pinf_P[no_atp]: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 58 | |
| 60539 | 59 | lemma nmi_lt[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | 
| 60 | lemma nmi_gt[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 61 | by (auto simp add: le_less) | 
| 60539 | 62 | lemma nmi_le[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | 
| 63 | lemma nmi_ge[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 64 | lemma nmi_eq[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 65 | lemma nmi_neq[no_atp]: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 66 | lemma nmi_P[no_atp]: "\<forall>x. ~P \<and> P \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 67 | lemma nmi_conj[no_atp]: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists>u\<in> U. u \<le> x) ; | |
| 68 | \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists>u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow> | |
| 69 | \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 70 | lemma nmi_disj[no_atp]: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists>u\<in> U. u \<le> x) ; | |
| 71 | \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists>u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow> | |
| 72 | \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists>u\<in> U. u \<le> x)" by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 73 | |
| 60539 | 74 | lemma npi_lt[no_atp]: "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) | 
| 75 | lemma npi_gt[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 76 | lemma npi_le[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x \<le> t \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 77 | lemma npi_ge[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 78 | lemma npi_eq[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 79 | lemma npi_neq[no_atp]: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists>u\<in> U. x \<le> u )" by auto | |
| 80 | lemma npi_P[no_atp]: "\<forall>x. ~P \<and> P \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 81 | lemma npi_conj[no_atp]: "\<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> | |
| 82 | \<Longrightarrow> \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 83 | lemma npi_disj[no_atp]: "\<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> | |
| 84 | \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists>u\<in> U. x \<le> u)" by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 85 | |
| 49519 | 86 | lemma lin_dense_lt[no_atp]: | 
| 87 | "t \<in> U \<Longrightarrow> | |
| 60539 | 88 | \<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)" | 
| 89 | proof clarsimp | |
| 49519 | 90 | fix x l u y | 
| 60539 | 91 | assume tU: "t \<in> U" | 
| 92 | and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" | |
| 93 | and lx: "l < x" | |
| 94 | and xu: "x < u" | |
| 95 | and px: "x < t" | |
| 96 | and ly: "l < y" | |
| 97 | and yu: "y < u" | |
| 98 | from tU noU ly yu have tny: "t \<noteq> y" by auto | |
| 99 | have False if H: "t < y" | |
| 100 | proof - | |
| 101 | from less_trans[OF lx px] less_trans[OF H yu] have "l < t \<and> t < u" | |
| 102 | by simp | |
| 103 | with tU noU show ?thesis | |
| 104 | by auto | |
| 105 | qed | |
| 106 | then have "\<not> t < y" | |
| 107 | by auto | |
| 108 | then have "y \<le> t" | |
| 109 | by (simp add: not_less) | |
| 110 | then show "y < t" | |
| 111 | using tny by (simp add: less_le) | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 112 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 113 | |
| 49519 | 114 | lemma lin_dense_gt[no_atp]: | 
| 115 | "t \<in> U \<Longrightarrow> | |
| 60539 | 116 | \<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)" | 
| 117 | proof clarsimp | |
| 49519 | 118 | fix x l u y | 
| 60539 | 119 | assume tU: "t \<in> U" | 
| 120 | and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" | |
| 121 | and lx: "l < x" | |
| 122 | and xu: "x < u" | |
| 123 | and px: "t < x" | |
| 124 | and ly: "l < y" | |
| 125 | and yu: "y < u" | |
| 126 | from tU noU ly yu have tny: "t \<noteq> y" by auto | |
| 127 | have False if H: "y < t" | |
| 128 | proof - | |
| 129 | from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" | |
| 130 | by simp | |
| 131 | with tU noU show ?thesis | |
| 132 | by auto | |
| 133 | qed | |
| 134 | then have "\<not> y < t" | |
| 135 | by auto | |
| 136 | then have "t \<le> y" | |
| 137 | by (auto simp add: not_less) | |
| 138 | then show "t < y" | |
| 139 | using tny by (simp add: less_le) | |
| 49519 | 140 | qed | 
| 141 | ||
| 142 | lemma lin_dense_le[no_atp]: | |
| 143 | "t \<in> U \<Longrightarrow> | |
| 60539 | 144 | \<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)" | 
| 145 | proof clarsimp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 146 | fix x l u y | 
| 60539 | 147 | assume tU: "t \<in> U" | 
| 148 | and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" | |
| 149 | and lx: "l < x" | |
| 150 | and xu: "x < u" | |
| 151 | and px: "x \<le> t" | |
| 152 | and ly: "l < y" | |
| 153 | and yu: "y < u" | |
| 154 | from tU noU ly yu have tny: "t \<noteq> y" by auto | |
| 155 | have False if H: "t < y" | |
| 156 | proof - | |
| 49519 | 157 | from less_le_trans[OF lx px] less_trans[OF H yu] | 
| 158 | have "l < t \<and> t < u" by simp | |
| 60539 | 159 | with tU noU show ?thesis by auto | 
| 160 | qed | |
| 49519 | 161 | then have "\<not> t < y" by auto | 
| 162 | then show "y \<le> t" by (simp add: not_less) | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 163 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 164 | |
| 49519 | 165 | lemma lin_dense_ge[no_atp]: | 
| 166 | "t \<in> U \<Longrightarrow> | |
| 60539 | 167 | \<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)" | 
| 168 | proof clarsimp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 169 | fix x l u y | 
| 60539 | 170 | assume tU: "t \<in> U" | 
| 171 | and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" | |
| 172 | and lx: "l < x" | |
| 173 | and xu: "x < u" | |
| 174 | and px: "t \<le> x" | |
| 175 | and ly: "l < y" | |
| 176 | and yu: "y < u" | |
| 177 | from tU noU ly yu have tny: "t \<noteq> y" by auto | |
| 178 | have False if H: "y < t" | |
| 179 | proof - | |
| 49519 | 180 | from less_trans[OF ly H] le_less_trans[OF px xu] | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 181 | have "l < t \<and> t < u" by simp | 
| 60539 | 182 | with tU noU show ?thesis by auto | 
| 183 | qed | |
| 184 | then have "\<not> y < t" by auto | |
| 49519 | 185 | then show "t \<le> y" by (simp add: not_less) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 186 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 187 | |
| 49519 | 188 | lemma lin_dense_eq[no_atp]: | 
| 189 | "t \<in> U \<Longrightarrow> | |
| 60539 | 190 | \<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)" | 
| 49519 | 191 | by auto | 
| 192 | ||
| 193 | lemma lin_dense_neq[no_atp]: | |
| 194 | "t \<in> U \<Longrightarrow> | |
| 60539 | 195 | \<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)" | 
| 49519 | 196 | by auto | 
| 197 | ||
| 198 | lemma lin_dense_P[no_atp]: | |
| 60539 | 199 | "\<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)" | 
| 49519 | 200 | by auto | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 201 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 202 | lemma lin_dense_conj[no_atp]: | 
| 60539 | 203 | "\<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 | 
| 204 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> P1 y) ; | |
| 205 | \<forall>x l u. (\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> P2 x | |
| 206 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow> | |
| 207 | \<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) | |
| 208 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 209 | by blast | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 210 | lemma lin_dense_disj[no_atp]: | 
| 60539 | 211 | "\<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 | 
| 212 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> P1 y) ; | |
| 213 | \<forall>x l u. (\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> P2 x | |
| 214 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow> | |
| 215 | \<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) | |
| 216 | \<longrightarrow> (\<forall>y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 217 | by blast | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 218 | |
| 60539 | 219 | lemma npmibnd[no_atp]: "\<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> | 
| 220 | \<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')" | |
| 49519 | 221 | by auto | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 222 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 223 | lemma finite_set_intervals[no_atp]: | 
| 60539 | 224 | assumes px: "P x" | 
| 225 | and lx: "l \<le> x" | |
| 226 | and xu: "x \<le> u" | |
| 227 | and linS: "l\<in> S" | |
| 228 | and uinS: "u \<in> S" | |
| 229 | and fS:"finite S" | |
| 230 | and lS: "\<forall>x\<in> S. l \<le> x" | |
| 231 | and Su: "\<forall>x\<in> S. x \<le> u" | |
| 232 | 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" | |
| 49519 | 233 | proof - | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 234 |   let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 235 |   let ?xM = "{y. y\<in> S \<and> x \<le> y}"
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 236 | let ?a = "Max ?Mx" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 237 | let ?b = "Min ?xM" | 
| 60539 | 238 | have MxS: "?Mx \<subseteq> S" | 
| 239 | by blast | |
| 240 | then have fMx: "finite ?Mx" | |
| 241 | using fS finite_subset by auto | |
| 242 | from lx linS have linMx: "l \<in> ?Mx" | |
| 243 | by blast | |
| 244 |   then have Mxne: "?Mx \<noteq> {}"
 | |
| 245 | by blast | |
| 246 | have xMS: "?xM \<subseteq> S" | |
| 247 | by blast | |
| 248 | then have fxM: "finite ?xM" | |
| 249 | using fS finite_subset by auto | |
| 250 | from xu uinS have linxM: "u \<in> ?xM" | |
| 251 | by blast | |
| 252 |   then have xMne: "?xM \<noteq> {}"
 | |
| 253 | by blast | |
| 254 | have ax: "?a \<le> x" | |
| 255 | using Mxne fMx by auto | |
| 256 | have xb: "x \<le> ?b" | |
| 257 | using xMne fxM by auto | |
| 258 | have "?a \<in> ?Mx" | |
| 60708 | 259 | using Max_in[OF fMx Mxne] by simp | 
| 260 | then have ainS: "?a \<in> S" | |
| 261 | using MxS by blast | |
| 60539 | 262 | have "?b \<in> ?xM" | 
| 60708 | 263 | using Min_in[OF fxM xMne] by simp | 
| 264 | then have binS: "?b \<in> S" | |
| 265 | using xMS by blast | |
| 60539 | 266 | have noy: "\<forall>y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S" | 
| 267 | proof clarsimp | |
| 49519 | 268 | fix y | 
| 269 | assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S" | |
| 60708 | 270 | from yS have "y \<in> ?Mx \<or> y \<in> ?xM" | 
| 271 | by (auto simp add: linear) | |
| 60539 | 272 | then show False | 
| 273 | proof | |
| 49519 | 274 | assume "y \<in> ?Mx" | 
| 60708 | 275 | then have "y \<le> ?a" | 
| 276 | using Mxne fMx by auto | |
| 277 | with ay show ?thesis | |
| 278 | by (simp add: not_le[symmetric]) | |
| 60539 | 279 | next | 
| 49519 | 280 | assume "y \<in> ?xM" | 
| 60708 | 281 | then have "?b \<le> y" | 
| 282 | using xMne fxM by auto | |
| 283 | with yb show ?thesis | |
| 284 | by (simp add: not_le[symmetric]) | |
| 60539 | 285 | qed | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 286 | qed | 
| 60708 | 287 | from ainS binS noy ax xb px show ?thesis | 
| 288 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 289 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 290 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 291 | lemma finite_set_intervals2[no_atp]: | 
| 60539 | 292 | assumes px: "P x" | 
| 293 | and lx: "l \<le> x" | |
| 294 | and xu: "x \<le> u" | |
| 295 | and linS: "l\<in> S" | |
| 296 | and uinS: "u \<in> S" | |
| 297 | and fS: "finite S" | |
| 298 | and lS: "\<forall>x\<in> S. l \<le> x" | |
| 299 | and Su: "\<forall>x\<in> S. x \<le> u" | |
| 300 | 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)" | |
| 301 | proof - | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 302 | from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] | 
| 60539 | 303 | obtain a and b where as: "a \<in> S" and bs: "b \<in> S" | 
| 304 | and noS: "\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S" | |
| 305 | and axb: "a \<le> x \<and> x \<le> b \<and> P x" | |
| 306 | by auto | |
| 307 | from axb have "x = a \<or> x = b \<or> (a < x \<and> x < b)" | |
| 308 | by (auto simp add: le_less) | |
| 309 | then show ?thesis | |
| 310 | using px as bs noS by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 311 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 312 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 313 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 314 | |
| 49519 | 315 | |
| 60533 | 316 | section \<open>The classical QE after Langford for dense linear orders\<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 317 | |
| 53215 
5e47c31c6f7c
renamed typeclass dense_linorder to unbounded_dense_linorder
 hoelzl parents: 
53072diff
changeset | 318 | context unbounded_dense_linorder | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 319 | begin | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 320 | |
| 49519 | 321 | lemma interval_empty_iff: "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
 | 
| 27825 | 322 | by (auto dest: dense) | 
| 323 | ||
| 60539 | 324 | lemma dlo_qe_bnds[no_atp]: | 
| 325 |   assumes ne: "L \<noteq> {}"
 | |
| 326 |     and neU: "U \<noteq> {}"
 | |
| 327 | and fL: "finite L" | |
| 328 | and fU: "finite U" | |
| 329 | 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)" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 330 | proof (simp only: atomize_eq, rule iffI) | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 331 | assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" | 
| 60708 | 332 | then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" | 
| 333 | by blast | |
| 60539 | 334 | have "l < u" if l: "l \<in> L" and u: "u \<in> U" for l u | 
| 335 | proof - | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 336 | have "l < x" using xL l by blast | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 337 | also have "x < u" using xU u by blast | 
| 60539 | 338 | finally show ?thesis . | 
| 339 | qed | |
| 49519 | 340 | then show "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 341 | next | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 342 | assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 343 | let ?ML = "Max L" | 
| 60539 | 344 | let ?MU = "Min U" | 
| 60708 | 345 | from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" | 
| 346 | by auto | |
| 347 | from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" | |
| 348 | by auto | |
| 349 | from th1 th2 H have "?ML < ?MU" | |
| 350 | by auto | |
| 351 | with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" | |
| 352 | by blast | |
| 353 | from th3 th1' have "\<forall>l \<in> L. l < w" | |
| 354 | by auto | |
| 355 | moreover from th4 th2' have "\<forall>u \<in> U. w < u" | |
| 356 | by auto | |
| 357 | ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" | |
| 358 | by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 359 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 360 | |
| 60539 | 361 | lemma dlo_qe_noub[no_atp]: | 
| 362 |   assumes ne: "L \<noteq> {}"
 | |
| 363 | and fL: "finite L" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 364 |   shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
 | 
| 60539 | 365 | proof (simp add: atomize_eq) | 
| 60708 | 366 | from gt_ex[of "Max L"] obtain M where M: "Max L < M" | 
| 367 | by blast | |
| 368 | from ne fL have "\<forall>x \<in> L. x \<le> Max L" | |
| 369 | by simp | |
| 370 | with M have "\<forall>x\<in>L. x < M" | |
| 371 | by (auto intro: le_less_trans) | |
| 372 | then show "\<exists>x. \<forall>y\<in>L. y < x" | |
| 373 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 374 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 375 | |
| 60539 | 376 | lemma dlo_qe_nolb[no_atp]: | 
| 377 |   assumes ne: "U \<noteq> {}"
 | |
| 378 | and fU: "finite U" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 379 |   shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
 | 
| 60539 | 380 | proof (simp add: atomize_eq) | 
| 60708 | 381 | from lt_ex[of "Min U"] obtain M where M: "M < Min U" | 
| 382 | by blast | |
| 383 | from ne fU have "\<forall>x \<in> U. Min U \<le> x" | |
| 384 | by simp | |
| 385 | with M have "\<forall>x\<in>U. M < x" | |
| 386 | by (auto intro: less_le_trans) | |
| 387 | then show "\<exists>x. \<forall>y\<in>U. x < y" | |
| 388 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 389 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 390 | |
| 60539 | 391 | lemma exists_neq[no_atp]: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 392 | using gt_ex[of t] by auto | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 393 | |
| 60539 | 394 | lemmas dlo_simps[no_atp] = order_refl less_irrefl not_less not_le exists_neq | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 395 | le_less neq_iff linear less_not_permute | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 396 | |
| 53215 
5e47c31c6f7c
renamed typeclass dense_linorder to unbounded_dense_linorder
 hoelzl parents: 
53072diff
changeset | 397 | lemma axiom[no_atp]: "class.unbounded_dense_linorder (op \<le>) (op <)" | 
| 
5e47c31c6f7c
renamed typeclass dense_linorder to unbounded_dense_linorder
 hoelzl parents: 
53072diff
changeset | 398 | by (rule unbounded_dense_linorder_axioms) | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 399 | lemma atoms[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 400 | shows "TERM (less :: 'a \<Rightarrow> _)" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 401 | and "TERM (less_eq :: 'a \<Rightarrow> _)" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 402 | and "TERM (op = :: 'a \<Rightarrow> _)" . | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 403 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 404 | declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms] | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 405 | declare dlo_simps[langfordsimp] | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 406 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 407 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 408 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 409 | (* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *) | 
| 41849 | 410 | lemmas dnf[no_atp] = conj_disj_distribL conj_disj_distribR | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 411 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 412 | lemmas weak_dnf_simps[no_atp] = simp_thms dnf | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 413 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 414 | lemma nnf_simps[no_atp]: | 
| 60708 | 415 | "(\<not> (P \<and> Q)) \<longleftrightarrow> (\<not> P \<or> \<not> Q)" | 
| 416 | "(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)" | |
| 417 | "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)" | |
| 418 | "(P \<longleftrightarrow> Q) \<longleftrightarrow> ((P \<and> Q) \<or> (\<not> P \<and> \<not> Q))" | |
| 419 | "(\<not> \<not> P) \<longleftrightarrow> P" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 420 | by blast+ | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 421 | |
| 60708 | 422 | lemma ex_distrib[no_atp]: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" | 
| 423 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 424 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 425 | lemmas dnf_simps[no_atp] = weak_dnf_simps nnf_simps ex_distrib | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 426 | |
| 48891 | 427 | ML_file "langford.ML" | 
| 60533 | 428 | method_setup dlo = \<open> | 
| 55848 | 429 | Scan.succeed (SIMPLE_METHOD' o Langford.dlo_tac) | 
| 60533 | 430 | \<close> "Langford's algorithm for quantifier elimination in dense linear orders" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 431 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 432 | |
| 60533 | 433 | section \<open>Contructive dense linear orders yield QE for linear arithmetic over ordered Fields\<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 434 | |
| 60533 | 435 | text \<open>Linear order without upper bounds\<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 436 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 437 | locale linorder_stupid_syntax = linorder | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 438 | begin | 
| 49519 | 439 | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 440 | notation | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 441 |   less_eq  ("op \<sqsubseteq>") and
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 442 |   less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 443 |   less  ("op \<sqsubset>") and
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 444 |   less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 445 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 446 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 447 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 448 | locale linorder_no_ub = linorder_stupid_syntax + | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 449 | assumes gt_ex: "\<exists>y. less x y" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 450 | begin | 
| 49519 | 451 | |
| 60539 | 452 | lemma ge_ex[no_atp]: "\<exists>y. x \<sqsubseteq> y" | 
| 453 | using gt_ex by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 454 | |
| 61586 | 455 | text \<open>Theorems for \<open>\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)\<close>\<close> | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 456 | lemma pinf_conj[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 457 | assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 60539 | 458 | and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 459 | shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))" | 
| 60539 | 460 | proof - | 
| 60708 | 461 | from ex1 ex2 obtain z1 and z2 | 
| 462 | where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | |
| 463 | and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 464 | by blast | |
| 465 | from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" | |
| 466 | by blast | |
| 467 | from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" | |
| 468 | by simp_all | |
| 60539 | 469 | have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" if H: "z \<sqsubset> x" for x | 
| 470 | using less_trans[OF zz1 H] less_trans[OF zz2 H] z1 zz1 z2 zz2 by auto | |
| 60708 | 471 | then show ?thesis | 
| 472 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 473 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 474 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 475 | lemma pinf_disj[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 476 | assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 49519 | 477 | and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 478 | shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 479 | proof- | 
| 60708 | 480 | from ex1 ex2 obtain z1 and z2 | 
| 481 | where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | |
| 482 | and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 483 | by blast | |
| 484 | from gt_ex obtain z where z: "ord.max less_eq z1 z2 \<sqsubset> z" | |
| 485 | by blast | |
| 486 | from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" | |
| 487 | by simp_all | |
| 60539 | 488 | have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" if H: "z \<sqsubset> x" for x | 
| 489 | using less_trans[OF zz1 H] less_trans[OF zz2 H] z1 zz1 z2 zz2 by auto | |
| 60708 | 490 | then show ?thesis | 
| 491 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 492 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 493 | |
| 60539 | 494 | lemma pinf_ex[no_atp]: | 
| 495 | assumes ex: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" | |
| 496 | and p1: P1 | |
| 497 | shows "\<exists>x. P x" | |
| 49519 | 498 | proof - | 
| 60708 | 499 | from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" | 
| 500 | by blast | |
| 501 | from gt_ex obtain x where x: "z \<sqsubset> x" | |
| 502 | by blast | |
| 503 | from z x p1 show ?thesis | |
| 504 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 505 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 506 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 507 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 508 | |
| 60533 | 509 | text \<open>Linear order without upper bounds\<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 510 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 511 | locale linorder_no_lb = linorder_stupid_syntax + | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 512 | assumes lt_ex: "\<exists>y. less y x" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 513 | begin | 
| 49519 | 514 | |
| 60708 | 515 | lemma le_ex[no_atp]: "\<exists>y. y \<sqsubseteq> x" | 
| 516 | using lt_ex by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 517 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 518 | |
| 61586 | 519 | text \<open>Theorems for \<open>\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)\<close>\<close> | 
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 520 | lemma minf_conj[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 521 | assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 49519 | 522 | and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 523 | shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))" | 
| 60539 | 524 | proof - | 
| 60708 | 525 | from ex1 ex2 obtain z1 and z2 | 
| 526 | where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | |
| 60539 | 527 | and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | 
| 528 | by blast | |
| 60708 | 529 | from lt_ex obtain z where z: "z \<sqsubset> ord.min less_eq z1 z2" | 
| 530 | by blast | |
| 531 | from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" | |
| 532 | by simp_all | |
| 60539 | 533 | have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" if H: "x \<sqsubset> z" for x | 
| 534 | using less_trans[OF H zz1] less_trans[OF H zz2] z1 zz1 z2 zz2 by auto | |
| 60708 | 535 | then show ?thesis | 
| 536 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 537 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 538 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 539 | lemma minf_disj[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 540 | assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 49519 | 541 | and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 542 | shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))" | 
| 49519 | 543 | proof - | 
| 60708 | 544 | from ex1 ex2 obtain z1 and z2 | 
| 545 | where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | |
| 546 | and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 547 | by blast | |
| 548 | from lt_ex obtain z where z: "z \<sqsubset> ord.min less_eq z1 z2" | |
| 549 | by blast | |
| 550 | from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" | |
| 551 | by simp_all | |
| 60539 | 552 | have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" if H: "x \<sqsubset> z" for x | 
| 553 | using less_trans[OF H zz1] less_trans[OF H zz2] z1 zz1 z2 zz2 by auto | |
| 60708 | 554 | then show ?thesis | 
| 555 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 556 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 557 | |
| 49519 | 558 | lemma minf_ex[no_atp]: | 
| 559 | assumes ex: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" | |
| 560 | and p1: P1 | |
| 60539 | 561 | shows "\<exists>x. P x" | 
| 49519 | 562 | proof - | 
| 60708 | 563 | from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" | 
| 564 | by blast | |
| 565 | from lt_ex obtain x where x: "x \<sqsubset> z" | |
| 566 | by blast | |
| 567 | from z x p1 show ?thesis | |
| 568 | by blast | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 569 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 570 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 571 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 572 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 573 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34974diff
changeset | 574 | locale constr_dense_linorder = linorder_no_lb + linorder_no_ub + | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 575 | fixes between | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 576 | assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y" | 
| 49519 | 577 | and between_same: "between x x = x" | 
| 54868 | 578 | begin | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 579 | |
| 60539 | 580 | sublocale dlo: unbounded_dense_linorder | 
| 61166 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 wenzelm parents: 
61089diff
changeset | 581 | proof (unfold_locales, goal_cases) | 
| 60708 | 582 | case (1 x y) | 
| 583 | then show ?case | |
| 584 | using between_less [of x y] by auto | |
| 585 | next | |
| 586 | case 2 | |
| 587 | then show ?case by (rule lt_ex) | |
| 588 | next | |
| 589 | case 3 | |
| 590 | then show ?case by (rule gt_ex) | |
| 591 | qed | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 592 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 593 | lemma rinf_U[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 594 | assumes fU: "finite U" | 
| 60539 | 595 | 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 | 
| 596 | \<longrightarrow> (\<forall>y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )" | |
| 597 | 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')" | |
| 598 | and nmi: "\<not> MP" and npi: "\<not> PP" and ex: "\<exists>x. P x" | |
| 599 | shows "\<exists>u\<in> U. \<exists>u' \<in> U. P (between u u')" | |
| 49519 | 600 | proof - | 
| 60539 | 601 | from ex obtain x where px: "P x" | 
| 602 | by blast | |
| 603 | from px nmi npi nmpiU have "\<exists>u\<in> U. \<exists>u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" | |
| 604 | by auto | |
| 605 | 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'" | |
| 606 | by auto | |
| 607 |   from uU have Une: "U \<noteq> {}"
 | |
| 608 | by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 609 | let ?l = "linorder.Min less_eq U" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 610 | let ?u = "linorder.Max less_eq U" | 
| 60539 | 611 | have linM: "?l \<in> U" | 
| 612 | using fU Une by simp | |
| 613 | have uinM: "?u \<in> U" | |
| 614 | using fU Une by simp | |
| 615 | have lM: "\<forall>t\<in> U. ?l \<sqsubseteq> t" | |
| 616 | using Une fU by auto | |
| 617 | have Mu: "\<forall>t\<in> U. t \<sqsubseteq> ?u" | |
| 618 | using Une fU by auto | |
| 619 | have th: "?l \<sqsubseteq> u" | |
| 620 | using uU Une lM by auto | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 621 | from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" . | 
| 60539 | 622 | have th: "u' \<sqsubseteq> ?u" | 
| 623 | using uU' Une Mu by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 624 | from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" . | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 625 | from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu] | 
| 60539 | 626 | consider u where "u \<in> U" "P u" | | 
| 627 | t1 t2 where "t1 \<in> U" "t2 \<in> U" "\<forall>y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" "t1 \<sqsubset> x" "x \<sqsubset> t2" "P x" | |
| 628 | by blast | |
| 629 | then show ?thesis | |
| 630 | proof cases | |
| 60567 | 631 | case u: 1 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 632 | have "between u u = u" by (simp add: between_same) | 
| 60567 | 633 | with u have "P (between u u)" by simp | 
| 634 | with u show ?thesis by blast | |
| 60539 | 635 | next | 
| 636 | case 2 | |
| 637 | note t1M = \<open>t1 \<in> U\<close> and t2M = \<open>t2\<in> U\<close> | |
| 638 | and noM = \<open>\<forall>y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U\<close> | |
| 639 | and t1x = \<open>t1 \<sqsubset> x\<close> and xt2 = \<open>x \<sqsubset> t2\<close> | |
| 640 | and px = \<open>P x\<close> | |
| 49519 | 641 | from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" . | 
| 642 | let ?u = "between t1 t2" | |
| 643 | from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto | |
| 644 | from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast | |
| 60539 | 645 | with t1M t2M show ?thesis by blast | 
| 646 | qed | |
| 49519 | 647 | qed | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 648 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 649 | theorem fr_eq[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 650 | assumes fU: "finite U" | 
| 60539 | 651 | 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 | 
| 652 | \<longrightarrow> (\<forall>y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )" | |
| 653 | and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists>u\<in> U. u \<sqsubseteq> x)" | |
| 654 | and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists>u\<in> U. x \<sqsubseteq> u)" | |
| 49519 | 655 | 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)" | 
| 60539 | 656 | shows "(\<exists>x. P x) \<equiv> (MP \<or> PP \<or> (\<exists>u \<in> U. \<exists>u'\<in> U. P (between u u')))" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 657 | (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D") | 
| 49519 | 658 | proof - | 
| 60539 | 659 | have "?E \<longleftrightarrow> ?D" | 
| 660 | proof | |
| 661 | show ?D if px: ?E | |
| 662 | proof - | |
| 663 | consider "MP \<or> PP" | "\<not> MP" "\<not> PP" by blast | |
| 664 | then show ?thesis | |
| 665 | proof cases | |
| 666 | case 1 | |
| 667 | then show ?thesis by blast | |
| 668 | next | |
| 669 | case 2 | |
| 670 | from npmibnd[OF nmibnd npibnd] | |
| 671 | 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')" . | |
| 672 | from rinf_U[OF fU lin_dense nmpiU \<open>\<not> MP\<close> \<open>\<not> PP\<close> px] show ?thesis | |
| 673 | by blast | |
| 674 | qed | |
| 675 | qed | |
| 676 | show ?E if ?D | |
| 677 | proof - | |
| 678 | from that consider MP | PP | ?F by blast | |
| 679 | then show ?thesis | |
| 680 | proof cases | |
| 681 | case 1 | |
| 60567 | 682 | from minf_ex[OF mi this] show ?thesis . | 
| 60539 | 683 | next | 
| 684 | case 2 | |
| 60567 | 685 | from pinf_ex[OF pi this] show ?thesis . | 
| 60539 | 686 | next | 
| 687 | case 3 | |
| 688 | then show ?thesis by blast | |
| 689 | qed | |
| 690 | qed | |
| 691 | qed | |
| 692 | then show "?E \<equiv> ?D" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 693 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 694 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 695 | lemmas minf_thms[no_atp] = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 696 | lemmas pinf_thms[no_atp] = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 697 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 698 | lemmas nmi_thms[no_atp] = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 699 | lemmas npi_thms[no_atp] = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P | 
| 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 700 | lemmas lin_dense_thms[no_atp] = 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 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 701 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 702 | lemma ferrack_axiom[no_atp]: "constr_dense_linorder less_eq less between" | 
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34974diff
changeset | 703 | by (rule constr_dense_linorder_axioms) | 
| 49519 | 704 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35267diff
changeset | 705 | lemma atoms[no_atp]: | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 706 | shows "TERM (less :: 'a \<Rightarrow> _)" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 707 | and "TERM (less_eq :: 'a \<Rightarrow> _)" | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 708 | and "TERM (op = :: 'a \<Rightarrow> _)" . | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 709 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 710 | declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 711 | nmi: nmi_thms npi: npi_thms lindense: | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 712 | lin_dense_thms qe: fr_eq atoms: atoms] | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 713 | |
| 60533 | 714 | declaration \<open> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 715 | let | 
| 49519 | 716 |   fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
 | 
| 717 | fun generic_whatis phi = | |
| 718 | let | |
| 719 |       val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
 | |
| 720 | fun h x t = | |
| 59582 | 721 | case Thm.term_of t of | 
| 49519 | 722 |           Const(@{const_name HOL.eq}, _)$y$z =>
 | 
| 59582 | 723 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq | 
| 49519 | 724 | else Ferrante_Rackoff_Data.Nox | 
| 725 |        | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
 | |
| 59582 | 726 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq | 
| 49519 | 727 | else Ferrante_Rackoff_Data.Nox | 
| 728 | | b$y$z => if Term.could_unify (b, lt) then | |
| 59582 | 729 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt | 
| 730 | else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt | |
| 49519 | 731 | else Ferrante_Rackoff_Data.Nox | 
| 732 | else if Term.could_unify (b, le) then | |
| 59582 | 733 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le | 
| 734 | else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge | |
| 49519 | 735 | else Ferrante_Rackoff_Data.Nox | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 736 | else Ferrante_Rackoff_Data.Nox | 
| 49519 | 737 | | _ => Ferrante_Rackoff_Data.Nox | 
| 738 | in h end | |
| 61089 | 739 | fun ss phi ctxt = | 
| 740 | simpset_of (put_simpset HOL_ss ctxt addsimps (simps phi)) | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 741 | in | 
| 49519 | 742 |   Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
 | 
| 743 |     {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
 | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 744 | end | 
| 60533 | 745 | \<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 746 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 747 | end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 748 | |
| 48891 | 749 | ML_file "ferrante_rackoff.ML" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 750 | |
| 60533 | 751 | method_setup ferrack = \<open> | 
| 30549 | 752 | Scan.succeed (SIMPLE_METHOD' o FerranteRackoff.dlo_tac) | 
| 60533 | 753 | \<close> "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders" | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 754 | |
| 49519 | 755 | |
| 60533 | 756 | subsection \<open>Ferrante and Rackoff algorithm over ordered fields\<close> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 757 | |
| 60539 | 758 | lemma neg_prod_lt: | 
| 759 | fixes c :: "'a::linordered_field" | |
| 760 | assumes "c < 0" | |
| 761 | shows "c * x < 0 \<equiv> x > 0" | |
| 49519 | 762 | proof - | 
| 60539 | 763 | have "c * x < 0 \<longleftrightarrow> 0 / c < x" | 
| 764 | by (simp only: neg_divide_less_eq[OF \<open>c < 0\<close>] algebra_simps) | |
| 765 | also have "\<dots> \<longleftrightarrow> 0 < x" by simp | |
| 766 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 767 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 768 | |
| 60539 | 769 | lemma pos_prod_lt: | 
| 770 | fixes c :: "'a::linordered_field" | |
| 771 | assumes "c > 0" | |
| 772 | shows "c * x < 0 \<equiv> x < 0" | |
| 49519 | 773 | proof - | 
| 60539 | 774 | have "c * x < 0 \<longleftrightarrow> 0 /c > x" | 
| 775 | by (simp only: pos_less_divide_eq[OF \<open>c > 0\<close>] algebra_simps) | |
| 776 | also have "\<dots> \<longleftrightarrow> 0 > x" by simp | |
| 777 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 778 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 779 | |
| 60539 | 780 | lemma neg_prod_sum_lt: | 
| 781 | fixes c :: "'a::linordered_field" | |
| 782 | assumes "c < 0" | |
| 783 | shows "c * x + t < 0 \<equiv> x > (- 1 / c) * t" | |
| 49519 | 784 | proof - | 
| 60539 | 785 | have "c * x + t < 0 \<longleftrightarrow> c * x < - t" | 
| 786 | by (subst less_iff_diff_less_0 [of "c * x" "- t"]) simp | |
| 787 | also have "\<dots> \<longleftrightarrow> - t / c < x" | |
| 788 | by (simp only: neg_divide_less_eq[OF \<open>c < 0\<close>] algebra_simps) | |
| 789 | also have "\<dots> \<longleftrightarrow> (- 1 / c) * t < x" by simp | |
| 790 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 791 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 792 | |
| 60539 | 793 | lemma pos_prod_sum_lt: | 
| 794 | fixes c :: "'a::linordered_field" | |
| 795 | assumes "c > 0" | |
| 796 | shows "c * x + t < 0 \<equiv> x < (- 1 / c) * t" | |
| 49519 | 797 | proof - | 
| 60539 | 798 | have "c * x + t < 0 \<longleftrightarrow> c * x < - t" | 
| 799 | by (subst less_iff_diff_less_0 [of "c * x" "- t"]) simp | |
| 800 | also have "\<dots> \<longleftrightarrow> - t / c > x" | |
| 801 | by (simp only: pos_less_divide_eq[OF \<open>c > 0\<close>] algebra_simps) | |
| 802 | also have "\<dots> \<longleftrightarrow> (- 1 / c) * t > x" by simp | |
| 803 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 804 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 805 | |
| 60539 | 806 | lemma sum_lt: | 
| 807 | fixes x :: "'a::ordered_ab_group_add" | |
| 808 | shows "x + t < 0 \<equiv> x < - t" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 809 | using less_diff_eq[where a= x and b=t and c=0] by simp | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 810 | |
| 60539 | 811 | lemma neg_prod_le: | 
| 812 | fixes c :: "'a::linordered_field" | |
| 813 | assumes "c < 0" | |
| 814 | shows "c * x \<le> 0 \<equiv> x \<ge> 0" | |
| 49519 | 815 | proof - | 
| 60539 | 816 | have "c * x \<le> 0 \<longleftrightarrow> 0 / c \<le> x" | 
| 817 | by (simp only: neg_divide_le_eq[OF \<open>c < 0\<close>] algebra_simps) | |
| 818 | also have "\<dots> \<longleftrightarrow> 0 \<le> x" by simp | |
| 819 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 820 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 821 | |
| 60539 | 822 | lemma pos_prod_le: | 
| 823 | fixes c :: "'a::linordered_field" | |
| 824 | assumes "c > 0" | |
| 825 | shows "c * x \<le> 0 \<equiv> x \<le> 0" | |
| 49519 | 826 | proof - | 
| 60539 | 827 | have "c * x \<le> 0 \<longleftrightarrow> 0 / c \<ge> x" | 
| 828 | by (simp only: pos_le_divide_eq[OF \<open>c > 0\<close>] algebra_simps) | |
| 829 | also have "\<dots> \<longleftrightarrow> 0 \<ge> x" by simp | |
| 830 | finally show "PROP ?thesis" by simp | |
| 831 | qed | |
| 832 | ||
| 833 | lemma neg_prod_sum_le: | |
| 834 | fixes c :: "'a::linordered_field" | |
| 835 | assumes "c < 0" | |
| 836 | shows "c * x + t \<le> 0 \<equiv> x \<ge> (- 1 / c) * t" | |
| 837 | proof - | |
| 838 | have "c * x + t \<le> 0 \<longleftrightarrow> c * x \<le> -t" | |
| 839 | by (subst le_iff_diff_le_0 [of "c*x" "-t"]) simp | |
| 840 | also have "\<dots> \<longleftrightarrow> - t / c \<le> x" | |
| 841 | by (simp only: neg_divide_le_eq[OF \<open>c < 0\<close>] algebra_simps) | |
| 842 | also have "\<dots> \<longleftrightarrow> (- 1 / c) * t \<le> x" by simp | |
| 843 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 844 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 845 | |
| 60539 | 846 | lemma pos_prod_sum_le: | 
| 847 | fixes c :: "'a::linordered_field" | |
| 848 | assumes "c > 0" | |
| 849 | shows "c * x + t \<le> 0 \<equiv> x \<le> (- 1 / c) * t" | |
| 49519 | 850 | proof - | 
| 60539 | 851 | have "c * x + t \<le> 0 \<longleftrightarrow> c * x \<le> - t" | 
| 852 | by (subst le_iff_diff_le_0 [of "c*x" "-t"]) simp | |
| 853 | also have "\<dots> \<longleftrightarrow> - t / c \<ge> x" | |
| 854 | by (simp only: pos_le_divide_eq[OF \<open>c > 0\<close>] algebra_simps) | |
| 855 | also have "\<dots> \<longleftrightarrow> (- 1 / c) * t \<ge> x" by simp | |
| 856 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 857 | qed | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 858 | |
| 60539 | 859 | lemma sum_le: | 
| 860 | fixes x :: "'a::ordered_ab_group_add" | |
| 861 | shows "x + t \<le> 0 \<equiv> x \<le> - t" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 862 | using le_diff_eq[where a= x and b=t and c=0] by simp | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 863 | |
| 60539 | 864 | lemma nz_prod_eq: | 
| 865 | fixes c :: "'a::linordered_field" | |
| 866 | assumes "c \<noteq> 0" | |
| 867 | shows "c * x = 0 \<equiv> x = 0" | |
| 868 | using assms by simp | |
| 49519 | 869 | |
| 60539 | 870 | lemma nz_prod_sum_eq: | 
| 871 | fixes c :: "'a::linordered_field" | |
| 872 | assumes "c \<noteq> 0" | |
| 873 | shows "c * x + t = 0 \<equiv> x = (- 1/c) * t" | |
| 49519 | 874 | proof - | 
| 60539 | 875 | have "c * x + t = 0 \<longleftrightarrow> c * x = - t" | 
| 876 | by (subst eq_iff_diff_eq_0 [of "c*x" "-t"]) simp | |
| 877 | also have "\<dots> \<longleftrightarrow> x = - t / c" | |
| 878 | by (simp only: nonzero_eq_divide_eq[OF \<open>c \<noteq> 0\<close>] algebra_simps) | |
| 879 | finally show "PROP ?thesis" by simp | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 880 | qed | 
| 49519 | 881 | |
| 60539 | 882 | lemma sum_eq: | 
| 883 | fixes x :: "'a::ordered_ab_group_add" | |
| 884 | shows "x + t = 0 \<equiv> x = - t" | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 885 | using eq_diff_eq[where a= x and b=t and c=0] by simp | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 886 | |
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35028diff
changeset | 887 | interpretation class_dense_linordered_field: constr_dense_linorder | 
| 60539 | 888 | "op \<le>" "op <" "\<lambda>x y. 1/2 * ((x::'a::linordered_field) + y)" | 
| 49519 | 889 | by unfold_locales (dlo, dlo, auto) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 890 | |
| 60539 | 891 | declaration \<open> | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 892 | let | 
| 49519 | 893 | fun earlier [] x y = false | 
| 894 | | earlier (h::t) x y = | |
| 895 | if h aconvc y then false else if h aconvc x then true else earlier t x y; | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 896 | |
| 49519 | 897 | fun dest_frac ct = | 
| 59582 | 898 | case Thm.term_of ct of | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59586diff
changeset | 899 |     Const (@{const_name Rings.divide},_) $ a $ b=>
 | 
| 63201 | 900 | Rat.make (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | 
| 901 |   | Const(@{const_name inverse}, _)$a => Rat.make(1, HOLogic.dest_number a |> snd)
 | |
| 902 | | t => Rat.of_int (snd (HOLogic.dest_number t)) | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 903 | |
| 59582 | 904 | fun whatis x ct = case Thm.term_of ct of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 905 |   Const(@{const_name Groups.plus}, _)$(Const(@{const_name Groups.times},_)$_$y)$_ =>
 | 
| 59582 | 906 |      if y aconv Thm.term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 907 |      else ("Nox",[])
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 908 | | Const(@{const_name Groups.plus}, _)$y$_ =>
 | 
| 59582 | 909 |      if y aconv Thm.term_of x then ("x+t",[Thm.dest_arg ct])
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 910 |      else ("Nox",[])
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 911 | | Const(@{const_name Groups.times}, _)$_$y =>
 | 
| 59582 | 912 |      if y aconv Thm.term_of x then ("c*x",[Thm.dest_arg1 ct])
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 913 |      else ("Nox",[])
 | 
| 59582 | 914 | | t => if t aconv Thm.term_of x then ("x",[]) else ("Nox",[]);
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 915 | |
| 36945 | 916 | fun xnormalize_conv ctxt [] ct = Thm.reflexive ct | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 917 | | xnormalize_conv ctxt (vs as (x::_)) ct = | 
| 59582 | 918 | case Thm.term_of ct of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 919 |    Const(@{const_name Orderings.less},_)$_$Const(@{const_name Groups.zero},_) =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 920 | (case whatis x (Thm.dest_arg1 ct) of | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 921 |     ("c*x+t",[c,t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 922 | let | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 923 | val cr = dest_frac c | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 924 | val clt = Thm.dest_fun2 ct | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 925 | val cz = Thm.dest_arg ct | 
| 63205 | 926 | val neg = cr < @0 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 927 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 928 |                (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 929 | (if neg then Thm.apply (Thm.apply clt c) cz | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 930 | else Thm.apply (Thm.apply clt cz) c)) | 
| 36945 | 931 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 60801 | 932 | val th = Thm.implies_elim (Thm.instantiate' [SOME (Thm.ctyp_of_cterm x)] (map SOME [c,x,t]) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 933 |              (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 934 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 935 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 936 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 937 |     | ("x+t",[t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 938 | let | 
| 59586 | 939 | val T = Thm.ctyp_of_cterm x | 
| 60801 | 940 |         val th = Thm.instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 941 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 942 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 943 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 944 |     | ("c*x",[c]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 945 | let | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 946 | val cr = dest_frac c | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 947 | val clt = Thm.dest_fun2 ct | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 948 | val cz = Thm.dest_arg ct | 
| 63205 | 949 | val neg = cr < @0 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 950 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 951 |                (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 952 | (if neg then Thm.apply (Thm.apply clt c) cz | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 953 | else Thm.apply (Thm.apply clt cz) c)) | 
| 36945 | 954 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 60801 | 955 | val th = Thm.implies_elim (Thm.instantiate' [SOME (Thm.ctyp_of_cterm x)] (map SOME [c,x]) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 956 |              (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 957 | val rth = th | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 958 | in rth end | 
| 36945 | 959 | | _ => Thm.reflexive ct) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 960 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 961 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 962 | |  Const(@{const_name Orderings.less_eq},_)$_$Const(@{const_name Groups.zero},_) =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 963 | (case whatis x (Thm.dest_arg1 ct) of | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 964 |     ("c*x+t",[c,t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 965 | let | 
| 61075 | 966 | val T = Thm.typ_of_cterm x | 
| 967 | val cT = Thm.ctyp_of_cterm x | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 968 | val cr = dest_frac c | 
| 61075 | 969 |         val clt = Thm.cterm_of ctxt (Const (@{const_name ord_class.less}, T --> T --> @{typ bool}))
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 970 | val cz = Thm.dest_arg ct | 
| 63205 | 971 | val neg = cr < @0 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 972 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 973 |                (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 974 | (if neg then Thm.apply (Thm.apply clt c) cz | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 975 | else Thm.apply (Thm.apply clt cz) c)) | 
| 36945 | 976 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 61075 | 977 | val th = Thm.implies_elim (Thm.instantiate' [SOME cT] (map SOME [c,x,t]) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 978 |              (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 979 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 980 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 981 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 982 |     | ("x+t",[t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 983 | let | 
| 59586 | 984 | val T = Thm.ctyp_of_cterm x | 
| 60801 | 985 |         val th = Thm.instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 986 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 987 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 988 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 989 |     | ("c*x",[c]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 990 | let | 
| 61075 | 991 | val T = Thm.typ_of_cterm x | 
| 992 | val cT = Thm.ctyp_of_cterm x | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 993 | val cr = dest_frac c | 
| 61075 | 994 |         val clt = Thm.cterm_of ctxt (Const (@{const_name ord_class.less}, T --> T --> @{typ bool}))
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 995 | val cz = Thm.dest_arg ct | 
| 63205 | 996 | val neg = cr < @0 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 997 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 998 |                (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 999 | (if neg then Thm.apply (Thm.apply clt c) cz | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 1000 | else Thm.apply (Thm.apply clt cz) c)) | 
| 36945 | 1001 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 60801 | 1002 | val th = Thm.implies_elim (Thm.instantiate' [SOME (Thm.ctyp_of_cterm x)] (map SOME [c,x]) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1003 |              (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1004 | val rth = th | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1005 | in rth end | 
| 36945 | 1006 | | _ => Thm.reflexive ct) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1007 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38549diff
changeset | 1008 | |  Const(@{const_name HOL.eq},_)$_$Const(@{const_name Groups.zero},_) =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1009 | (case whatis x (Thm.dest_arg1 ct) of | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1010 |     ("c*x+t",[c,t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1011 | let | 
| 59586 | 1012 | val T = Thm.ctyp_of_cterm x | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1013 | val cr = dest_frac c | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1014 | val ceq = Thm.dest_fun2 ct | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1015 | val cz = Thm.dest_arg ct | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 1016 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 1017 |             (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 1018 |              (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply ceq c) cz)))
 | 
| 36945 | 1019 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 1020 | val th = Thm.implies_elim | |
| 60801 | 1021 |                  (Thm.instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1022 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 1023 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1024 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1025 |     | ("x+t",[t]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1026 | let | 
| 59586 | 1027 | val T = Thm.ctyp_of_cterm x | 
| 60801 | 1028 |         val th = Thm.instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1029 | val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv | 
| 36753 
5cf4e9128f22
renamed Normalizer to the more specific Semiring_Normalizer
 haftmann parents: 
36635diff
changeset | 1030 | (Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1031 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1032 |     | ("c*x",[c]) =>
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1033 | let | 
| 59586 | 1034 | val T = Thm.ctyp_of_cterm x | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1035 | val cr = dest_frac c | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1036 | val ceq = Thm.dest_fun2 ct | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1037 | val cz = Thm.dest_arg ct | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 1038 | val cthp = Simplifier.rewrite ctxt | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 1039 |             (Thm.apply @{cterm "Trueprop"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45620diff
changeset | 1040 |              (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply ceq c) cz)))
 | 
| 36945 | 1041 | val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI | 
| 1042 | val rth = Thm.implies_elim | |
| 60801 | 1043 |                  (Thm.instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1044 | in rth end | 
| 36945 | 1045 | | _ => Thm.reflexive ct); | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1046 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1047 | local | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1048 |   val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1049 |   val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1050 |   val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
 | 
| 52090 | 1051 |   val ss = simpset_of @{context}
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1052 | in | 
| 59582 | 1053 | fun field_isolate_conv phi ctxt vs ct = case Thm.term_of ct of | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 1054 |   Const(@{const_name Orderings.less},_)$a$b =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1055 | let val (ca,cb) = Thm.dest_binop ct | 
| 59586 | 1056 | val T = Thm.ctyp_of_cterm ca | 
| 60801 | 1057 | val th = Thm.instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1058 | val nth = Conv.fconv_rule | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1059 | (Conv.arg_conv (Conv.arg1_conv | 
| 52090 | 1060 | (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th | 
| 36945 | 1061 | val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1062 | in rth end | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 1063 | | Const(@{const_name Orderings.less_eq},_)$a$b =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1064 | let val (ca,cb) = Thm.dest_binop ct | 
| 59586 | 1065 | val T = Thm.ctyp_of_cterm ca | 
| 60801 | 1066 | val th = Thm.instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1067 | val nth = Conv.fconv_rule | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1068 | (Conv.arg_conv (Conv.arg1_conv | 
| 52090 | 1069 | (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th | 
| 36945 | 1070 | val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1071 | in rth end | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1072 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38549diff
changeset | 1073 | | Const(@{const_name HOL.eq},_)$a$b =>
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1074 | let val (ca,cb) = Thm.dest_binop ct | 
| 59586 | 1075 | val T = Thm.ctyp_of_cterm ca | 
| 60801 | 1076 | val th = Thm.instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1077 | val nth = Conv.fconv_rule | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1078 | (Conv.arg_conv (Conv.arg1_conv | 
| 52090 | 1079 | (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th | 
| 36945 | 1080 | val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1081 | in rth end | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38549diff
changeset | 1082 | | @{term "Not"} $(Const(@{const_name HOL.eq},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
 | 
| 36945 | 1083 | | _ => Thm.reflexive ct | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1084 | end; | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1085 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1086 | fun classfield_whatis phi = | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1087 | let | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1088 | fun h x t = | 
| 59582 | 1089 | case Thm.term_of t of | 
| 1090 |      Const(@{const_name HOL.eq}, _)$y$z =>
 | |
| 1091 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq | |
| 1092 | else Ferrante_Rackoff_Data.Nox | |
| 1093 |    | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
 | |
| 1094 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq | |
| 1095 | else Ferrante_Rackoff_Data.Nox | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 1096 |    | Const(@{const_name Orderings.less},_)$y$z =>
 | 
| 59582 | 1097 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt | 
| 1098 | else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt | |
| 1099 | else Ferrante_Rackoff_Data.Nox | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 1100 |    | Const (@{const_name Orderings.less_eq},_)$y$z =>
 | 
| 59582 | 1101 | if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le | 
| 1102 | else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge | |
| 1103 | else Ferrante_Rackoff_Data.Nox | |
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1104 | | _ => Ferrante_Rackoff_Data.Nox | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1105 | in h end; | 
| 61089 | 1106 | fun class_field_ss phi ctxt = | 
| 1107 | simpset_of (put_simpset HOL_basic_ss ctxt | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 1108 |     addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49756diff
changeset | 1109 |     |> fold Splitter.add_split [@{thm "abs_split"}, @{thm "split_max"}, @{thm "split_min"}])
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1110 | |
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1111 | in | 
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35028diff
changeset | 1112 | Ferrante_Rackoff_Data.funs @{thm "class_dense_linordered_field.ferrack_axiom"}
 | 
| 26161 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1113 |   {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
 | 
| 
34cb0b457dcc
Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
 chaieb parents: diff
changeset | 1114 | end | 
| 60533 | 1115 | \<close> | 
| 29841 
86d94bb79226
A generic decision procedure for linear rea arithmetic and normed vector spaces
 chaieb parents: 
29823diff
changeset | 1116 | |
| 60539 | 1117 | end |