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