| author | haftmann | 
| Fri, 22 Feb 2008 12:01:57 +0100 | |
| changeset 26113 | ba5909699cc3 | 
| parent 26040 | 08d52e2dba07 | 
| child 26733 | 47224a933c14 | 
| permissions | -rw-r--r-- | 
| 23453 | 1 | (* | 
| 2 | ID: $Id$ | |
| 3 | Author: Amine Chaieb, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 23470 | 6 | header {* Dense linear order without endpoints
 | 
| 23453 | 7 | and a quantifier elimination procedure in Ferrante and Rackoff style *} | 
| 8 | ||
| 9 | theory Dense_Linear_Order | |
| 10 | imports Finite_Set | |
| 11 | uses | |
| 23466 
886655a150f6
moved quantifier elimination tools to Tools/Qelim/;
 wenzelm parents: 
23453diff
changeset | 12 | "Tools/Qelim/qelim.ML" | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 13 | "Tools/Qelim/langford_data.ML" | 
| 23466 
886655a150f6
moved quantifier elimination tools to Tools/Qelim/;
 wenzelm parents: 
23453diff
changeset | 14 | "Tools/Qelim/ferrante_rackoff_data.ML" | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 15 |   ("Tools/Qelim/langford.ML")
 | 
| 23466 
886655a150f6
moved quantifier elimination tools to Tools/Qelim/;
 wenzelm parents: 
23453diff
changeset | 16 |   ("Tools/Qelim/ferrante_rackoff.ML")
 | 
| 23453 | 17 | begin | 
| 18 | ||
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 19 | setup Langford_Data.setup | 
| 23453 | 20 | setup Ferrante_Rackoff_Data.setup | 
| 21 | ||
| 24344 | 22 | context linorder | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 23 | begin | 
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 24 | |
| 25062 | 25 | lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear) | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 26 | |
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 27 | lemma gather_simps: | 
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 28 | shows | 
| 25062 | 29 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)" | 
| 30 | and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)" | |
| 31 | "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))" | |
| 32 | and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))" by auto | |
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 33 | |
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 34 | lemma | 
| 25062 | 35 |   gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
 | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 36 | by simp | 
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 37 | |
| 25062 | 38 | text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
 | 
| 39 | lemma minf_lt: "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto | |
| 40 | lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow> (t < x \<longleftrightarrow> False)" | |
| 23453 | 41 | by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) | 
| 42 | ||
| 25062 | 43 | lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le) | 
| 44 | lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)" | |
| 23453 | 45 | by (auto simp add: less_le not_less not_le) | 
| 25062 | 46 | lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto | 
| 47 | lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto | |
| 48 | lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast | |
| 23453 | 49 | |
| 25062 | 50 | text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
 | 
