src/HOL/Decision_Procs/Dense_Linear_Order.thy
author wenzelm
Tue, 31 Mar 2015 20:18:10 +0200
changeset 59884 bbf49d7dfd6f
parent 59586 ddf6deaadfe8
child 60352 d46de31a50c4
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30439
57c68b3af2ea Updated paths in Decision_Procs comments and NEWS
hoelzl
parents: 29841
diff changeset
     1
(*  Title       : HOL/Decision_Procs/Dense_Linear_Order.thy
29655
ac31940cfb69 Plain, Main form meeting points in import hierarchy
haftmann
parents: 29509
diff changeset
     2
    Author      : Amine Chaieb, TU Muenchen
ac31940cfb69 Plain, Main form meeting points in import hierarchy
haftmann
parents: 29509
diff 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
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 55848
diff changeset
     5
section {* Dense linear order without endpoints
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
     6
  and a quantifier elimination procedure in Ferrante and Rackoff style *}
34cb0b457dcc Old HOL/Dense_Linear_Order.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
0842e906300c normalized imports
haftmann
parents: 30652
diff changeset
     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
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
    12
ML_file "langford_data.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
    13
ML_file "ferrante_rackoff_data.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
    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
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    18
lemma less_not_permute[no_atp]: "\<not> (x < y \<and> y < x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    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
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    21
lemma gather_simps[no_atp]: 
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    22
  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    23
    (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    24
  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    25
    (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    26
  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    27
    (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    28
  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    29
    (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    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
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    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
53072
acc495621d72 prefer plain subscript;
wenzelm
parents: 52090
diff changeset
    35
text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"}*}
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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
53072
acc495621d72 prefer plain subscript;
wenzelm
parents: 52090
diff changeset
    47
text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"}*}
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    48
lemma pinf_gt[no_atp]:  "\<exists>z. \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    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: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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: 35267
diff 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
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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)
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    66
lemma  nmi_P[no_atp]: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    67
lemma  nmi_conj[no_atp]: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 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
    68
  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
    69
  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    70
lemma  nmi_disj[no_atp]: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 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
    71
  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
    72
  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" 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
    73
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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)
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    80
lemma  npi_P[no_atp]: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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>
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
    82
  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
    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>
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
    84
  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" 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
    85
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    86
lemma lin_dense_lt[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    87
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    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)"
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
    89
proof(clarsimp)
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    90
  fix x l u y
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    91
  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < 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
    92
    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < 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
    93
  from tU noU ly yu have tny: "t\<noteq>y" by auto
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    94
  { assume H: "t < 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
    95
    from less_trans[OF lx px] less_trans[OF H yu]
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    96
    have "l < t \<and> t < u" by simp
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    97
    with tU noU have "False" by auto }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    98
  then have "\<not> t < y" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
    99
  then have "y \<le> t" by (simp add: not_less)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   100
  then show "y < t" 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
   101
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
   102
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   103
lemma lin_dense_gt[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   104
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   105
    \<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)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   106
proof(clarsimp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   107
  fix x l u y
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   108
  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"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   109
    and px: "t < x" and ly: "l<y" and yu:"y < u"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   110
  from tU noU ly yu have tny: "t\<noteq>y" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   111
  { assume H: "y< t"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   112
    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   113
    with tU noU have "False" by auto }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   114
  then have "\<not> y<t" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   115
  then have "t \<le> y" by (auto simp add: not_less)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   116
  then show "t < y" using tny by (simp add: less_le)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   117
qed
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   118
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   119
lemma lin_dense_le[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   120
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   121
    \<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)"
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
   122
proof(clarsimp)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   123
  fix x l u 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
   124
  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"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   125
    and px: "x \<le> t" and ly: "l<y" and yu:"y < 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
   126
  from tU noU ly yu have tny: "t\<noteq>y" by auto
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   127
  { assume H: "t < y"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   128
    from less_le_trans[OF lx px] less_trans[OF H yu]
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   129
    have "l < t \<and> t < u" by simp
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   130
    with tU noU have "False" by auto }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   131
  then have "\<not> t < y" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   132
  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
   133
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
   134
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   135
lemma lin_dense_ge[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   136
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   137
    \<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)"
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
   138
proof(clarsimp)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   139
  fix x l u 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
   140
  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"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   141
    and px: "t \<le> x" and ly: "l<y" and yu:"y < 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
   142
  from tU noU ly yu have tny: "t\<noteq>y" by auto
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   143
  { assume H: "y< t"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   144
    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
   145
    have "l < t \<and> t < u" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   146
    with tU noU have "False" by auto }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   147
  then have "\<not> y<t" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   148
  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
   149
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
   150
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   151
lemma lin_dense_eq[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   152
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   153
    \<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)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   154
  by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   155
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   156
lemma lin_dense_neq[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   157
  "t \<in> U \<Longrightarrow>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   158
    \<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)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   159
  by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   160
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   161
lemma lin_dense_P[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   162
  "\<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)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   163
  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
   164
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   165
lemma lin_dense_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
   166
  "\<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
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   167
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 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
   168
  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 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
   169
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   170
  \<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)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   171
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 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
   172
  by blast
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   173
lemma lin_dense_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
   174
  "\<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
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   175
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 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
   176
  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 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
   177
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   178
  \<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)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   179
  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 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
   180
  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
   181
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   182
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>
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
   183
  \<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
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   184
  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
   185
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   186
lemma finite_set_intervals[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
   187
  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   188
    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"
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
   189
  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
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   190
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
   191
  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
   192
  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
   193
  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
   194
  let ?b = "Min ?xM"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   195
  have MxS: "?Mx \<subseteq> S" 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
   196
  hence fMx: "finite ?Mx" using fS finite_subset 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
   197
  from lx linS have linMx: "l \<in> ?Mx" 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
   198
  hence Mxne: "?Mx \<noteq> {}" 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
   199
  have xMS: "?xM \<subseteq> S" 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
   200
  hence fxM: "finite ?xM" using fS finite_subset 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
   201
  from xu uinS have linxM: "u \<in> ?xM" 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
   202
  hence xMne: "?xM \<noteq> {}" 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
   203
  have ax:"?a \<le> x" using Mxne fMx 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
   204
  have xb:"x \<le> ?b" using xMne fxM 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
   205
  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS 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
   206
  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS 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
   207
  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   208
  proof(clarsimp)
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   209
    fix y
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   210
    assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
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
   211
    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   212
    moreover {
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   213
      assume "y \<in> ?Mx"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   214
      hence "y \<le> ?a" using Mxne fMx by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   215
      with ay have "False" by (simp add: not_le[symmetric])
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   216
    }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   217
    moreover {
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   218
      assume "y \<in> ?xM"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   219
      hence "?b \<le> y" using xMne fxM by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   220
      with yb have "False" by (simp add: not_le[symmetric])
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   221
    }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   222
    ultimately show False 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
   223
  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
   224
  from ainS binS noy ax xb px show ?thesis 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
   225
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
   226
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   227
lemma finite_set_intervals2[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
   228
  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   229
    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"
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
   230
  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)"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   231
proof-
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   232
  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   233
  obtain a and b where as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   234
    and axb: "a \<le> x \<and> x \<le> b \<and> P 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
   235
  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  thus ?thesis using px as bs noS 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
   237
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
   238
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   239
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
   240
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   241
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
   242
section {* The classical QE after Langford for dense linear orders *}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   243
53215
5e47c31c6f7c renamed typeclass dense_linorder to unbounded_dense_linorder
hoelzl
parents: 53072
diff changeset
   244
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
   245
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
   246
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   247
lemma interval_empty_iff: "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
27825
12254665fc41 re-arranged class dense_linear_order
haftmann
parents: 27487
diff changeset
   248
  by (auto dest: dense)
12254665fc41 re-arranged class dense_linear_order
haftmann
parents: 27487
diff changeset
   249
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   250
lemma dlo_qe_bnds[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
   251
  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite 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
   252
  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)"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   253
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
   254
  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. 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
   255
  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   256
  { fix l u assume l: "l \<in> L" and u: "u \<in> 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
   257
    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
   258
    also have "x < u" using xU u by blast
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   259
    finally (less_trans) have "l < u" . }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   260
  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
   261
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
   262
  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
   263
  let ?ML = "Max L"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   264
  let ?MU = "Min 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
   265
  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" 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
   266
  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" 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
   267
  from th1 th2 H have "?ML < ?MU" 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
   268
  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" 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
   269
  from th3 th1' have "\<forall>l \<in> L. l < w" 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
   270
  moreover from th4 th2' have "\<forall>u \<in> U. w < u" 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
   271
  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" 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
   272
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
   273
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   274
lemma dlo_qe_noub[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
   275
  assumes ne: "L \<noteq> {}" and fL: "finite L"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   276
  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   277
proof(simp add: atomize_eq)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   278
  from gt_ex[of "Max L"] obtain M where M: "Max L < M" 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
   279
  from ne fL have "\<forall>x \<in> L. x \<le> Max L" 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
   280
  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   281
  thus "\<exists>x. \<forall>y\<in>L. y < x" 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
   282
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
   283
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   284
lemma dlo_qe_nolb[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
   285
  assumes ne: "U \<noteq> {}" and fU: "finite 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
   286
  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   287
proof(simp add: atomize_eq)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   288
  from lt_ex[of "Min U"] obtain M where M: "M < Min U" 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
   289
  from ne fU have "\<forall>x \<in> U. Min U \<le> x" 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
   290
  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   291
  thus "\<exists>x. \<forall>y\<in>U. x < y" 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
   292
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
   293
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   294
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
   295
  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
   296
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   297
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
   298
  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
   299
53215
5e47c31c6f7c renamed typeclass dense_linorder to unbounded_dense_linorder
hoelzl
parents: 53072
diff changeset
   300
lemma axiom[no_atp]: "class.unbounded_dense_linorder (op \<le>) (op <)"
5e47c31c6f7c renamed typeclass dense_linorder to unbounded_dense_linorder
hoelzl
parents: 53072
diff changeset
   301
  by (rule unbounded_dense_linorder_axioms)
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   302
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
   303
  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
   304
    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
   305
    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
   306
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   307
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
   308
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
   309
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   310
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
   311
34cb0b457dcc Old HOL/Dense_Linear_Order.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
(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 38864
diff changeset
   313
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
   314
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   315
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
   316
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   317
lemma nnf_simps[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
   318
    "(\<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)"
34cb0b457dcc Old HOL/Dense_Linear_Order.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
    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  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
   321
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   322
lemma ex_distrib[no_atp]: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" 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
   323
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   324
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
   325
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
   326
ML_file "langford.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
   327
method_setup dlo = {*
55848
1bfe72d14630 more standard module name;
wenzelm
parents: 55792
diff changeset
   328
  Scan.succeed (SIMPLE_METHOD' o Langford.dlo_tac)
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
   329
*} "Langford's algorithm for quantifier elimination in dense linear orders"
34cb0b457dcc Old HOL/Dense_Linear_Order.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
34cb0b457dcc Old HOL/Dense_Linear_Order.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
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30549
diff changeset
   332
section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields *}
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
   333
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   334
text {* Linear order without upper bounds *}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   335
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29252
diff changeset
   336
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
   337
begin
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   338
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
   339
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
   340
  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
   341
  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
   342
  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
   343
  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
   344
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   345
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
   346
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29252
diff changeset
   347
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
   348
  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
   349
begin
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   350
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   351
lemma ge_ex[no_atp]: "\<exists>y. x \<sqsubseteq> y" 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
   352
53072
acc495621d72 prefer plain subscript;
wenzelm
parents: 52090
diff changeset
   353
text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>+\<^sub>\<infinity>)"} *}
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   354
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
   355
  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   356
  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> 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
   357
  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> 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
   358
proof-
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
34cb0b457dcc Old HOL/Dense_Linear_Order.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
     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" 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
   361
  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" 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
   362
  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   363
  { fix x assume H: "z \<sqsubset> 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
   364
    from less_trans[OF zz1 H] less_trans[OF zz2 H]
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   365
    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 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
   366
  }
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   367
  thus ?thesis 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
   368
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
   369
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   370
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
   371
  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   372
    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
   373
  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
   374
proof-
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   376
     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" 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
   377
  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" 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
   378
  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   379
  { fix x assume H: "z \<sqsubset> 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
   380
    from less_trans[OF zz1 H] less_trans[OF zz2 H]
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   381
    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 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
   382
  }
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   383
  thus ?thesis 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
   384
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
   385
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   386
lemma pinf_ex[no_atp]: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   387
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
   388
  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" 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
   389
  from gt_ex obtain x where x: "z \<sqsubset> x" 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
   390
  from z x p1 show ?thesis 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
   391
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
   392
34cb0b457dcc Old HOL/Dense_Linear_Order.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
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
   394
34cb0b457dcc Old HOL/Dense_Linear_Order.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
text {* Linear order without upper bounds *}
34cb0b457dcc Old HOL/Dense_Linear_Order.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
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29252
diff changeset
   397
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
   398
  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
   399
begin
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   400
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   401
lemma le_ex[no_atp]: "\<exists>y. y \<sqsubseteq> x" 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
   402
34cb0b457dcc Old HOL/Dense_Linear_Order.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
53072
acc495621d72 prefer plain subscript;
wenzelm
parents: 52090
diff changeset
   404
text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^sub>-\<^sub>\<infinity>)"} *}
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   405
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
   406
  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   407
    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
   408
  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> 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
   409
proof-
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   410
  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
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" 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
   412
  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   413
  { fix x assume H: "x \<sqsubset> z"
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
   414
    from less_trans[OF H zz1] less_trans[OF H zz2]
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   415
    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 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
   416
  }
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   417
  thus ?thesis 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
   418
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
   419
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   420
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
   421
  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   422
    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
   423
  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   424
proof -
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   425
  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   426
    and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" 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
   427
  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" 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
   428
  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   429
  { fix x assume H: "x \<sqsubset> z"
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
   430
    from less_trans[OF H zz1] less_trans[OF H zz2]
34cb0b457dcc Old HOL/Dense_Linear_Order.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
    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 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
   432
  }
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   433
  thus ?thesis 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
   434
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
   435
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   436
lemma minf_ex[no_atp]:
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   437
  assumes ex: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   438
    and p1: P1
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   439
  shows "\<exists> x. P x"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   440
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
   441
  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" 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
   442
  from lt_ex obtain x where x: "x \<sqsubset> z" 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
   443
  from z x p1 show ?thesis 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
   444
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
   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
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   448
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   449
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
   450
  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
   451
  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   452
    and between_same: "between x x = x"
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 53215
diff changeset
   453
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
   454
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 53215
diff changeset
   455
sublocale dlo: 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
   456
  apply unfold_locales
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  using gt_ex lt_ex between_less
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   458
  apply auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   459
  apply (rule_tac x="between x y" in exI)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   460
  apply simp
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   461
  done
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
   462
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   463
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
   464
  assumes fU: "finite U"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   465
    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
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   466
      \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   467
    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')"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   468
    and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  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
   469
  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   470
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
   471
  from ex obtain x where px: "P x" 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
   472
  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" 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
   473
  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
34cb0b457dcc Old HOL/Dense_Linear_Order.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
  from uU have Une: "U \<noteq> {}" 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
   475
  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
   476
  let ?u = "linorder.Max 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
   477
  have linM: "?l \<in> U" using fU Une 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
   478
  have uinM: "?u \<in> U" using fU Une 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
   479
  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU 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
   480
  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU 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
   481
  have th:"?l \<sqsubseteq> u" using uU Une lM 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
   482
  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> 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
   483
  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu 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
   484
  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
   485
  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   486
  have "(\<exists> s\<in> U. P s) \<or>
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   487
      (\<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)" .
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   488
  moreover {
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   489
    fix u assume um: "u\<in>U" and pu: "P 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
   490
    have "between u u = u" by (simp add: between_same)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   491
    with um pu have "P (between u u)" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   492
    with um have ?thesis by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   493
  moreover {
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
   494
    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"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   495
    then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   496
      and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   497
      and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x" by blast
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   498
    from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   499
    let ?u = "between t1 t2"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   500
    from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   501
    from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   502
    with t1M t2M have ?thesis by blast
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   503
  }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   504
  ultimately show ?thesis by blast
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   505
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
   506
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   507
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
   508
  assumes fU: "finite U"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   509
    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
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   510
     \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   511
    and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   512
    and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   513
    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)"
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
   514
  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u 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
   515
  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   516
proof -
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   517
  { assume px: "\<exists> x. P x"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   518
    have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   519
    moreover { assume "MP \<or> PP" hence "?D" by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   520
    moreover {
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   521
      assume nmi: "\<not> MP" and npi: "\<not> PP"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   522
      from npmibnd[OF nmibnd npibnd]
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   523
      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')" .
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   524
      from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   525
    ultimately have "?D" by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   526
  moreover
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   527
  { assume "?D"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   528
    moreover { assume m:"MP" from minf_ex[OF mi m] have "?E" . }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   529
    moreover { assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   530
    moreover { assume f:"?F" hence "?E" by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   531
    ultimately have "?E" by blast }
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   532
  ultimately have "?E = ?D" by blast thus "?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
   533
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
   534
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   535
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: 35267
diff changeset
   536
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
   537
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   538
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: 35267
diff changeset
   539
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: 35267
diff changeset
   540
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
   541
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   542
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: 34974
diff changeset
   543
  by (rule constr_dense_linorder_axioms)
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   544
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35267
diff changeset
   545
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
   546
  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
   547
    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
   548
    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
   549
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   550
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
   551
    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
   552
    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
   553
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   554
declaration {*
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   555
let
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   556
  fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   557
  fun generic_whatis phi =
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   558
    let
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   559
      val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   560
      fun h x t =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   561
        case Thm.term_of t of
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   562
          Const(@{const_name HOL.eq}, _)$y$z =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   563
            if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   564
            else Ferrante_Rackoff_Data.Nox
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   565
       | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   566
            if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   567
            else Ferrante_Rackoff_Data.Nox
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   568
       | b$y$z => if Term.could_unify (b, lt) then
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   569
                     if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   570
                     else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   571
                     else Ferrante_Rackoff_Data.Nox
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   572
                 else if Term.could_unify (b, le) then
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   573
                     if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   574
                     else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   575
                     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
   576
                 else Ferrante_Rackoff_Data.Nox
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   577
       | _ => Ferrante_Rackoff_Data.Nox
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   578
  in h end
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   579
  fun ss phi =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   580
    simpset_of (put_simpset HOL_ss @{context} 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
   581
in
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   582
  Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   583
    {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
   584
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
   585
*}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   586
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   587
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
   588
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
   589
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
   590
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   591
method_setup ferrack = {*
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   592
  Scan.succeed (SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
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
   593
*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
34cb0b457dcc Old HOL/Dense_Linear_Order.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
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   595
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
   596
subsection {* Ferrante and Rackoff algorithm over ordered fields *}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   597
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   598
lemma neg_prod_lt:"(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   599
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
   600
  assume H: "c < 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   601
  have "c*x < 0 = (0/c < x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   602
    by (simp only: neg_divide_less_eq[OF H] algebra_simps)
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
   603
  also have "\<dots> = (0 < x)" 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
   604
  finally show  "(c*x < 0) == (x > 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
   605
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
   606
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   607
lemma pos_prod_lt:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   608
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
   609
  assume H: "c > 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   610
  then have "c*x < 0 = (0/c > x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   611
    by (simp only: pos_less_divide_eq[OF H] algebra_simps)
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
   612
  also have "\<dots> = (0 > x)" 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
   613
  finally show  "(c*x < 0) == (x < 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
   614
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
   615
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   616
lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   617
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
   618
  assume H: "c < 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   619
  have "c*x + t< 0 = (c*x < -t)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   620
    by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   621
  also have "\<dots> = (-t/c < x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   622
    by (simp only: neg_divide_less_eq[OF H] algebra_simps)
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
   623
  also have "\<dots> = ((- 1/c)*t < x)" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   624
  finally show "(c*x + t < 0) == (x > (- 1/c)*t)" 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
   625
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
   626
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   627
lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   628
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
   629
  assume H: "c > 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   630
  have "c*x + t< 0 = (c*x < -t)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   631
    by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   632
  also have "\<dots> = (-t/c > x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   633
    by (simp only: pos_less_divide_eq[OF H] algebra_simps)
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
   634
  also have "\<dots> = ((- 1/c)*t > x)" 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
   635
  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" 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
   636
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
   637
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   638
lemma sum_lt:"((x::'a::ordered_ab_group_add) + t < 0) == (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
   639
  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
   640
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   641
lemma neg_prod_le:"(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   642
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
   643
  assume H: "c < 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   644
  have "c*x <= 0 = (0/c <= x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   645
    by (simp only: neg_divide_le_eq[OF H] algebra_simps)
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
   646
  also have "\<dots> = (0 <= x)" 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
   647
  finally show  "(c*x <= 0) == (x >= 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
   648
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
   649
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   650
lemma pos_prod_le:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   651
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
   652
  assume H: "c > 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   653
  hence "c*x <= 0 = (0/c >= x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   654
    by (simp only: pos_le_divide_eq[OF H] algebra_simps)
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
   655
  also have "\<dots> = (0 >= x)" 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
   656
  finally show  "(c*x <= 0) == (x <= 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
   657
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
   658
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   659
lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>linordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   660
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
   661
  assume H: "c < 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   662
  have "c*x + t <= 0 = (c*x <= -t)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   663
    by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   664
  also have "\<dots> = (-t/c <= x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   665
    by (simp only: neg_divide_le_eq[OF H] algebra_simps)
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
   666
  also have "\<dots> = ((- 1/c)*t <= x)" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   667
  finally show "(c*x + t <= 0) == (x >= (- 1/c)*t)" 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
   668
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
   669
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   670
lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>linordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   671
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
   672
  assume H: "c > 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   673
  have "c*x + t <= 0 = (c*x <= -t)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   674
    by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   675
  also have "\<dots> = (-t/c >= x)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   676
    by (simp only: pos_le_divide_eq[OF H] algebra_simps)
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
   677
  also have "\<dots> = ((- 1/c)*t >= x)" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   678
  finally show "(c*x + t <= 0) == (x <= (- 1/c)*t)" 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
   679
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
   680
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   681
lemma sum_le:"((x::'a::ordered_ab_group_add) + t <= 0) == (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
   682
  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
   683
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   684
lemma nz_prod_eq:"(c\<Colon>'a\<Colon>linordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   685
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   686
lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>linordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   687
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
   688
  assume H: "c \<noteq> 0"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   689
  have "c*x + t = 0 = (c*x = -t)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   690
    by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   691
  also have "\<dots> = (x = -t/c)"
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   692
    by (simp only: nonzero_eq_divide_eq[OF H] algebra_simps)
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
  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" 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
   694
qed
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   695
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   696
lemma sum_eq:"((x::'a::ordered_ab_group_add) + t = 0) == (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
   697
  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
   698
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   699
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35028
diff changeset
   700
interpretation class_dense_linordered_field: constr_dense_linorder
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29252
diff changeset
   701
 "op <=" "op <"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
   702
   "\<lambda> x y. 1/2 * ((x::'a::{linordered_field}) + y)"
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   703
  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
   704
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   705
declaration{*
34cb0b457dcc Old HOL/Dense_Linear_Order.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
let
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   707
  fun earlier [] x y = false
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   708
    | earlier (h::t) x y =
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   709
        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
   710
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   711
fun dest_frac ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   712
  case Thm.term_of ct of
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   713
    Const (@{const_name Fields.divide},_) $ a $ b=>
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   714
      Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   715
  | Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   716
  | t => Rat.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
   717
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   718
fun mk_frac phi cT x =
49519
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   719
  let val (a, b) = Rat.quotient_of_rat x
4d2c93e1d9c8 misc tuning;
wenzelm
parents: 48891
diff changeset
   720
  in if b = 1 then Numeral.mk_cnumber cT a
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: 45620
diff changeset
   721
    else Thm.apply
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: 45620
diff changeset
   722
         (Thm.apply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
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
   723
                     (Numeral.mk_cnumber cT a))
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   724
         (Numeral.mk_cnumber cT b)
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   725
 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
   726
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   727
fun whatis x ct = case Thm.term_of ct of
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   728
  Const(@{const_name Groups.plus}, _)$(Const(@{const_name Groups.times},_)$_$y)$_ =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   729
     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
   730
     else ("Nox",[])
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   731
| Const(@{const_name Groups.plus}, _)$y$_ =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   732
     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
   733
     else ("Nox",[])
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   734
| Const(@{const_name Groups.times}, _)$_$y =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   735
     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
   736
     else ("Nox",[])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   737
| 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
   738
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   739
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
   740
| xnormalize_conv ctxt (vs as (x::_)) ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   741
   case Thm.term_of ct of
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   742
   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
   743
    (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
   744
    ("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
   745
       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
   746
        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
   747
        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
   748
        val cz = Thm.dest_arg 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
   749
        val neg = cr </ Rat.zero
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   750
        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: 45620
diff changeset
   751
               (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: 45620
diff changeset
   752
                  (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: 45620
diff changeset
   753
                    else Thm.apply (Thm.apply clt cz) c))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   754
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   755
        val th = Thm.implies_elim (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
   756
             (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
   757
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   758
                   (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
   759
      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
   760
    | ("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
   761
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   762
        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
   763
        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   764
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   765
              (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
   766
       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
   767
    | ("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
   768
       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
   769
        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
   770
        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
   771
        val cz = Thm.dest_arg 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
   772
        val neg = cr </ Rat.zero
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   773
        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: 45620
diff changeset
   774
               (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: 45620
diff changeset
   775
                  (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: 45620
diff changeset
   776
                    else Thm.apply (Thm.apply clt cz) c))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   777
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   778
        val th = Thm.implies_elim (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
   779
             (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
   780
        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
   781
      in rth end
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   782
    | _ => 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
   783
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   784
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   785
|  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
   786
   (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
   787
    ("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
   788
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   789
        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
   790
        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
   791
        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
34cb0b457dcc Old HOL/Dense_Linear_Order.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
        val cz = Thm.dest_arg 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
   793
        val neg = cr </ Rat.zero
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   794
        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: 45620
diff changeset
   795
               (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: 45620
diff changeset
   796
                  (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: 45620
diff changeset
   797
                    else Thm.apply (Thm.apply clt cz) c))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   798
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   799
        val th = Thm.implies_elim (instantiate' [SOME T] (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
   800
             (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
   801
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   802
                   (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
   803
      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
   804
    | ("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
   805
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   806
        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
   807
        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_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
   808
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   809
              (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
   810
       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
   811
    | ("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
   812
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   813
        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
   814
        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
   815
        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   816
        val cz = Thm.dest_arg 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
   817
        val neg = cr </ Rat.zero
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   818
        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: 45620
diff changeset
   819
               (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: 45620
diff changeset
   820
                  (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: 45620
diff changeset
   821
                    else Thm.apply (Thm.apply clt cz) c))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   822
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   823
        val th = Thm.implies_elim (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
   824
             (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
   825
        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
   826
      in rth end
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   827
    | _ => 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
   828
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38549
diff changeset
   829
|  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
   830
   (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
   831
    ("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
   832
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   833
        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
   834
        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
   835
        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
   836
        val cz = Thm.dest_arg ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   837
        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: 45620
diff changeset
   838
            (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: 45620
diff changeset
   839
             (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply ceq c) cz)))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   840
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   841
        val th = Thm.implies_elim
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
   842
                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) 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
   843
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   844
                   (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
   845
      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
   846
    | ("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
   847
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   848
        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
   849
        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   850
        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
36753
5cf4e9128f22 renamed Normalizer to the more specific Semiring_Normalizer
haftmann
parents: 36635
diff changeset
   851
              (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
   852
       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
   853
    | ("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
   854
       let
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   855
        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
   856
        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
   857
        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
   858
        val cz = Thm.dest_arg ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   859
        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: 45620
diff changeset
   860
            (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: 45620
diff changeset
   861
             (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply ceq c) cz)))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   862
        val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   863
        val rth = Thm.implies_elim
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
   864
                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) 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
   865
      in rth end
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   866
    | _ => 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
   867
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   868
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
   869
  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
   870
  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
   871
  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
52090
ff1ec795604b proper context;
wenzelm
parents: 51717
diff changeset
   872
  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
   873
in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   874
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: 35084
diff changeset
   875
  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
   876
   let val (ca,cb) = Thm.dest_binop ct
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   877
       val T = Thm.ctyp_of_cterm ca
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
   878
       val th = instantiate' [SOME T] [SOME ca, SOME cb] 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
   879
       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
   880
         (Conv.arg_conv (Conv.arg1_conv
52090
ff1ec795604b proper context;
wenzelm
parents: 51717
diff changeset
   881
              (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   882
       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
   883
   in rth end
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 35084
diff changeset
   884
| 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
   885
   let val (ca,cb) = Thm.dest_binop ct
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   886
       val T = Thm.ctyp_of_cterm ca
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
   887
       val th = instantiate' [SOME T] [SOME ca, SOME cb] 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
   888
       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
   889
         (Conv.arg_conv (Conv.arg1_conv
52090
ff1ec795604b proper context;
wenzelm
parents: 51717
diff changeset
   890
              (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   891
       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
   892
   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
   893
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38549
diff changeset
   894
| 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
   895
   let val (ca,cb) = Thm.dest_binop ct
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   896
       val T = Thm.ctyp_of_cterm ca
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
   897
       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_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
   898
       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
   899
         (Conv.arg_conv (Conv.arg1_conv
52090
ff1ec795604b proper context;
wenzelm
parents: 51717
diff changeset
   900
              (Semiring_Normalizer.semiring_normalize_ord_conv (put_simpset ss ctxt) (earlier vs)))) th
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   901
       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
   902
   in rth end
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38549
diff changeset
   903
| @{term "Not"} $(Const(@{const_name HOL.eq},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36753
diff changeset
   904
| _ => 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
   905
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
   906
34cb0b457dcc Old HOL/Dense_Linear_Order.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
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
   908
 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
   909
  fun h x t =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   910
   case Thm.term_of t of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   911
     Const(@{const_name HOL.eq}, _)$y$z =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   912
      if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   913
      else Ferrante_Rackoff_Data.Nox
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   914
   | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   915
      if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   916
      else Ferrante_Rackoff_Data.Nox
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 35084
diff changeset
   917
   | Const(@{const_name Orderings.less},_)$y$z =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   918
       if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   919
       else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   920
       else Ferrante_Rackoff_Data.Nox
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 35084
diff changeset
   921
   | Const (@{const_name Orderings.less_eq},_)$y$z =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   922
       if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   923
       else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58889
diff changeset
   924
       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
   925
   | _ => 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
   926
 in h 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
   927
fun class_field_ss phi =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   928
  simpset_of (put_simpset HOL_basic_ss @{context}
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   929
    addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49756
diff changeset
   930
    |> 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
   931
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   932
in
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35028
diff changeset
   933
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
   934
  {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
   935
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
   936
*}
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 38864
diff changeset
   937
(*
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   938
lemma upper_bound_finite_set:
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   939
  assumes fS: "finite S"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   940
  shows "\<exists>(a::'a::linorder). \<forall>x \<in> S. f x \<le> a"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   941
proof(induct rule: finite_induct[OF fS])
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   942
  case 1 thus ?case by simp
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   943
next
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   944
  case (2 x F)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   945
  from "2.hyps" obtain a where a:"\<forall>x \<in>F. f x \<le> a" by blast
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   946
  let ?a = "max a (f x)"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   947
  have m: "a \<le> ?a" "f x \<le> ?a" by simp_all
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   948
  {fix y assume y: "y \<in> insert x F"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   949
    {assume "y = x" hence "f y \<le> ?a" using m by simp}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   950
    moreover
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   951
    {assume yF: "y\<in> F" from a[rule_format, OF yF] m have "f y \<le> ?a" by (simp add: max_def)}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   952
    ultimately have "f y \<le> ?a" using y by blast}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   953
  then show ?case by blast
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   954
qed
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   955
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   956
lemma lower_bound_finite_set:
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   957
  assumes fS: "finite S"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   958
  shows "\<exists>(a::'a::linorder). \<forall>x \<in> S. f x \<ge> a"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   959
proof(induct rule: finite_induct[OF fS])
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   960
  case 1 thus ?case by simp
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   961
next
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   962
  case (2 x F)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   963
  from "2.hyps" obtain a where a:"\<forall>x \<in>F. f x \<ge> a" by blast
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   964
  let ?a = "min a (f x)"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   965
  have m: "a \<ge> ?a" "f x \<ge> ?a" by simp_all
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   966
  {fix y assume y: "y \<in> insert x F"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   967
    {assume "y = x" hence "f y \<ge> ?a" using m by simp}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   968
    moreover
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   969
    {assume yF: "y\<in> F" from a[rule_format, OF yF] m have "f y \<ge> ?a" by (simp add: min_def)}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   970
    ultimately have "f y \<ge> ?a" using y by blast}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   971
  then show ?case by blast
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   972
qed
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   973
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   974
lemma bound_finite_set: assumes f: "finite S"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   975
  shows "\<exists>a. \<forall>x \<in>S. (f x:: 'a::linorder) \<le> a"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   976
proof-
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   977
  let ?F = "f ` S"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   978
  from f have fF: "finite ?F" by simp
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   979
  let ?a = "Max ?F"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   980
  {assume "S = {}" hence ?thesis by blast}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   981
  moreover
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   982
  {assume Se: "S \<noteq> {}" hence Fe: "?F \<noteq> {}" by simp
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   983
  {fix x assume x: "x \<in> S"
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   984
    hence th0: "f x \<in> ?F" by simp
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   985
    hence "f x \<le> ?a" using Max_ge[OF fF th0] ..}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   986
  hence ?thesis by blast}
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   987
ultimately show ?thesis by blast
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29823
diff changeset
   988
qed
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 38864
diff changeset
   989
*)
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
   990
34cb0b457dcc Old HOL/Dense_Linear_Order.thy and the setup in Arith_Tools for Ferrante and Rackoff's Quantifier elimination for linear arithmetic over ordered Fields.
chaieb
parents:
diff changeset
   991
end