| 51 | lemma pinf_gt: "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto | |
| 52 | lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow> (x < t \<longleftrightarrow> False)" | |
| 23453 | 53 | by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) | 
| 54 | ||
| 25062 | 55 | lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le) | 
| 56 | lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)" | |
| 23453 | 57 | by (auto simp add: less_le not_less not_le) | 
| 25062 | 58 | lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto | 
| 59 | lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto | |
| 60 | lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast | |
| 23453 | 61 | |
| 25062 | 62 | lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | 
| 63 | lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" | |
| 23453 | 64 | by (auto simp add: le_less) | 
| 25062 | 65 | lemma nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | 
| 66 | lemma nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 67 | lemma nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 68 | lemma nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 69 | lemma nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 70 | lemma nmi_conj: "\<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' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 73 | lemma nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x) ; | |
| 74 | \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow> | |
| 75 | \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists> u\<in> U. u \<le> x)" by auto | |
| 23453 | 76 | |
| 25062 | 77 | lemma npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x < t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less) | 
| 78 | lemma npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 79 | lemma npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x \<le> t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 80 | lemma npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 81 | lemma npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> x = t \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 82 | lemma npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow> (\<exists> u\<in> U. x \<le> u )" by auto | |
| 83 | lemma npi_P: "\<forall> x. ~P \<and> P \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 84 | lemma npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk> | |
| 85 | \<Longrightarrow> \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 86 | lemma npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk> | |
| 87 | \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow> (\<exists> u\<in> U. x \<le> u)" by auto | |
| 23453 | 88 | |
| 25062 | 89 | lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)" | 
| 23453 | 90 | proof(clarsimp) | 
| 25062 | 91 | fix x l u y assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" | 
| 92 | and xu: "x<u" and px: "x < t" and ly: "l<y" and yu:"y < u" | |
| 23453 | 93 | from tU noU ly yu have tny: "t\<noteq>y" by auto | 
| 25062 | 94 |   {assume H: "t < y"
 | 
| 23453 | 95 | from less_trans[OF lx px] less_trans[OF H yu] | 
| 25062 | 96 | have "l < t \<and> t < u" by simp | 
| 23453 | 97 | with tU noU have "False" by auto} | 
| 25062 | 98 | hence "\<not> t < y" by auto hence "y \<le> t" by (simp add: not_less) | 
| 99 | thus "y < t" using tny by (simp add: less_le) | |
| 23453 | 100 | qed | 
| 101 | ||
| 25062 | 102 | lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)" | 
| 23453 | 103 | proof(clarsimp) | 
| 104 | fix x l u y | |
| 25062 | 105 | assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u" | 
| 106 | and px: "t < x" and ly: "l<y" and yu:"y < u" | |
| 23453 | 107 | from tU noU ly yu have tny: "t\<noteq>y" by auto | 
| 25062 | 108 |   {assume H: "y< t"
 | 
| 109 | from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp | |
| 23453 | 110 | with tU noU have "False" by auto} | 
| 25062 | 111 | hence "\<not> y<t" by auto hence "t \<le> y" by (auto simp add: not_less) | 
| 112 | thus "t < y" using tny by (simp add:less_le) | |
| 23453 | 113 | qed | 
| 114 | ||
| 25062 | 115 | lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)" | 
| 23453 | 116 | proof(clarsimp) | 
| 117 | fix x l u y | |
| 25062 | 118 | assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u" | 
| 119 | and px: "x \<le> t" and ly: "l<y" and yu:"y < u" | |
| 23453 | 120 | from tU noU ly yu have tny: "t\<noteq>y" by auto | 
| 25062 | 121 |   {assume H: "t < y"
 | 
| 23453 | 122 | from less_le_trans[OF lx px] less_trans[OF H yu] | 
| 25062 | 123 | have "l < t \<and> t < u" by simp | 
| 23453 | 124 | with tU noU have "False" by auto} | 
| 25062 | 125 | hence "\<not> t < y" by auto thus "y \<le> t" by (simp add: not_less) | 
| 23453 | 126 | qed | 
| 127 | ||
| 25062 | 128 | lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)" | 
| 23453 | 129 | proof(clarsimp) | 
| 130 | fix x l u y | |
| 25062 | 131 | assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u" | 
| 132 | and px: "t \<le> x" and ly: "l<y" and yu:"y < u" | |
| 23453 | 133 | from tU noU ly yu have tny: "t\<noteq>y" by auto | 
| 25062 | 134 |   {assume H: "y< t"
 | 
| 23453 | 135 | from less_trans[OF ly H] le_less_trans[OF px xu] | 
| 25062 | 136 | have "l < t \<and> t < u" by simp | 
| 23453 | 137 | with tU noU have "False" by auto} | 
| 25062 | 138 | hence "\<not> y<t" by auto thus "t \<le> y" by (simp add: not_less) | 
| 23453 | 139 | qed | 
| 25062 | 140 | lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)" by auto | 
| 141 | lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)" by auto | |
| 142 | lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)" by auto | |
| 23453 | 143 | |
| 144 | lemma lin_dense_conj: | |
| 25062 | 145 | "\<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 | 
| 146 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ; | |
| 147 | \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x | |
| 148 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow> | |
| 149 | \<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) | |
| 150 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))" | |
| 23453 | 151 | by blast | 
| 152 | lemma lin_dense_disj: | |
| 25062 | 153 | "\<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 | 
| 154 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ; | |
| 155 | \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x | |
| 156 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow> | |
| 157 | \<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) | |
| 158 | \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))" | |
| 23453 | 159 | by blast | 
| 160 | ||
| 25062 | 161 | lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk> | 
| 162 | \<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')" | |
| 23453 | 163 | by auto | 
| 164 | ||
| 165 | lemma finite_set_intervals: | |
| 25062 | 166 | assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S" | 
| 167 | and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u" | |
| 168 | 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" | |
| 23453 | 169 | proof- | 
| 25062 | 170 |   let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
 | 
| 171 |   let ?xM = "{y. y\<in> S \<and> x \<le> y}"
 | |
| 23453 | 172 | let ?a = "Max ?Mx" | 
| 173 | let ?b = "Min ?xM" | |
| 174 | have MxS: "?Mx \<subseteq> S" by blast | |
| 175 | hence fMx: "finite ?Mx" using fS finite_subset by auto | |
| 176 | from lx linS have linMx: "l \<in> ?Mx" by blast | |
| 177 |   hence Mxne: "?Mx \<noteq> {}" by blast
 | |
| 178 | have xMS: "?xM \<subseteq> S" by blast | |
| 179 | hence fxM: "finite ?xM" using fS finite_subset by auto | |
| 180 | from xu uinS have linxM: "u \<in> ?xM" by blast | |
| 181 |   hence xMne: "?xM \<noteq> {}" by blast
 | |
| 25062 | 182 | have ax:"?a \<le> x" using Mxne fMx by auto | 
| 183 | have xb:"x \<le> ?b" using xMne fxM by auto | |
| 23453 | 184 | have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast | 
| 185 | have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast | |
| 25062 | 186 | have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S" | 
| 23453 | 187 | proof(clarsimp) | 
| 25062 | 188 | fix y assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S" | 
| 23453 | 189 | from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear) | 
| 25062 | 190 |     moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
 | 
| 191 |     moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
 | |
| 23453 | 192 | ultimately show "False" by blast | 
| 193 | qed | |
| 194 | from ainS binS noy ax xb px show ?thesis by blast | |
| 195 | qed | |
| 196 | ||
| 197 | lemma finite_set_intervals2: | |
| 25062 | 198 | assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S" | 
| 199 | and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u" | |
| 200 | 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)" | |
| 23453 | 201 | proof- | 
| 202 | from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] | |
| 203 | obtain a and b where | |
| 25062 | 204 | as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S" | 
| 205 | and axb: "a \<le> x \<and> x \<le> b \<and> P x" by auto | |
| 206 | from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less) | |
| 23453 | 207 | thus ?thesis using px as bs noS by blast | 
| 208 | qed | |
| 209 | ||
| 210 | end | |
| 211 | ||
| 24344 | 212 | section {* The classical QE after Langford for dense linear orders *}
 | 
| 213 | ||
| 24422 | 214 | context dense_linear_order | 
| 24344 | 215 | begin | 
| 216 | ||
| 217 | lemma dlo_qe_bnds: | |
| 218 |   assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
 | |
| 25062 | 219 | 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)" | 
| 24344 | 220 | proof (simp only: atomize_eq, rule iffI) | 
| 25062 | 221 | assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" | 
| 222 | then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast | |
| 24344 | 223 |   {fix l u assume l: "l \<in> L" and u: "u \<in> U"
 | 
| 25106 | 224 | have "l < x" using xL l by blast | 
| 225 | also have "x < u" using xU u by blast | |
| 226 | finally (less_trans) have "l < u" .} | |
| 25062 | 227 | thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast | 
| 24344 | 228 | next | 
| 25062 | 229 | assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u" | 
| 24344 | 230 | let ?ML = "Max L" | 
| 231 | let ?MU = "Min U" | |
| 25062 | 232 | from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto | 
| 233 | from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto | |
| 234 | from th1 th2 H have "?ML < ?MU" by auto | |
| 235 | with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast | |
| 236 | from th3 th1' have "\<forall>l \<in> L. l < w" by auto | |
| 237 | moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto | |
| 238 | ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto | |
| 24344 | 239 | qed | 
| 240 | ||
| 241 | lemma dlo_qe_noub: | |
| 242 |   assumes ne: "L \<noteq> {}" and fL: "finite L"
 | |
| 25062 | 243 |   shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
 | 
| 24344 | 244 | proof(simp add: atomize_eq) | 
| 25106 | 245 | from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast | 
| 25062 | 246 | from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp | 
| 247 | with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans) | |
| 248 | thus "\<exists>x. \<forall>y\<in>L. y < x" by blast | |
| 24344 | 249 | qed | 
| 250 | ||
| 251 | lemma dlo_qe_nolb: | |
| 252 |   assumes ne: "U \<noteq> {}" and fU: "finite U"
 | |
| 25062 | 253 |   shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
 | 
| 24344 | 254 | proof(simp add: atomize_eq) | 
| 25106 | 255 | from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast | 
| 25062 | 256 | from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp | 
| 257 | with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans) | |
| 258 | thus "\<exists>x. \<forall>y\<in>U. x < y" by blast | |
| 24344 | 259 | qed | 
| 260 | ||
| 261 | lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" | |
| 25106 | 262 | using gt_ex[of t] by auto | 
| 24344 | 263 | |
| 264 | lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq | |
| 265 | le_less neq_iff linear less_not_permute | |
| 266 | ||
| 25106 | 267 | lemma axiom: "dense_linear_order (op \<le>) (op <)" by fact | 
| 24748 | 268 | lemma atoms: | 
| 269 | includes meta_term_syntax | |
| 270 | shows "TERM (less :: 'a \<Rightarrow> _)" | |
| 271 | and "TERM (less_eq :: 'a \<Rightarrow> _)" | |
| 272 | and "TERM (op = :: 'a \<Rightarrow> _)" . | |
| 24344 | 273 | |
| 274 | declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms] | |
| 275 | declare dlo_simps[langfordsimp] | |
| 276 | ||
| 277 | end | |
| 278 | ||
| 279 | (* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *) | |
| 280 | lemma dnf: | |
| 281 | "(P & (Q | R)) = ((P&Q) | (P&R))" | |
| 282 | "((Q | R) & P) = ((Q&P) | (R&P))" | |
| 283 | by blast+ | |
| 284 | ||
| 285 | lemmas weak_dnf_simps = simp_thms dnf | |
| 286 | ||
| 287 | lemma nnf_simps: | |
| 288 | "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" | |
| 289 | "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P" | |
| 290 | by blast+ | |
| 291 | ||
| 292 | lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast | |
| 293 | ||
| 294 | lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib | |
| 295 | ||
| 296 | use "Tools/Qelim/langford.ML" | |
| 297 | method_setup dlo = {*
 | |
| 298 | Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac) | |
| 299 | *} "Langford's algorithm for quantifier elimination in dense linear orders" | |
| 300 | ||
| 301 | ||
| 302 | section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
 | |
| 303 | ||
| 23453 | 304 | text {* Linear order without upper bounds *}
 | 
| 305 | ||
| 26040 | 306 | locale linorder_stupid_syntax = linorder | 
| 23453 | 307 | begin | 
| 26040 | 308 | notation | 
| 309 |   less_eq  ("op \<sqsubseteq>") and
 | |
| 310 |   less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
 | |
| 311 |   less  ("op \<sqsubset>") and
 | |
| 312 |   less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
 | |
| 23453 | 313 | |
| 26040 | 314 | end | 
| 23453 | 315 | |
| 26040 | 316 | locale linorder_no_ub = linorder_stupid_syntax + | 
| 317 | assumes gt_ex: "\<exists>y. less x y" | |
| 318 | begin | |
| 319 | lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto | |
| 320 | ||
| 321 | text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
 | |
| 23453 | 322 | lemma pinf_conj: | 
| 26040 | 323 | assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 324 | and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 325 | shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))" | |
| 23453 | 326 | proof- | 
| 26040 | 327 | from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 328 | and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast | |
| 329 | from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast | |
| 330 | from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all | |
| 331 |   {fix x assume H: "z \<sqsubset> x"
 | |
| 23453 | 332 | from less_trans[OF zz1 H] less_trans[OF zz2 H] | 
| 333 | have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" using z1 zz1 z2 zz2 by auto | |
| 334 | } | |
| 335 | thus ?thesis by blast | |
| 336 | qed | |
| 337 | ||
| 338 | lemma pinf_disj: | |
| 26040 | 339 | assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 340 | and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 341 | shows "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))" | |
| 23453 | 342 | proof- | 
| 26040 | 343 | from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 344 | and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast | |
| 345 | from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast | |
| 346 | from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all | |
| 347 |   {fix x assume H: "z \<sqsubset> x"
 | |
| 23453 | 348 | from less_trans[OF zz1 H] less_trans[OF zz2 H] | 
| 349 | have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" using z1 zz1 z2 zz2 by auto | |
| 350 | } | |
| 351 | thus ?thesis by blast | |
| 352 | qed | |
| 353 | ||
| 26040 | 354 | lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x" | 
| 23453 | 355 | proof- | 
| 26040 | 356 | from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast | 
| 357 | from gt_ex obtain x where x: "z \<sqsubset> x" by blast | |
| 23453 | 358 | from z x p1 show ?thesis by blast | 
| 359 | qed | |
| 360 | ||
| 361 | end | |
| 362 | ||
| 363 | text {* Linear order without upper bounds *}
 | |
| 364 | ||
| 26040 | 365 | locale linorder_no_lb = linorder_stupid_syntax + | 
| 366 | assumes lt_ex: "\<exists>y. less y x" | |
| 23453 | 367 | begin | 
| 26040 | 368 | lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto | 
| 23453 | 369 | |
| 370 | ||
| 26040 | 371 | text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
 | 
| 23453 | 372 | lemma minf_conj: | 
| 26040 | 373 | assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 374 | and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 375 | shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))" | |
| 23453 | 376 | proof- | 
| 26040 | 377 | from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast | 
| 378 | from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast | |
| 379 | from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all | |
| 380 |   {fix x assume H: "x \<sqsubset> z"
 | |
| 23453 | 381 | from less_trans[OF H zz1] less_trans[OF H zz2] | 
| 382 | have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')" using z1 zz1 z2 zz2 by auto | |
| 383 | } | |
| 384 | thus ?thesis by blast | |
| 385 | qed | |
| 386 | ||
| 387 | lemma minf_disj: | |
| 26040 | 388 | assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')" | 
| 389 | and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" | |
| 390 | shows "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))" | |
| 23453 | 391 | proof- | 
| 26040 | 392 | from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast | 
| 393 | from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast | |
| 394 | from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all | |
| 395 |   {fix x assume H: "x \<sqsubset> z"
 | |
| 23453 | 396 | from less_trans[OF H zz1] less_trans[OF H zz2] | 
| 397 | have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')" using z1 zz1 z2 zz2 by auto | |
| 398 | } | |
| 399 | thus ?thesis by blast | |
| 400 | qed | |
| 401 | ||
| 26040 | 402 | lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x" | 
| 23453 | 403 | proof- | 
| 26040 | 404 | from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast | 
| 405 | from lt_ex obtain x where x: "x \<sqsubset> z" by blast | |
| 23453 | 406 | from z x p1 show ?thesis by blast | 
| 407 | qed | |
| 408 | ||
| 409 | end | |
| 410 | ||
| 24344 | 411 | |
| 26040 | 412 | locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub + | 
| 23453 | 413 | fixes between | 
| 26040 | 414 | assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y" | 
| 24398 
8d83b1e7c3dd
Axioms for class no longer use object-universal quantifiers
 chaieb parents: 
24344diff
changeset | 415 | and between_same: "between x x = x" | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 416 | |
| 26040 | 417 | interpretation constr_dense_linear_order < dense_linear_order | 
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 418 | apply unfold_locales | 
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 419 | using gt_ex lt_ex between_less | 
| 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 420 | by (auto, rule_tac x="between x y" in exI, simp) | 
| 23453 | 421 | |
| 26040 | 422 | context constr_dense_linear_order | 
| 423 | begin | |
| 424 | ||
| 23453 | 425 | lemma rinf_U: | 
| 426 | assumes fU: "finite U" | |
| 26040 | 427 | 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 | 
| 428 | \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )" | |
| 429 | 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')" | |
| 23453 | 430 | and nmi: "\<not> MP" and npi: "\<not> PP" and ex: "\<exists> x. P x" | 
| 431 | shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')" | |
| 432 | proof- | |
| 433 | from ex obtain x where px: "P x" by blast | |
| 26040 | 434 | from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto | 
| 435 | then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto | |
| 23453 | 436 |   from uU have Une: "U \<noteq> {}" by auto
 | 
| 26040 | 437 | term "linorder.Min less_eq" | 
| 438 | let ?l = "linorder.Min less_eq U" | |
| 439 | let ?u = "linorder.Max less_eq U" | |
| 23453 | 440 | have linM: "?l \<in> U" using fU Une by simp | 
| 441 | have uinM: "?u \<in> U" using fU Une by simp | |
| 26040 | 442 | have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto | 
| 443 | have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto | |
| 444 | have th:"?l \<sqsubseteq> u" using uU Une lM by auto | |
| 445 | from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" . | |
| 446 | have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp | |
| 447 | from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" . | |
| 23453 | 448 | from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu] | 
| 449 | have "(\<exists> s\<in> U. P s) \<or> | |
| 26040 | 450 | (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" . | 
| 23453 | 451 |   moreover { fix u assume um: "u\<in>U" and pu: "P u"
 | 
| 452 | have "between u u = u" by (simp add: between_same) | |
| 453 | with um pu have "P (between u u)" by simp | |
| 454 | with um have ?thesis by blast} | |
| 455 |   moreover{
 | |
| 26040 | 456 | assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x" | 
| 23453 | 457 | then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U" | 
| 26040 | 458 | and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x" | 
| 23453 | 459 | by blast | 
| 26040 | 460 | from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" . | 
| 23453 | 461 | let ?u = "between t1 t2" | 
| 26040 | 462 | from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto | 
| 25106 | 463 | from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast | 
| 23453 | 464 | with t1M t2M have ?thesis by blast} | 
| 465 | ultimately show ?thesis by blast | |
| 466 | qed | |
| 467 | ||
| 468 | theorem fr_eq: | |
| 469 | assumes fU: "finite U" | |
| 26040 | 470 | 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 | 
| 471 | \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )" | |
| 472 | and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)" | |
| 473 | and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)" | |
| 474 | 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)" | |
| 23453 | 475 | shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))" | 
| 476 | (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D") | |
| 477 | proof- | |
| 478 |  {
 | |
| 479 | assume px: "\<exists> x. P x" | |
| 480 | have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast | |
| 481 |    moreover {assume "MP \<or> PP" hence "?D" by blast}
 | |
| 482 |    moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
 | |
| 483 | from npmibnd[OF nmibnd npibnd] | |
| 26040 | 484 | 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')" . | 
| 23453 | 485 | from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast} | 
| 486 | ultimately have "?D" by blast} | |
| 487 | moreover | |
| 488 |  { assume "?D"
 | |
| 489 |    moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
 | |
| 490 |    moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
 | |
| 491 |    moreover {assume f:"?F" hence "?E" by blast}
 | |
| 492 | ultimately have "?E" by blast} | |
| 493 | ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp | |
| 494 | qed | |
| 495 | ||
| 496 | lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P | |
| 497 | lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P | |
| 498 | ||
| 499 | lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P | |
| 500 | lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P | |
| 501 | lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P | |
| 502 | ||
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 503 | lemma ferrack_axiom: "constr_dense_linear_order less_eq less between" by fact | 
| 24748 | 504 | lemma atoms: | 
| 505 | includes meta_term_syntax | |
| 506 | shows "TERM (less :: 'a \<Rightarrow> _)" | |
| 507 | and "TERM (less_eq :: 'a \<Rightarrow> _)" | |
| 508 | and "TERM (op = :: 'a \<Rightarrow> _)" . | |
| 23453 | 509 | |
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 510 | declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms | 
| 23453 | 511 | nmi: nmi_thms npi: npi_thms lindense: | 
| 512 | lin_dense_thms qe: fr_eq atoms: atoms] | |
| 513 | ||
| 514 | declaration {*
 | |
| 515 | let | |
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 516 | fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
 | 
| 23453 | 517 | fun generic_whatis phi = | 
| 518 | let | |
| 26040 | 519 |   val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
 | 
| 23453 | 520 | fun h x t = | 
| 521 | case term_of t of | |
| 522 |      Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
 | |
| 523 | else Ferrante_Rackoff_Data.Nox | |
| 524 |    | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
 | |
| 525 | else Ferrante_Rackoff_Data.Nox | |
| 526 | | b$y$z => if Term.could_unify (b, lt) then | |
| 527 | if term_of x aconv y then Ferrante_Rackoff_Data.Lt | |
| 528 | else if term_of x aconv z then Ferrante_Rackoff_Data.Gt | |
| 529 | else Ferrante_Rackoff_Data.Nox | |
| 530 | else if Term.could_unify (b, le) then | |
| 531 | if term_of x aconv y then Ferrante_Rackoff_Data.Le | |
| 532 | else if term_of x aconv z then Ferrante_Rackoff_Data.Ge | |
| 533 | else Ferrante_Rackoff_Data.Nox | |
| 534 | else Ferrante_Rackoff_Data.Nox | |
| 535 | | _ => Ferrante_Rackoff_Data.Nox | |
| 536 | in h end | |
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 537 | fun ss phi = HOL_ss addsimps (simps phi) | 
| 23453 | 538 | in | 
| 539 |  Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
 | |
| 540 |   {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
 | |
| 541 | end | |
| 542 | *} | |
| 543 | ||
| 544 | end | |
| 545 | ||
| 23466 
886655a150f6
moved quantifier elimination tools to Tools/Qelim/;
 wenzelm parents: 
23453diff
changeset | 546 | use "Tools/Qelim/ferrante_rackoff.ML" | 
| 23453 | 547 | |
| 23902 
c69069242a51
Added quantifier elimination in dense linear orders after Langford; locale dense_linear_order renamed to constr_dense_linear_order (since it requires the beween constant). locale dense_linear_order is now the classical definition of DLO. constr_dense_linear_order is an instance of dense_linear_order; Method dlo now applies the langford QE, odl Method dlo renamed to ferrack, since it ia a QE only in interpretations where between is interpreted in a manner to vanish after substitution.
 chaieb parents: 
23470diff
changeset | 548 | method_setup ferrack = {*
 | 
| 23453 | 549 | Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac) | 
| 550 | *} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders" | |
| 551 | ||
| 24081 | 552 | end |