src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
author wenzelm
Sun, 03 Dec 2017 22:28:19 +0100
changeset 67123 3fe40ff1b921
parent 66809 f6a30d48aab0
child 67399 eab6ce8368fa
permissions -rw-r--r--
misc tuning and modernization;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     1
(*  Title:      HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     2
    Author:     Amine Chaieb
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     3
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     4
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
     5
section \<open>A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     6
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
     7
theory Parametric_Ferrante_Rackoff
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
     8
imports
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
     9
  Reflected_Multivariate_Polynomial
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    10
  Dense_Linear_Order
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    11
  DP_Library
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64240
diff changeset
    12
  "HOL-Library.Code_Target_Numeral"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    13
begin
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    14
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
    15
subsection \<open>Terms\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    16
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    17
datatype (plugins del: size) tm = CP poly | Bound nat | Add tm tm | Mul poly tm
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    18
  | Neg tm | Sub tm tm | CNP nat poly tm
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    19
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    20
instantiation tm :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    21
begin
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    22
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    23
primrec size_tm :: "tm \<Rightarrow> nat"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    24
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    25
    "size_tm (CP c) = polysize c"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    26
  | "size_tm (Bound n) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    27
  | "size_tm (Neg a) = 1 + size_tm a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    28
  | "size_tm (Add a b) = 1 + size_tm a + size_tm b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    29
  | "size_tm (Sub a b) = 3 + size_tm a + size_tm b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    30
  | "size_tm (Mul c a) = 1 + polysize c + size_tm a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    31
  | "size_tm (CNP n c a) = 3 + polysize c + size_tm a "
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    32
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    33
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    34
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
    35
end
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    36
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    37
text \<open>Semantics of terms tm.\<close>
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    38
primrec Itm :: "'a::{field_char_0,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    39
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    40
    "Itm vs bs (CP c) = (Ipoly vs c)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    41
  | "Itm vs bs (Bound n) = bs!n"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    42
  | "Itm vs bs (Neg a) = -(Itm vs bs a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    43
  | "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    44
  | "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    45
  | "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    46
  | "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    47
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    48
fun allpolys :: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    49
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    50
    "allpolys P (CP c) = P c"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    51
  | "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    52
  | "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    53
  | "allpolys P (Neg p) = allpolys P p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    54
  | "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    55
  | "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    56
  | "allpolys P p = True"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    57
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    58
primrec tmboundslt :: "nat \<Rightarrow> tm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    59
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    60
    "tmboundslt n (CP c) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    61
  | "tmboundslt n (Bound m) = (m < n)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    62
  | "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    63
  | "tmboundslt n (Neg a) = tmboundslt n a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    64
  | "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    65
  | "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    66
  | "tmboundslt n (Mul i a) = tmboundslt n a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    67
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    68
primrec tmbound0 :: "tm \<Rightarrow> bool"  \<comment> \<open>a \<open>tm\<close> is \<^emph>\<open>independent\<close> of Bound 0\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    69
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    70
    "tmbound0 (CP c) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    71
  | "tmbound0 (Bound n) = (n>0)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    72
  | "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    73
  | "tmbound0 (Neg a) = tmbound0 a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    74
  | "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    75
  | "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    76
  | "tmbound0 (Mul i a) = tmbound0 a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    77
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    78
lemma tmbound0_I:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    79
  assumes "tmbound0 a"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    80
  shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    81
  using assms by (induct a rule: tm.induct) auto
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    82
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    83
primrec tmbound :: "nat \<Rightarrow> tm \<Rightarrow> bool"  \<comment> \<open>a \<open>tm\<close> is \<^emph>\<open>independent\<close> of Bound n\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    84
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    85
    "tmbound n (CP c) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    86
  | "tmbound n (Bound m) = (n \<noteq> m)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    87
  | "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    88
  | "tmbound n (Neg a) = tmbound n a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    89
  | "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    90
  | "tmbound n (Sub a b) = (tmbound n a \<and> tmbound n b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
    91
  | "tmbound n (Mul i a) = tmbound n a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    92
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    93
lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    94
  by (induct t) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    95
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    96
lemma tmbound_I:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    97
  assumes bnd: "tmboundslt (length bs) t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    98
    and nb: "tmbound n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    99
    and le: "n \<le> length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   100
  shows "Itm vs (bs[n:=x]) t = Itm vs bs t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   101
  using nb le bnd
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   102
  by (induct t rule: tm.induct) auto
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   103
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   104
fun decrtm0 :: "tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   105
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   106
    "decrtm0 (Bound n) = Bound (n - 1)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   107
  | "decrtm0 (Neg a) = Neg (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   108
  | "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   109
  | "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   110
  | "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   111
  | "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   112
  | "decrtm0 a = a"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   113
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   114
fun incrtm0 :: "tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   115
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   116
    "incrtm0 (Bound n) = Bound (n + 1)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   117
  | "incrtm0 (Neg a) = Neg (incrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   118
  | "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   119
  | "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   120
  | "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   121
  | "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   122
  | "incrtm0 a = a"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   123
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   124
lemma decrtm0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   125
  assumes nb: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   126
  shows "Itm vs (x # bs) t = Itm vs bs (decrtm0 t)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   127
  using nb by (induct t rule: decrtm0.induct) simp_all
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   128
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   129
lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   130
  by (induct t rule: decrtm0.induct) simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   131
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   132
primrec decrtm :: "nat \<Rightarrow> tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   133
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   134
    "decrtm m (CP c) = (CP c)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   135
  | "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   136
  | "decrtm m (Neg a) = Neg (decrtm m a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   137
  | "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   138
  | "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   139
  | "decrtm m (Mul c a) = Mul c (decrtm m a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   140
  | "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   141
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   142
primrec removen :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   143
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   144
    "removen n [] = []"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   145
  | "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   146
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   147
lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   148
  by (induct xs arbitrary: n) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   149
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   150
lemma nth_length_exceeds: "n \<ge> length xs \<Longrightarrow> xs!n = []!(n - length xs)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   151
  by (induct xs arbitrary: n) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   152
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   153
lemma removen_length: "length (removen n xs) = (if n \<ge> length xs then length xs else length xs - 1)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   154
  by (induct xs arbitrary: n) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   155
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   156
lemma removen_nth:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   157
  "(removen n xs)!m =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   158
    (if n \<ge> length xs then xs!m
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   159
     else if m < n then xs!m
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   160
     else if m \<le> length xs then xs!(Suc m)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   161
     else []!(m - (length xs - 1)))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   162
proof (induct xs arbitrary: n m)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   163
  case Nil
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   164
  then show ?case by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   165
next
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   166
  case (Cons x xs)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   167
  let ?l = "length (x # xs)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   168
  consider "n \<ge> ?l" | "n < ?l" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   169
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   170
  proof cases
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   171
    case 1
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   172
    with removen_same[OF this] show ?thesis by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   173
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   174
    case nl: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   175
    consider "m < n" | "m \<ge> n" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   176
    then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   177
    proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   178
      case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   179
      then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   180
        using Cons by (cases m) auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   181
    next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   182
      case 2
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   183
      consider "m \<le> ?l" | "m > ?l" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   184
      then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   185
      proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   186
        case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   187
        then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   188
          using Cons by (cases m) auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   189
      next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   190
        case ml: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   191
        have th: "length (removen n (x # xs)) = length xs"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   192
          using removen_length[where n = n and xs= "x # xs"] nl by simp
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   193
        with ml have "m \<ge> length (removen n (x # xs))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   194
          by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   195
        from th nth_length_exceeds[OF this] have "(removen n (x # xs))!m = [] ! (m - length xs)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   196
           by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   197
        then have "(removen n (x # xs))!m = [] ! (m - (length (x # xs) - 1))"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
   198
          by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   199
        then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   200
          using ml nl by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   201
      qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   202
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   203
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   204
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   205
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   206
lemma decrtm:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   207
  assumes bnd: "tmboundslt (length bs) t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   208
    and nb: "tmbound m t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   209
    and nle: "m \<le> length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   210
  shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   211
  using bnd nb nle by (induct t rule: tm.induct) (auto simp add: removen_nth)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   212
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   213
primrec tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   214
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   215
    "tmsubst0 t (CP c) = CP c"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   216
  | "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   217
  | "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   218
  | "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   219
  | "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   220
  | "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   221
  | "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   222
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   223
lemma tmsubst0: "Itm vs (x # bs) (tmsubst0 t a) = Itm vs (Itm vs (x # bs) t # bs) a"
41842
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41838
diff changeset
   224
  by (induct a rule: tm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   225
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   226
lemma tmsubst0_nb: "tmbound0 t \<Longrightarrow> tmbound0 (tmsubst0 t a)"
41842
d8f76db6a207 added simp lemma nth_Cons_pos to List
nipkow
parents: 41838
diff changeset
   227
  by (induct a rule: tm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   228
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   229
primrec tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   230
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   231
    "tmsubst n t (CP c) = CP c"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   232
  | "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   233
  | "tmsubst n t (CNP m c a) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   234
      (if n = m then Add (Mul c t) (tmsubst n t a) else CNP m c (tmsubst n t a))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   235
  | "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   236
  | "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   237
  | "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   238
  | "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   239
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   240
lemma tmsubst:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   241
  assumes nb: "tmboundslt (length bs) a"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   242
    and nlt: "n \<le> length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   243
  shows "Itm vs bs (tmsubst n t a) = Itm vs (bs[n:= Itm vs bs t]) a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   244
  using nb nlt
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   245
  by (induct a rule: tm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   246
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   247
lemma tmsubst_nb0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   248
  assumes tnb: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   249
  shows "tmbound0 (tmsubst 0 t a)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   250
  using tnb
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   251
  by (induct a rule: tm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   252
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   253
lemma tmsubst_nb:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   254
  assumes tnb: "tmbound m t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   255
  shows "tmbound m (tmsubst m t a)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   256
  using tnb
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   257
  by (induct a rule: tm.induct) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   258
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   259
lemma incrtm0_tmbound: "tmbound n t \<Longrightarrow> tmbound (Suc n) (incrtm0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   260
  by (induct t) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   261
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   262
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   263
text \<open>Simplification.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   264
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   265
fun tmadd:: "tm \<Rightarrow> tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   266
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   267
    "tmadd (CNP n1 c1 r1) (CNP n2 c2 r2) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   268
      (if n1 = n2 then
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   269
        let c = c1 +\<^sub>p c2
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   270
        in if c = 0\<^sub>p then tmadd r1 r2 else CNP n1 c (tmadd r1 r2)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   271
      else if n1 \<le> n2 then (CNP n1 c1 (tmadd r1 (CNP n2 c2 r2)))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   272
      else (CNP n2 c2 (tmadd (CNP n1 c1 r1) r2)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   273
  | "tmadd (CNP n1 c1 r1) t = CNP n1 c1 (tmadd r1 t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   274
  | "tmadd t (CNP n2 c2 r2) = CNP n2 c2 (tmadd t r2)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   275
  | "tmadd (CP b1) (CP b2) = CP (b1 +\<^sub>p b2)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   276
  | "tmadd a b = Add a b"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   277
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   278
lemma tmadd [simp]: "Itm vs bs (tmadd t s) = Itm vs bs (Add t s)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   279
  apply (induct t s rule: tmadd.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   280
                      apply (simp_all add: Let_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   281
  apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   282
   apply (case_tac "n1 \<le> n2")
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   283
    apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   284
   apply (case_tac "n1 = n2")
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   285
    apply (simp_all add: algebra_simps)
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   286
  apply (simp only: distrib_left [symmetric] polyadd [symmetric])
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   287
  apply simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   288
  done
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   289
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   290
lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd t s)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   291
  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   292
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   293
lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd t s)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   294
  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   295
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   296
lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd t s)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   297
  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   298
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   299
lemma tmadd_allpolys_npoly[simp]:
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   300
  "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd t s)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   301
  by (induct t s rule: tmadd.induct) (simp_all add: Let_def polyadd_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   302
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   303
fun tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   304
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   305
    "tmmul (CP j) = (\<lambda>i. CP (i *\<^sub>p j))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   306
  | "tmmul (CNP n c a) = (\<lambda>i. CNP n (i *\<^sub>p c) (tmmul a i))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   307
  | "tmmul t = (\<lambda>i. Mul i t)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   308
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   309
lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   310
  by (induct t arbitrary: i rule: tmmul.induct) (simp_all add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   311
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   312
lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   313
  by (induct t arbitrary: i rule: tmmul.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   314
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   315
lemma tmmul_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   316
  by (induct t arbitrary: n rule: tmmul.induct) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   317
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   318
lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   319
  by (induct t arbitrary: i rule: tmmul.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   320
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   321
lemma tmmul_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   322
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   323
  shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   324
  by (induct t rule: tmmul.induct) (simp_all add: Let_def polymul_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   325
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   326
definition tmneg :: "tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   327
  where "tmneg t \<equiv> tmmul t (C (- 1,1))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   328
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   329
definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   330
  where "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd s (tmneg t))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   331
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   332
lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   333
  using tmneg_def[of t] by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   334
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   335
lemma tmneg_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   336
  using tmneg_def by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   337
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   338
lemma tmneg_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   339
  using tmneg_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   340
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   341
lemma tmneg_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   342
  using tmneg_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   343
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   344
lemma [simp]: "isnpoly (C (-1, 1))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   345
  by (simp add: isnpoly_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   346
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   347
lemma tmneg_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   348
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   349
  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   350
  by (auto simp: tmneg_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   351
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   352
lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   353
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   354
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   355
lemma tmsub_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmsub t s)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   356
  using tmsub_def by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   357
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   358
lemma tmsub_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmsub t s)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   359
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   360
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   361
lemma tmsub_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmsub t s)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   362
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   363
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   364
lemma tmsub_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   365
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   366
  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   367
  by (simp add: tmsub_def isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   368
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   369
fun simptm :: "tm \<Rightarrow> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   370
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   371
    "simptm (CP j) = CP (polynate j)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   372
  | "simptm (Bound n) = CNP n (1)\<^sub>p (CP 0\<^sub>p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   373
  | "simptm (Neg t) = tmneg (simptm t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   374
  | "simptm (Add t s) = tmadd (simptm t) (simptm s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   375
  | "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   376
  | "simptm (Mul i t) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   377
      (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   378
  | "simptm (CNP n c t) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   379
      (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p)) (simptm t))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   380
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   381
lemma polynate_stupid:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   382
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
   383
  shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   384
  apply (subst polynate[symmetric])
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   385
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   386
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   387
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   388
lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   389
  by (induct t rule: simptm.induct) (auto simp add: Let_def polynate_stupid)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   390
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   391
lemma simptm_tmbound0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   392
  by (induct t rule: simptm.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   393
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   394
lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   395
  by (induct t rule: simptm.induct) (auto simp add: Let_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   396
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   397
lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   398
  by (induct t rule: simptm.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   399
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   400
lemma [simp]: "isnpoly 0\<^sub>p"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   401
  and [simp]: "isnpoly (C (1, 1))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   402
  by (simp_all add: isnpoly_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   403
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   404
lemma simptm_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   405
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   406
  shows "allpolys isnpoly (simptm p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   407
  by (induct p rule: simptm.induct) (auto simp add: Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   408
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   409
declare let_cong[fundef_cong del]
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   410
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   411
fun split0 :: "tm \<Rightarrow> poly \<times> tm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   412
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   413
    "split0 (Bound 0) = ((1)\<^sub>p, CP 0\<^sub>p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   414
  | "split0 (CNP 0 c t) = (let (c', t') = split0 t in (c +\<^sub>p c', t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   415
  | "split0 (Neg t) = (let (c, t') = split0 t in (~\<^sub>p c, Neg t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   416
  | "split0 (CNP n c t) = (let (c', t') = split0 t in (c', CNP n c t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   417
  | "split0 (Add s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   418
  | "split0 (Sub s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   419
  | "split0 (Mul c t) = (let (c', t') = split0 t in (c *\<^sub>p c', Mul c t'))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   420
  | "split0 t = (0\<^sub>p, t)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   421
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   422
declare let_cong[fundef_cong]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   423
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   424
lemma split0_stupid[simp]: "\<exists>x y. (x, y) = split0 p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   425
  apply (rule exI[where x="fst (split0 p)"])
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   426
  apply (rule exI[where x="snd (split0 p)"])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   427
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   428
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   429
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   430
lemma split0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   431
  "tmbound 0 (snd (split0 t)) \<and> Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   432
  apply (induct t rule: split0.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   433
          apply simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   434
         apply (simp add: Let_def split_def field_simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   435
        apply (simp add: Let_def split_def field_simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   436
       apply (simp add: Let_def split_def field_simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   437
      apply (simp add: Let_def split_def field_simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   438
     apply (simp add: Let_def split_def field_simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   439
    apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric])
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   440
   apply (simp add: Let_def split_def field_simps)
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
   441
  apply (simp add: Let_def split_def field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   442
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   443
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   444
lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   445
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   446
  fix c' t'
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   447
  assume "split0 t = (c', t')"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   448
  then have "c' = fst (split0 t)" "t' = snd (split0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   449
    by auto
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   450
  with split0[where t="t" and bs="bs"] show "Itm vs bs t = Itm vs bs (CNP 0 c' t')"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   451
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   452
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   453
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   454
lemma split0_nb0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   455
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   456
  shows "split0 t = (c',t') \<Longrightarrow>  tmbound 0 t'"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   457
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   458
  fix c' t'
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   459
  assume "split0 t = (c', t')"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   460
  then have "c' = fst (split0 t)" "t' = snd (split0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   461
    by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   462
  with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   463
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   464
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   465
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   466
lemma split0_nb0'[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   467
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   468
  shows "tmbound0 (snd (split0 t))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   469
  using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   470
  by (simp add: tmbound0_tmbound_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   471
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   472
lemma split0_nb:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   473
  assumes nb: "tmbound n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   474
  shows "tmbound n (snd (split0 t))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   475
  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   476
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   477
lemma split0_blt:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   478
  assumes nb: "tmboundslt n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   479
  shows "tmboundslt n (snd (split0 t))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   480
  using nb by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   481
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   482
lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   483
  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   484
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   485
lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0 \<or> n > 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   486
  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   487
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   488
lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   489
  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   490
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   491
lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   492
  by (induct p rule: split0.induct) (auto simp  add: isnpoly_def Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   493
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   494
lemma isnpoly_fst_split0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   495
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   496
  shows "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   497
  by (induct p rule: split0.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   498
    (auto simp  add: polyadd_norm polysub_norm polyneg_norm polymul_norm Let_def split_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   499
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   500
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   501
subsection \<open>Formulae\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   502
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   503
datatype (plugins del: size) fm = T | F | Le tm | Lt tm | Eq tm | NEq tm |
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   504
  NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   505
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   506
instantiation fm :: size
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   507
begin
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   508
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   509
primrec size_fm :: "fm \<Rightarrow> nat"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   510
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   511
    "size_fm (NOT p) = 1 + size_fm p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   512
  | "size_fm (And p q) = 1 + size_fm p + size_fm q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   513
  | "size_fm (Or p q) = 1 + size_fm p + size_fm q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   514
  | "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   515
  | "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   516
  | "size_fm (E p) = 1 + size_fm p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   517
  | "size_fm (A p) = 4 + size_fm p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   518
  | "size_fm T = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   519
  | "size_fm F = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   520
  | "size_fm (Le _) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   521
  | "size_fm (Lt _) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   522
  | "size_fm (Eq _) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   523
  | "size_fm (NEq _) = 1"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   524
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   525
instance ..
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   526
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   527
end
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   528
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   529
lemma fmsize_pos [simp]: "size p > 0" for p :: fm
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
   530
  by (induct p) simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   531
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   532
text \<open>Semantics of formulae (fm).\<close>
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   533
primrec Ifm ::"'a::linordered_field list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   534
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   535
    "Ifm vs bs T = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   536
  | "Ifm vs bs F = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   537
  | "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   538
  | "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   539
  | "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   540
  | "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   541
  | "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   542
  | "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   543
  | "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   544
  | "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   545
  | "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   546
  | "Ifm vs bs (E p) = (\<exists>x. Ifm vs (x#bs) p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   547
  | "Ifm vs bs (A p) = (\<forall>x. Ifm vs (x#bs) p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   548
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   549
fun not:: "fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   550
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   551
    "not (NOT (NOT p)) = not p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   552
  | "not (NOT p) = p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   553
  | "not T = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   554
  | "not F = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   555
  | "not (Lt t) = Le (tmneg t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   556
  | "not (Le t) = Lt (tmneg t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   557
  | "not (Eq t) = NEq t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   558
  | "not (NEq t) = Eq t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   559
  | "not p = NOT p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   560
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   561
lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   562
  by (induct p rule: not.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   563
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   564
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   565
  where "conj p q \<equiv>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   566
    (if p = F \<or> q = F then F
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   567
     else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   568
     else if q = T then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   569
     else if p = q then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   570
     else And p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   571
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   572
lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   573
  by (cases "p=F \<or> q=F", simp_all add: conj_def) (cases p, simp_all)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   574
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   575
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   576
  where "disj p q \<equiv>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   577
    (if (p = T \<or> q = T) then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   578
     else if p = F then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   579
     else if q = F then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   580
     else if p = q then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   581
     else Or p q)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   582
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   583
lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   584
  by (cases "p = T \<or> q = T", simp_all add: disj_def) (cases p, simp_all)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   585
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   586
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   587
  where "imp p q \<equiv>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   588
    (if p = F \<or> q = T \<or> p = q then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   589
     else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   590
     else if q = F then not p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   591
     else Imp p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   592
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   593
lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   594
  by (cases "p = F \<or> q = T") (simp_all add: imp_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   595
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   596
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   597
  where "iff p q \<equiv>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   598
   (if p = q then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   599
    else if p = NOT q \<or> NOT p = q then F
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   600
    else if p = F then not q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   601
    else if q = F then not p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   602
    else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   603
    else if q = T then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   604
    else Iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   605
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   606
lemma iff[simp]: "Ifm vs bs (iff p q) = Ifm vs bs (Iff p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   607
  by (unfold iff_def, cases "p = q", simp, cases "p = NOT q", simp) (cases "NOT p= q", auto)
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   608
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   609
text \<open>Quantifier freeness.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   610
fun qfree:: "fm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   611
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   612
    "qfree (E p) = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   613
  | "qfree (A p) = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   614
  | "qfree (NOT p) = qfree p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   615
  | "qfree (And p q) = (qfree p \<and> qfree q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   616
  | "qfree (Or  p q) = (qfree p \<and> qfree q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   617
  | "qfree (Imp p q) = (qfree p \<and> qfree q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   618
  | "qfree (Iff p q) = (qfree p \<and> qfree q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   619
  | "qfree p = True"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   620
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   621
text \<open>Boundedness and substitution.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   622
primrec boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   623
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   624
    "boundslt n T = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   625
  | "boundslt n F = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   626
  | "boundslt n (Lt t) = tmboundslt n t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   627
  | "boundslt n (Le t) = tmboundslt n t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   628
  | "boundslt n (Eq t) = tmboundslt n t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   629
  | "boundslt n (NEq t) = tmboundslt n t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   630
  | "boundslt n (NOT p) = boundslt n p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   631
  | "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   632
  | "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   633
  | "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   634
  | "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   635
  | "boundslt n (E p) = boundslt (Suc n) p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   636
  | "boundslt n (A p) = boundslt (Suc n) p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   637
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   638
fun bound0:: "fm \<Rightarrow> bool"  \<comment> \<open>a formula is independent of Bound 0\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   639
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   640
    "bound0 T = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   641
  | "bound0 F = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   642
  | "bound0 (Lt a) = tmbound0 a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   643
  | "bound0 (Le a) = tmbound0 a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   644
  | "bound0 (Eq a) = tmbound0 a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   645
  | "bound0 (NEq a) = tmbound0 a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   646
  | "bound0 (NOT p) = bound0 p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   647
  | "bound0 (And p q) = (bound0 p \<and> bound0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   648
  | "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   649
  | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   650
  | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   651
  | "bound0 p = False"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   652
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   653
lemma bound0_I:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   654
  assumes bp: "bound0 p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   655
  shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   656
  using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   657
  by (induct p rule: bound0.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   658
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   659
primrec bound:: "nat \<Rightarrow> fm \<Rightarrow> bool"  \<comment> \<open>a formula is independent of Bound n\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   660
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   661
    "bound m T = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   662
  | "bound m F = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   663
  | "bound m (Lt t) = tmbound m t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   664
  | "bound m (Le t) = tmbound m t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   665
  | "bound m (Eq t) = tmbound m t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   666
  | "bound m (NEq t) = tmbound m t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   667
  | "bound m (NOT p) = bound m p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   668
  | "bound m (And p q) = (bound m p \<and> bound m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   669
  | "bound m (Or p q) = (bound m p \<and> bound m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   670
  | "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   671
  | "bound m (Iff p q) = (bound m p \<and> bound m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   672
  | "bound m (E p) = bound (Suc m) p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   673
  | "bound m (A p) = bound (Suc m) p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   674
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   675
lemma bound_I:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   676
  assumes bnd: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   677
    and nb: "bound n p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   678
    and le: "n \<le> length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   679
  shows "Ifm vs (bs[n:=x]) p = Ifm vs bs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   680
  using bnd nb le tmbound_I[where bs=bs and vs = vs]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   681
proof (induct p arbitrary: bs n rule: fm.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   682
  case (E p bs n)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   683
  have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" for y
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   684
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   685
    from E have bnd: "boundslt (length (y#bs)) p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   686
      and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   687
    from E.hyps[OF bnd nb le tmbound_I] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   688
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   689
  then show ?case by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   690
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   691
  case (A p bs n)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   692
  have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" for y
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   693
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   694
    from A have bnd: "boundslt (length (y#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   695
      and nb: "bound (Suc n) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   696
      and le: "Suc n \<le> length (y#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   697
      by simp_all
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   698
    from A.hyps[OF bnd nb le tmbound_I] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   699
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   700
  then show ?case by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   701
qed auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   702
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   703
fun decr0 :: "fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   704
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   705
    "decr0 (Lt a) = Lt (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   706
  | "decr0 (Le a) = Le (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   707
  | "decr0 (Eq a) = Eq (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   708
  | "decr0 (NEq a) = NEq (decrtm0 a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   709
  | "decr0 (NOT p) = NOT (decr0 p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   710
  | "decr0 (And p q) = conj (decr0 p) (decr0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   711
  | "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   712
  | "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   713
  | "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   714
  | "decr0 p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   715
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   716
lemma decr0:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   717
  assumes "bound0 p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   718
  shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   719
  using assms by (induct p rule: decr0.induct) (simp_all add: decrtm0)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   720
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   721
primrec decr :: "nat \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   722
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   723
    "decr m T = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   724
  | "decr m F = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   725
  | "decr m (Lt t) = (Lt (decrtm m t))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   726
  | "decr m (Le t) = (Le (decrtm m t))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   727
  | "decr m (Eq t) = (Eq (decrtm m t))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   728
  | "decr m (NEq t) = (NEq (decrtm m t))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   729
  | "decr m (NOT p) = NOT (decr m p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   730
  | "decr m (And p q) = conj (decr m p) (decr m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   731
  | "decr m (Or p q) = disj (decr m p) (decr m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   732
  | "decr m (Imp p q) = imp (decr m p) (decr m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   733
  | "decr m (Iff p q) = iff (decr m p) (decr m q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   734
  | "decr m (E p) = E (decr (Suc m) p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   735
  | "decr m (A p) = A (decr (Suc m) p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   736
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   737
lemma decr:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   738
  assumes bnd: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   739
    and nb: "bound m p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   740
    and nle: "m < length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   741
  shows "Ifm vs (removen m bs) (decr m p) = Ifm vs bs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   742
  using bnd nb nle
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   743
proof (induct p arbitrary: bs m rule: fm.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   744
  case (E p bs m)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   745
  have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   746
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   747
    from E
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   748
    have bnd: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   749
      and nb: "bound (Suc m) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   750
      and nle: "Suc m < length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   751
      by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   752
    from E(1)[OF bnd nb nle] show ?thesis .
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   753
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   754
  then show ?case by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   755
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   756
  case (A p bs m)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   757
  have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   758
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   759
    from A
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   760
    have bnd: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   761
      and nb: "bound (Suc m) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   762
      and nle: "Suc m < length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   763
      by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   764
    from A(1)[OF bnd nb nle] show ?thesis .
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   765
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   766
  then show ?case by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   767
qed (auto simp add: decrtm removen_nth)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   768
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   769
primrec subst0 :: "tm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   770
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   771
    "subst0 t T = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   772
  | "subst0 t F = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   773
  | "subst0 t (Lt a) = Lt (tmsubst0 t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   774
  | "subst0 t (Le a) = Le (tmsubst0 t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   775
  | "subst0 t (Eq a) = Eq (tmsubst0 t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   776
  | "subst0 t (NEq a) = NEq (tmsubst0 t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   777
  | "subst0 t (NOT p) = NOT (subst0 t p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   778
  | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   779
  | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   780
  | "subst0 t (Imp p q) = Imp (subst0 t p)  (subst0 t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   781
  | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   782
  | "subst0 t (E p) = E p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   783
  | "subst0 t (A p) = A p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   784
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   785
lemma subst0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   786
  assumes qf: "qfree p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   787
  shows "Ifm vs (x # bs) (subst0 t p) = Ifm vs ((Itm vs (x # bs) t) # bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   788
  using qf tmsubst0[where x="x" and bs="bs" and t="t"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   789
  by (induct p rule: fm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   790
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   791
lemma subst0_nb:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   792
  assumes bp: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   793
    and qf: "qfree p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   794
  shows "bound0 (subst0 t p)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   795
  using qf tmsubst0_nb[OF bp] bp by (induct p rule: fm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   796
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   797
primrec subst:: "nat \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   798
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   799
    "subst n t T = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   800
  | "subst n t F = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   801
  | "subst n t (Lt a) = Lt (tmsubst n t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   802
  | "subst n t (Le a) = Le (tmsubst n t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   803
  | "subst n t (Eq a) = Eq (tmsubst n t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   804
  | "subst n t (NEq a) = NEq (tmsubst n t a)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   805
  | "subst n t (NOT p) = NOT (subst n t p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   806
  | "subst n t (And p q) = And (subst n t p) (subst n t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   807
  | "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   808
  | "subst n t (Imp p q) = Imp (subst n t p)  (subst n t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   809
  | "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   810
  | "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   811
  | "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   812
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   813
lemma subst:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   814
  assumes nb: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   815
    and nlm: "n \<le> length bs"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   816
  shows "Ifm vs bs (subst n t p) = Ifm vs (bs[n:= Itm vs bs t]) p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   817
  using nb nlm
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   818
proof (induct p arbitrary: bs n t rule: fm.induct)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   819
  case (E p bs n)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   820
  have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   821
        Ifm vs (x#bs[n:= Itm vs bs t]) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   822
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   823
    from E have bn: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   824
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   825
    from E have nlm: "Suc n \<le> length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   826
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   827
    from E(1)[OF bn nlm]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   828
    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   829
        Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   830
      by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   831
    then show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   832
      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   833
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   834
  then show ?case by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   835
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   836
  case (A p bs n)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   837
  have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   838
        Ifm vs (x#bs[n:= Itm vs bs t]) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   839
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   840
    from A have bn: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   841
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   842
    from A have nlm: "Suc n \<le> length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   843
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   844
    from A(1)[OF bn nlm]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   845
    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   846
        Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   847
      by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   848
    then show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   849
      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   850
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   851
  then show ?case by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   852
qed (auto simp add: tmsubst)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   853
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   854
lemma subst_nb:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   855
  assumes "tmbound m t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   856
  shows "bound m (subst m t p)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   857
  using assms tmsubst_nb incrtm0_tmbound by (induct p arbitrary: m t rule: fm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   858
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   859
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   860
  by (induct p rule: not.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   861
lemma not_bn0[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   862
  by (induct p rule: not.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   863
lemma not_nb[simp]: "bound n p \<Longrightarrow> bound n (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   864
  by (induct p rule: not.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   865
lemma not_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   866
  by (induct p rule: not.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   867
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   868
lemma conj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (conj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   869
  using conj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   870
lemma conj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (conj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   871
  using conj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   872
lemma conj_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (conj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   873
  using conj_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   874
lemma conj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   875
  using conj_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   876
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   877
lemma disj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (disj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   878
  using disj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   879
lemma disj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (disj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   880
  using disj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   881
lemma disj_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (disj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   882
  using disj_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   883
lemma disj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (disj p q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   884
  using disj_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   885
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   886
lemma imp_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (imp p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   887
  using imp_def by (cases "p = F \<or> q = T") (simp_all add: imp_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   888
lemma imp_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (imp p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   889
  using imp_def by (cases "p = F \<or> q = T \<or> p = q") (simp_all add: imp_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   890
lemma imp_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (imp p q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   891
  using imp_def by (cases "p = F \<or> q = T \<or> p = q") (simp_all add: imp_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   892
lemma imp_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (imp p q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   893
  using imp_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   894
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   895
lemma iff_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   896
  unfolding iff_def by (cases "p = q") auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   897
lemma iff_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   898
  using iff_def unfolding iff_def by (cases "p = q") auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   899
lemma iff_nb[simp]: "bound n p \<Longrightarrow> bound n q \<Longrightarrow> bound n (iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   900
  using iff_def unfolding iff_def by (cases "p = q") auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   901
lemma iff_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (iff p q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   902
  using iff_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   903
lemma decr0_qf: "bound0 p \<Longrightarrow> qfree (decr0 p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   904
  by (induct p) simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   905
61586
5197a2ecb658 isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   906
fun isatom :: "fm \<Rightarrow> bool"  \<comment> \<open>test for atomicity\<close>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   907
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   908
    "isatom T = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   909
  | "isatom F = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   910
  | "isatom (Lt a) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   911
  | "isatom (Le a) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   912
  | "isatom (Eq a) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   913
  | "isatom (NEq a) = True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   914
  | "isatom p = False"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   915
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   916
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   917
  by (induct p) simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   918
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   919
definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   920
  where "djf f p q \<equiv>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   921
    (if q = T then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   922
     else if q = F then f p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   923
     else (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   924
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   925
definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   926
  where "evaldjf f ps \<equiv> foldr (djf f) ps F"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   927
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   928
lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   929
  apply (cases "q = T")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   930
   apply (simp add: djf_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   931
  apply (cases "q = F")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   932
   apply (simp add: djf_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   933
  apply (cases "f p")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   934
              apply (simp_all add: Let_def djf_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   935
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   936
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   937
lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm vs bs (f p))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   938
  by (induct ps) (simp_all add: evaldjf_def djf_Or)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   939
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   940
lemma evaldjf_bound0:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   941
  assumes "\<forall>x\<in> set xs. bound0 (f x)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   942
  shows "bound0 (evaldjf f xs)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   943
  using assms
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   944
  apply (induct xs)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   945
   apply (auto simp add: evaldjf_def djf_def Let_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   946
  apply (case_tac "f a")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   947
              apply auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   948
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   949
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   950
lemma evaldjf_qf:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   951
  assumes "\<forall>x\<in> set xs. qfree (f x)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   952
  shows "qfree (evaldjf f xs)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   953
  using assms
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   954
  apply (induct xs)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   955
   apply (auto simp add: evaldjf_def djf_def Let_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   956
  apply (case_tac "f a")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   957
              apply auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   958
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   959
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   960
fun disjuncts :: "fm \<Rightarrow> fm list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   961
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   962
    "disjuncts (Or p q) = disjuncts p @ disjuncts q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   963
  | "disjuncts F = []"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   964
  | "disjuncts p = [p]"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   965
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   966
lemma disjuncts: "(\<exists>q \<in> set (disjuncts p). Ifm vs bs q) = Ifm vs bs p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   967
  by (induct p rule: disjuncts.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   968
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   969
lemma disjuncts_nb:
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   970
  assumes "bound0 p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   971
  shows "\<forall>q \<in> set (disjuncts p). bound0 q"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   972
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   973
  from assms have "list_all bound0 (disjuncts p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   974
    by (induct p rule: disjuncts.induct) auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   975
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   976
    by (simp only: list_all_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   977
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   978
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   979
lemma disjuncts_qf:
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   980
  assumes "qfree p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   981
  shows "\<forall>q \<in> set (disjuncts p). qfree q"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   982
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   983
  from assms have "list_all qfree (disjuncts p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   984
    by (induct p rule: disjuncts.induct) auto
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   985
  then show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
   986
    by (simp only: list_all_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   987
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   988
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   989
definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   990
  where "DJ f p \<equiv> evaldjf f (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   991
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   992
lemma DJ:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   993
  assumes fdj: "\<forall>p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   994
    and fF: "f F = F"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   995
  shows "Ifm vs bs (DJ f p) = Ifm vs bs (f p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   996
proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   997
  have "Ifm vs bs (DJ f p) = (\<exists>q \<in> set (disjuncts p). Ifm vs bs (f q))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   998
    by (simp add: DJ_def evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   999
  also have "\<dots> = Ifm vs bs (f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1000
    using fdj fF by (induct p rule: disjuncts.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1001
  finally show ?thesis .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1002
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1003
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1004
lemma DJ_qf:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1005
  assumes fqf: "\<forall>p. qfree p \<longrightarrow> qfree (f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1006
  shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1007
proof clarify
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1008
  fix  p
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1009
  assume qf: "qfree p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1010
  have th: "DJ f p = evaldjf f (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1011
    by (simp add: DJ_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1012
  from disjuncts_qf[OF qf] have "\<forall>q\<in> set (disjuncts p). qfree q" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1013
  with fqf have th':"\<forall>q\<in> set (disjuncts p). qfree (f q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1014
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1015
  from evaldjf_qf[OF th'] th show "qfree (DJ f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1016
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1017
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1018
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1019
lemma DJ_qe:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1020
  assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1021
  shows "\<forall>bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm vs bs ((DJ qe p)) = Ifm vs bs (E p))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1022
proof clarify
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1023
  fix p :: fm and bs
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1024
  assume qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1025
  from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1026
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1027
  from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1028
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1029
  have "Ifm vs bs (DJ qe p) \<longleftrightarrow> (\<exists>q\<in> set (disjuncts p). Ifm vs bs (qe q))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1030
    by (simp add: DJ_def evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1031
  also have "\<dots> = (\<exists>q \<in> set(disjuncts p). Ifm vs bs (E q))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1032
    using qe disjuncts_qf[OF qf] by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1033
  also have "\<dots> = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1034
    by (induct p rule: disjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1035
  finally show "qfree (DJ qe p) \<and> Ifm vs bs (DJ qe p) = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1036
    using qfth by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1037
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1038
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1039
fun conjuncts :: "fm \<Rightarrow> fm list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1040
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1041
    "conjuncts (And p q) = conjuncts p @ conjuncts q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1042
  | "conjuncts T = []"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1043
  | "conjuncts p = [p]"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1044
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1045
definition list_conj :: "fm list \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1046
  where "list_conj ps \<equiv> foldr conj ps T"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1047
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1048
definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1049
  where "CJNB f p \<equiv>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1050
    (let cjs = conjuncts p;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1051
      (yes, no) = partition bound0 cjs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1052
     in conj (decr0 (list_conj yes)) (f (list_conj no)))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1053
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1054
lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall>q \<in> set (conjuncts p). qfree q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1055
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1056
  assume qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1057
  then have "list_all qfree (conjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1058
    by (induct p rule: conjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1059
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1060
    by (simp only: list_all_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1061
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1062
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1063
lemma conjuncts: "(\<forall>q\<in> set (conjuncts p). Ifm vs bs q) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1064
  by (induct p rule: conjuncts.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1065
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1066
lemma conjuncts_nb:
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1067
  assumes "bound0 p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1068
  shows "\<forall>q \<in> set (conjuncts p). bound0 q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1069
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1070
  from assms have "list_all bound0 (conjuncts p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1071
    by (induct p rule:conjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1072
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1073
    by (simp only: list_all_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1074
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1075
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1076
fun islin :: "fm \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1077
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1078
    "islin (And p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1079
  | "islin (Or p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1080
  | "islin (Eq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1081
  | "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1082
  | "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1083
  | "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1084
  | "islin (NOT p) = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1085
  | "islin (Imp p q) = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1086
  | "islin (Iff p q) = False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1087
  | "islin p = bound0 p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1088
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1089
lemma islin_stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1090
  assumes nb: "tmbound0 p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1091
  shows "islin (Lt p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1092
    and "islin (Le p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1093
    and "islin (Eq p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1094
    and "islin (NEq p)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1095
  using nb by (cases p, auto, rename_tac nat a b, case_tac nat, auto)+
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1096
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1097
definition "lt p = (case p of CP (C c) \<Rightarrow> if 0>\<^sub>N c then T else F| _ \<Rightarrow> Lt p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1098
definition "le p = (case p of CP (C c) \<Rightarrow> if 0\<ge>\<^sub>N c then T else F | _ \<Rightarrow> Le p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1099
definition "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1100
definition "neq p = not (eq p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1101
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1102
lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1103
  apply (simp add: lt_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1104
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1105
        apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1106
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1107
         apply (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1108
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1109
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1110
lemma le: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (le p) = Ifm vs bs (Le p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1111
  apply (simp add: le_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1112
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1113
        apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1114
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1115
         apply (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1116
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1117
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1118
lemma eq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (eq p) = Ifm vs bs (Eq p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1119
  apply (simp add: eq_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1120
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1121
        apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1122
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1123
         apply (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1124
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1125
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1126
lemma neq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (neq p) = Ifm vs bs (NEq p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1127
  by (simp add: neq_def eq)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1128
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1129
lemma lt_lin: "tmbound0 p \<Longrightarrow> islin (lt p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1130
  apply (simp add: lt_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1131
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1132
        apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1133
   apply (rename_tac poly, case_tac poly)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1134
          apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1135
  apply (rename_tac nat a b, case_tac nat)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1136
   apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1137
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1138
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1139
lemma le_lin: "tmbound0 p \<Longrightarrow> islin (le p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1140
  apply (simp add: le_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1141
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1142
        apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1143
   apply (rename_tac poly, case_tac poly)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1144
          apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1145
  apply (rename_tac nat a b, case_tac nat)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1146
   apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1147
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1148
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1149
lemma eq_lin: "tmbound0 p \<Longrightarrow> islin (eq p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1150
  apply (simp add: eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1151
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1152
        apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1153
   apply (rename_tac poly, case_tac poly)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1154
          apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1155
  apply (rename_tac nat a b, case_tac nat)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1156
   apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1157
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1158
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1159
lemma neq_lin: "tmbound0 p \<Longrightarrow> islin (neq p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1160
  apply (simp add: neq_def eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1161
  apply (cases p)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1162
        apply simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1163
   apply (rename_tac poly, case_tac poly)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1164
          apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1165
  apply (rename_tac nat a b, case_tac nat)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1166
   apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1167
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1168
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1169
definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1170
definition "simple t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then le s else Le (CNP 0 c s))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1171
definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1172
definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1173
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1174
lemma simplt_islin [simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1175
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1176
  shows "islin (simplt t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1177
  unfolding simplt_def
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1178
  using split0_nb0'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1179
  by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1180
      islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1181
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1182
lemma simple_islin [simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1183
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1184
  shows "islin (simple t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1185
  unfolding simple_def
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1186
  using split0_nb0'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1187
  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1188
      islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1189
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1190
lemma simpeq_islin [simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1191
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1192
  shows "islin (simpeq t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1193
  unfolding simpeq_def
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1194
  using split0_nb0'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1195
  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1196
      islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1197
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1198
lemma simpneq_islin [simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1199
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1200
  shows "islin (simpneq t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1201
  unfolding simpneq_def
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1202
  using split0_nb0'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1203
  by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1204
      islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1205
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1206
lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1207
  by (cases "split0 s") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1208
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1209
lemma split0_npoly:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1210
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1211
    and *: "allpolys isnpoly t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1212
  shows "isnpoly (fst (split0 t))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1213
    and "allpolys isnpoly (snd (split0 t))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1214
  using *
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1215
  by (induct t rule: split0.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1216
    (auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1217
      polysub_norm really_stupid)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1218
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1219
lemma simplt[simp]: "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1220
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1221
  have *: "allpolys isnpoly (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1222
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1223
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1224
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1225
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1226
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1227
    then show ?thesis
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1228
      using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF *], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1229
      by (simp add: simplt_def Let_def split_def lt)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1230
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1231
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1232
    then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1233
      using split0[of "simptm t" vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1234
      by (simp add: simplt_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1235
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1236
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1237
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1238
lemma simple[simp]: "Ifm vs bs (simple t) = Ifm vs bs (Le t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1239
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1240
  have *: "allpolys isnpoly (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1241
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1242
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1243
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1244
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1245
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1246
    then show ?thesis
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1247
      using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF *], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1248
      by (simp add: simple_def Let_def split_def le)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1249
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1250
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1251
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1252
      using split0[of "simptm t" vs bs]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1253
      by (simp add: simple_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1254
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1255
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1256
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1257
lemma simpeq[simp]: "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1258
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1259
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1260
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1261
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1262
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1263
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1264
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1265
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1266
      using split0[of "simptm t" vs bs] eq[OF split0_npoly(2)[OF n], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1267
      by (simp add: simpeq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1268
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1269
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1270
    then show ?thesis using  split0[of "simptm t" vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1271
      by (simp add: simpeq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1272
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1273
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1274
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1275
lemma simpneq[simp]: "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1276
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1277
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1278
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1279
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1280
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1281
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1282
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1283
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1284
      using split0[of "simptm t" vs bs] neq[OF split0_npoly(2)[OF n], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1285
      by (simp add: simpneq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1286
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1287
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1288
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1289
      using split0[of "simptm t" vs bs] by (simp add: simpneq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1290
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1291
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1292
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1293
lemma lt_nb: "tmbound0 t \<Longrightarrow> bound0 (lt t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1294
  apply (simp add: lt_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1295
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1296
        apply auto
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1297
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1298
         apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1299
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1300
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1301
lemma le_nb: "tmbound0 t \<Longrightarrow> bound0 (le t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1302
  apply (simp add: le_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1303
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1304
        apply auto
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1305
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1306
         apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1307
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1308
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1309
lemma eq_nb: "tmbound0 t \<Longrightarrow> bound0 (eq t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1310
  apply (simp add: eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1311
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1312
        apply auto
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1313
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1314
         apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1315
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1316
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1317
lemma neq_nb: "tmbound0 t \<Longrightarrow> bound0 (neq t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1318
  apply (simp add: neq_def eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1319
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1320
        apply auto
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1321
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1322
         apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1323
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1324
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1325
lemma simplt_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1326
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1327
  shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1328
proof (simp add: simplt_def Let_def split_def)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1329
  assume "tmbound0 t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1330
  then have *: "tmbound0 (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1331
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1332
  let ?c = "fst (split0 (simptm t))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1333
  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1334
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1335
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1336
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1337
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1338
    by (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1339
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1340
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1341
  then show "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (lt (snd (split0 (simptm t))))) \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1342
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1343
    by (simp add: simplt_def Let_def split_def lt_nb)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1344
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1345
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1346
lemma simple_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1347
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1348
  shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1349
proof(simp add: simple_def Let_def split_def)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1350
  assume "tmbound0 t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1351
  then have *: "tmbound0 (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1352
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1353
  let ?c = "fst (split0 (simptm t))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1354
  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1355
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1356
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1357
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1358
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1359
    by (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1360
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1361
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1362
  then show "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (le (snd (split0 (simptm t))))) \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1363
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1364
    by (simp add: simplt_def Let_def split_def le_nb)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1365
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1366
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1367
lemma simpeq_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1368
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1369
  shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1370
proof (simp add: simpeq_def Let_def split_def)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1371
  assume "tmbound0 t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1372
  then have *: "tmbound0 (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1373
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1374
  let ?c = "fst (split0 (simptm t))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1375
  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1376
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1377
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1378
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1379
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1380
    by (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1381
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1382
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1383
  then show "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (eq (snd (split0 (simptm t))))) \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1384
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1385
    by (simp add: simpeq_def Let_def split_def eq_nb)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1386
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1387
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1388
lemma simpneq_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1389
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1390
  shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1391
proof (simp add: simpneq_def Let_def split_def)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1392
  assume "tmbound0 t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1393
  then have *: "tmbound0 (simptm t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1394
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1395
  let ?c = "fst (split0 (simptm t))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1396
  from tmbound_split0[OF *[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1397
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1398
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1399
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1400
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1401
    by (simp_all add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1402
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1403
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1404
  then show "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (neq (snd (split0 (simptm t))))) \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1405
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1406
    by (simp add: simpneq_def Let_def split_def neq_nb)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1407
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1408
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1409
fun conjs :: "fm \<Rightarrow> fm list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1410
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1411
    "conjs (And p q) = conjs p @ conjs q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1412
  | "conjs T = []"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1413
  | "conjs p = [p]"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1414
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1415
lemma conjs_ci: "(\<forall>q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1416
  by (induct p rule: conjs.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1417
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1418
definition list_disj :: "fm list \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1419
  where "list_disj ps \<equiv> foldr disj ps F"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1420
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1421
lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1422
  by (induct ps) (auto simp add: list_conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1423
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1424
lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1425
  by (induct ps) (auto simp add: list_conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1426
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1427
lemma list_disj: "Ifm vs bs (list_disj ps) = (\<exists>p\<in> set ps. Ifm vs bs p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1428
  by (induct ps) (auto simp add: list_disj_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1429
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1430
lemma conj_boundslt: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1431
  unfolding conj_def by auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1432
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1433
lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1434
  apply (induct p rule: conjs.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1435
              apply (unfold conjs.simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1436
              apply (unfold set_append)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1437
              apply (unfold ball_Un)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1438
              apply (unfold bound.simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1439
              apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1440
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1441
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1442
lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1443
  apply (induct p rule: conjs.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1444
              apply (unfold conjs.simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1445
              apply (unfold set_append)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1446
              apply (unfold ball_Un)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1447
              apply (unfold boundslt.simps)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1448
              apply blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1449
             apply simp_all
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1450
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1451
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1452
lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1453
  by (induct ps) (auto simp: list_conj_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1454
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1455
lemma list_conj_nb:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1456
  assumes "\<forall>p\<in> set ps. bound n p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1457
  shows "bound n (list_conj ps)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1458
  using assms by (induct ps) (auto simp: list_conj_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1459
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1460
lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1461
  by (induct ps) (auto simp: list_conj_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1462
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1463
lemma CJNB_qe:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1464
  assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1465
  shows "\<forall>bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm vs bs ((CJNB qe p)) = Ifm vs bs (E p))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1466
proof clarify
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1467
  fix bs p
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1468
  assume qfp: "qfree p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1469
  let ?cjs = "conjuncts p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1470
  let ?yes = "fst (partition bound0 ?cjs)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1471
  let ?no = "snd (partition bound0 ?cjs)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1472
  let ?cno = "list_conj ?no"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1473
  let ?cyes = "list_conj ?yes"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1474
  have part: "partition bound0 ?cjs = (?yes,?no)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1475
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1476
  from partition_P[OF part] have "\<forall>q\<in> set ?yes. bound0 q"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1477
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1478
  then have yes_nb: "bound0 ?cyes"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1479
    by (simp add: list_conj_nb')
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1480
  then have yes_qf: "qfree (decr0 ?cyes)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1481
    by (simp add: decr0_qf)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1482
  from conjuncts_qf[OF qfp] partition_set[OF part]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1483
  have " \<forall>q\<in> set ?no. qfree q"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1484
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1485
  then have no_qf: "qfree ?cno"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1486
    by (simp add: list_conj_qf)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1487
  with qe have cno_qf:"qfree (qe ?cno)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1488
    and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1489
    by blast+
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1490
  from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1491
    by (simp add: CJNB_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1492
  have "Ifm vs bs p = ((Ifm vs bs ?cyes) \<and> (Ifm vs bs ?cno))" for bs
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1493
  proof -
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1494
    from conjuncts have "Ifm vs bs p = (\<forall>q\<in> set ?cjs. Ifm vs bs q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1495
      by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1496
    also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm vs bs q) \<and> (\<forall>q\<in> set ?no. Ifm vs bs q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1497
      using partition_set[OF part] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1498
    finally show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1499
      using list_conj[of vs bs] by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1500
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1501
  then have "Ifm vs bs (E p) = (\<exists>x. (Ifm vs (x#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1502
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1503
  also fix y have "\<dots> = (\<exists>x. (Ifm vs (y#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1504
    using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1505
  also have "\<dots> = (Ifm vs bs (decr0 ?cyes) \<and> Ifm vs bs (E ?cno))"
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33268
diff changeset
  1506
    by (auto simp add: decr0[OF yes_nb] simp del: partition_filter_conv)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1507
  also have "\<dots> = (Ifm vs bs (conj (decr0 ?cyes) (qe ?cno)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1508
    using qe[rule_format, OF no_qf] by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1509
  finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1510
    by (simp add: Let_def CJNB_def split_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1511
  with qf show "qfree (CJNB qe p) \<and> Ifm vs bs (CJNB qe p) = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1512
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1513
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1514
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1515
fun simpfm :: "fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1516
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1517
    "simpfm (Lt t) = simplt (simptm t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1518
  | "simpfm (Le t) = simple (simptm t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1519
  | "simpfm (Eq t) = simpeq(simptm t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1520
  | "simpfm (NEq t) = simpneq(simptm t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1521
  | "simpfm (And p q) = conj (simpfm p) (simpfm q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1522
  | "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1523
  | "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1524
  | "simpfm (Iff p q) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1525
      disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1526
  | "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1527
  | "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1528
  | "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1529
  | "simpfm (NOT (Iff p q)) =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1530
      disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1531
  | "simpfm (NOT (Eq t)) = simpneq t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1532
  | "simpfm (NOT (NEq t)) = simpeq t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1533
  | "simpfm (NOT (Le t)) = simplt (Neg t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1534
  | "simpfm (NOT (Lt t)) = simple (Neg t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1535
  | "simpfm (NOT (NOT p)) = simpfm p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1536
  | "simpfm (NOT T) = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1537
  | "simpfm (NOT F) = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1538
  | "simpfm p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1539
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1540
lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1541
  by (induct p arbitrary: bs rule: simpfm.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1542
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1543
lemma simpfm_bound0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1544
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1545
  shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1546
  by (induct p rule: simpfm.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1547
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1548
lemma lt_qf[simp]: "qfree (lt t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1549
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1550
        apply (auto simp add: lt_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1551
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1552
         apply auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1553
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1554
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1555
lemma le_qf[simp]: "qfree (le t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1556
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1557
        apply (auto simp add: le_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1558
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1559
         apply auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1560
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1561
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1562
lemma eq_qf[simp]: "qfree (eq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1563
  apply (cases t)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1564
        apply (auto simp add: eq_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1565
  apply (rename_tac poly, case_tac poly)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1566
         apply auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1567
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1568
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1569
lemma neq_qf[simp]: "qfree (neq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1570
  by (simp add: neq_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1571
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1572
lemma simplt_qf[simp]: "qfree (simplt t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1573
  by (simp add: simplt_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1574
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1575
lemma simple_qf[simp]: "qfree (simple t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1576
  by (simp add: simple_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1577
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1578
lemma simpeq_qf[simp]: "qfree (simpeq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1579
  by (simp add: simpeq_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1580
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1581
lemma simpneq_qf[simp]: "qfree (simpneq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1582
  by (simp add: simpneq_def Let_def split_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1583
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1584
lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1585
  by (induct p rule: simpfm.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1586
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1587
lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1588
  by (simp add: disj_def)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1589
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1590
lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1591
  by (simp add: conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1592
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1593
lemma
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1594
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1595
  shows "qfree p \<Longrightarrow> islin (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1596
  by (induct p rule: simpfm.induct) (simp_all add: conj_lin disj_lin)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1597
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1598
fun prep :: "fm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1599
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1600
    "prep (E T) = T"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1601
  | "prep (E F) = F"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1602
  | "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1603
  | "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1604
  | "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1605
  | "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1606
  | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1607
  | "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1608
  | "prep (E p) = E (prep p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1609
  | "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1610
  | "prep (A p) = prep (NOT (E (NOT p)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1611
  | "prep (NOT (NOT p)) = prep p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1612
  | "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1613
  | "prep (NOT (A p)) = prep (E (NOT p))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1614
  | "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1615
  | "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1616
  | "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1617
  | "prep (NOT p) = not (prep p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1618
  | "prep (Or p q) = disj (prep p) (prep q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1619
  | "prep (And p q) = conj (prep p) (prep q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1620
  | "prep (Imp p q) = prep (Or (NOT p) q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1621
  | "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1622
  | "prep p = p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1623
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1624
lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1625
  by (induct p arbitrary: bs rule: prep.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1626
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1627
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1628
text \<open>Generic quantifier elimination.\<close>
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  1629
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1630
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1631
    "qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1632
  | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1633
  | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1634
  | "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1635
  | "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1636
  | "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1637
  | "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1638
  | "qelim p = (\<lambda>y. simpfm p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1639
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1640
lemma qelim:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1641
  assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1642
  shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm vs bs (qelim p qe) = Ifm vs bs p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1643
  using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1644
  by (induct p rule: qelim.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1645
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1646
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  1647
subsection \<open>Core Procedure\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1648
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1649
fun minusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>virtual substitution of \<open>-\<infinity>\<close>\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1650
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1651
    "minusinf (And p q) = conj (minusinf p) (minusinf q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1652
  | "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1653
  | "minusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1654
  | "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1655
  | "minusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1656
  | "minusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1657
  | "minusinf p = p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1658
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1659
fun plusinf:: "fm \<Rightarrow> fm"  \<comment> \<open>virtual substitution of \<open>+\<infinity>\<close>\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1660
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1661
    "plusinf (And p q) = conj (plusinf p) (plusinf q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1662
  | "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1663
  | "plusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1664
  | "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1665
  | "plusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1666
  | "plusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1667
  | "plusinf p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1668
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1669
lemma minusinf_inf:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1670
  assumes "islin p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1671
  shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1672
  using assms
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1673
proof (induct p rule: minusinf.induct)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1674
  case 1
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1675
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1676
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1677
    apply (rule_tac x="min z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1678
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1679
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1680
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1681
  case 2
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1682
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1683
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1684
    apply (rule_tac x="min z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1685
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1686
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1687
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1688
  case (3 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1689
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1690
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1691
  from 3 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1692
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1693
  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1694
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1695
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1696
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1697
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1698
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1699
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1700
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1701
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1702
      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1703
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1704
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1705
    have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1706
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1707
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1708
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1709
        using pos_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1710
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1711
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1712
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1713
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1714
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1715
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1716
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1717
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1718
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1719
    have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1720
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1721
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1722
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1723
        using neg_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1724
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1725
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1726
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1727
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1728
        using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1729
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1730
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1731
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1732
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1733
  case (4 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1734
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1735
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1736
  from 4 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1737
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1738
  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1739
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1740
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1741
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1742
  consider "?c = 0" | "?c > 0" | "?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1743
    by arith
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1744
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1745
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1746
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1747
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1748
      using eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1749
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1750
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1751
    have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1752
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1753
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1754
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1755
        using pos_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1756
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1757
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1758
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1759
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1760
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1761
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1762
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1763
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1764
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1765
    have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1766
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1767
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1768
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1769
        using neg_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1770
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1771
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1772
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1773
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1774
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1775
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1776
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1777
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1778
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1779
  case (5 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1780
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1781
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1782
  from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1783
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1784
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1785
    by (simp add: polyneg_norm)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1786
  note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1787
  let ?c = "Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1788
  fix y
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1789
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1790
  consider "?c = 0" | "?c > 0" | "?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1791
    by arith
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1792
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1793
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1794
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1795
    then show ?thesis using eqs by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1796
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1797
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1798
    have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1799
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1800
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1801
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1802
        using pos_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1803
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1804
      then have "?c * x + ?e < 0" by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1805
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1806
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1807
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1808
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1809
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1810
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1811
    have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1812
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1813
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1814
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1815
        using neg_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1816
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1817
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1818
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1819
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1820
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1821
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1822
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1823
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1824
next
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1825
  case (6 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1826
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1827
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1828
  from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1829
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1830
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1831
    by (simp add: polyneg_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1832
  note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1833
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1834
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1835
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1836
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1837
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1838
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1839
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1840
    then show ?thesis using eqs by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1841
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1842
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1843
    have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1844
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1845
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1846
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1847
        using pos_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1848
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1849
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1850
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1851
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1852
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1853
        by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1854
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1855
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1856
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1857
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1858
    have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1859
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1860
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1861
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1862
        using neg_less_divide_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1863
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1864
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1865
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1866
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1867
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1868
        by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1869
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1870
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1871
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1872
qed auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1873
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1874
lemma plusinf_inf:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1875
  assumes "islin p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1876
  shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  1877
  using assms
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1878
proof (induct p rule: plusinf.induct)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1879
  case 1
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1880
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1881
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1882
    apply (rule_tac x="max z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1883
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1884
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1885
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1886
  case 2
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1887
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1888
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1889
    apply (rule_tac x="max z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1890
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1891
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1892
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1893
  case (3 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1894
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1895
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1896
  from 3 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1897
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1898
  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1899
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1900
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1901
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1902
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1903
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1904
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1905
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1906
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1907
      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1908
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1909
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1910
    have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1911
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1912
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1913
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1914
        using pos_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1915
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1916
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1917
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1918
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1919
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1920
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1921
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1922
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1923
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1924
    have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1925
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1926
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1927
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1928
        using neg_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1929
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1930
      then have "?c * x + ?e < 0" by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1931
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1932
        using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1933
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1934
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1935
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1936
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1937
  case (4 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1938
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1939
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1940
  from 4 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1941
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1942
  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1943
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1944
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1945
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1946
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1947
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1948
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1949
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1950
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1951
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1952
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1953
    have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1954
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1955
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1956
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1957
        using pos_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1958
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1959
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1960
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1961
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1962
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1963
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1964
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1965
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1966
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1967
    have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1968
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1969
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1970
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1971
        using neg_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  1972
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1973
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1974
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1975
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1976
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1977
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1978
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1979
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1980
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1981
  case (5 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1982
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1983
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1984
  from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1985
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1986
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1987
    by (simp add: polyneg_norm)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1988
  note eqs = lt[OF nc(1), where ?'a = 'a] lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1989
  let ?c = "Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1990
  fix y
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1991
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1992
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1993
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1994
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1995
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1996
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1997
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1998
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1999
    have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2000
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2001
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2002
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2003
        using pos_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2004
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2005
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2006
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2007
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2008
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2009
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2010
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2011
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2012
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2013
    have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2014
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2015
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2016
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2017
        using neg_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2018
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2019
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2020
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2021
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2022
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2023
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2024
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2025
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2026
next
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2027
  case (6 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2028
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2029
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2030
  from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2031
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2032
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2033
    by (simp add: polyneg_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2034
  note eqs = lt[OF nc(1), where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2035
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2036
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2037
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2038
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2039
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2040
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2041
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2042
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2043
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2044
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2045
    have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2046
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2047
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2048
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2049
        using pos_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2050
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2051
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2052
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2053
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2054
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2055
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2056
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2057
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2058
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2059
    have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2060
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2061
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2062
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2063
        using neg_divide_less_eq[OF c, where a="x" and b="-?e"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2064
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2065
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2066
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2067
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2068
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2069
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2070
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2071
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2072
qed auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2073
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2074
lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2075
  by (induct p rule: minusinf.induct) (auto simp add: eq_nb lt_nb le_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2076
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2077
lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2078
  by (induct p rule: minusinf.induct) (auto simp add: eq_nb lt_nb le_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2079
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2080
lemma minusinf_ex:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2081
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2082
    and ex: "Ifm vs (x#bs) (minusinf p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2083
  shows "\<exists>x. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2084
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2085
  from bound0_I [OF minusinf_nb[OF lp], where bs ="bs"] ex
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2086
  have th: "\<forall>x. Ifm vs (x#bs) (minusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2087
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2088
  from minusinf_inf[OF lp, where bs="bs"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2089
  obtain z where z: "\<forall>x<z. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2090
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2091
  from th have "Ifm vs ((z - 1)#bs) (minusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2092
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2093
  moreover have "z - 1 < z"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2094
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2095
  ultimately show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2096
    using z by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2097
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2098
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2099
lemma plusinf_ex:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2100
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2101
    and ex: "Ifm vs (x#bs) (plusinf p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2102
  shows "\<exists>x. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2103
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2104
  from bound0_I [OF plusinf_nb[OF lp], where bs ="bs"] ex
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2105
  have th: "\<forall>x. Ifm vs (x#bs) (plusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2106
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2107
  from plusinf_inf[OF lp, where bs="bs"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2108
  obtain z where z: "\<forall>x>z. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2109
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2110
  from th have "Ifm vs ((z + 1)#bs) (plusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2111
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2112
  moreover have "z + 1 > z"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2113
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2114
  ultimately show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2115
    using z by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2116
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2117
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2118
fun uset :: "fm \<Rightarrow> (poly \<times> tm) list"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2119
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2120
    "uset (And p q) = uset p @ uset q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2121
  | "uset (Or p q) = uset p @ uset q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2122
  | "uset (Eq (CNP 0 a e)) = [(a, e)]"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2123
  | "uset (Le (CNP 0 a e)) = [(a, e)]"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2124
  | "uset (Lt (CNP 0 a e)) = [(a, e)]"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2125
  | "uset (NEq (CNP 0 a e)) = [(a, e)]"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2126
  | "uset p = []"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2127
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2128
lemma uset_l:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2129
  assumes lp: "islin p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2130
  shows "\<forall>(c,s) \<in> set (uset p). isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2131
  using lp by (induct p rule: uset.induct) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2132
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2133
lemma minusinf_uset0:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2134
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2135
    and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2136
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2137
  shows "\<exists>(c, s) \<in> set (uset p). x \<ge> - Itm vs (x#bs) s / Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2138
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2139
  have "\<exists>(c, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2140
      Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2141
      Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2142
    using lp nmi ex
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2143
    apply (induct p rule: minusinf.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2144
                        apply (auto simp add: eq le lt polyneg_norm)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2145
      apply (auto simp add: linorder_not_less order_le_less)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2146
    done
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2147
  then obtain c s where csU: "(c, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2148
    and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2149
      (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2150
  then have "x \<ge> (- Itm vs (x#bs) s) / Ipoly vs c"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2151
    using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2152
    by (auto simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2153
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2154
    using csU by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2155
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2156
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2157
lemma minusinf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2158
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2159
    and nmi: "\<not> (Ifm vs (a#bs) (minusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2160
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2161
  shows "\<exists>(c,s) \<in> set (uset p). x \<ge> - Itm vs (a#bs) s / Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2162
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2163
  from nmi have nmi': "\<not> Ifm vs (x#bs) (minusinf p)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2164
    by (simp add: bound0_I[OF minusinf_nb[OF lp], where b=x and b'=a])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2165
  from minusinf_uset0[OF lp nmi' ex]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2166
  obtain c s where csU: "(c,s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2167
    and th: "x \<ge> - Itm vs (x#bs) s / Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2168
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2169
  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2170
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2171
  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2172
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2173
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2174
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2175
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2176
lemma plusinf_uset0:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2177
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2178
    and nmi: "\<not> (Ifm vs (x#bs) (plusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2179
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2180
  shows "\<exists>(c, s) \<in> set (uset p). x \<le> - Itm vs (x#bs) s / Ipoly vs c"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  2181
proof -
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2182
  have "\<exists>(c, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2183
      Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2184
      Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2185
    using lp nmi ex
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2186
    apply (induct p rule: minusinf.induct)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2187
                        apply (auto simp add: eq le lt polyneg_norm)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2188
      apply (auto simp add: linorder_not_less order_le_less)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2189
    done
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2190
  then obtain c s
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2191
    where c_s: "(c, s) \<in> set (uset p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2192
      and "Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s \<or>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2193
        Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2194
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2195
  then have "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2196
    using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  2197
    by (auto simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2198
  then show ?thesis
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2199
    using c_s by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2200
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2201
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2202
lemma plusinf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2203
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2204
    and nmi: "\<not> (Ifm vs (a#bs) (plusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2205
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2206
  shows "\<exists>(c,s) \<in> set (uset p). x \<le> - Itm vs (a#bs) s / Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2207
proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2208
  from nmi have nmi': "\<not> (Ifm vs (x#bs) (plusinf p))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2209
    by (simp add: bound0_I[OF plusinf_nb[OF lp], where b=x and b'=a])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2210
  from plusinf_uset0[OF lp nmi' ex]
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2211
  obtain c s
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2212
    where c_s: "(c,s) \<in> set (uset p)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2213
      and x: "x \<le> - Itm vs (x#bs) s / Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2214
    by blast
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2215
  from uset_l[OF lp, rule_format, OF c_s] have nb: "tmbound0 s"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2216
    by simp
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2217
  from x tmbound0_I[OF nb, of vs x bs a] c_s show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2218
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2219
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2220
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2221
lemma lin_dense:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2222
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2223
    and noS: "\<forall>t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda>(c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2224
      (is "\<forall>t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda>(c,t). - ?Nt x t / ?N c) ` ?U p")
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2225
    and lx: "l < x" and xu: "x < u"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2226
    and px: "Ifm vs (x # bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2227
    and ly: "l < y" and yu: "y < u"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2228
  shows "Ifm vs (y#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2229
  using lp px noS
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2230
proof (induct p rule: islin.induct)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2231
  case (5 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2232
  from "5.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2233
  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2234
    and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2235
    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2236
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2237
  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2238
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2239
  then have ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2240
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2241
  consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2242
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2243
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2244
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2245
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2246
      using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2247
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2248
    case N: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2249
    from px pos_less_divide_eq[OF N, where a="x" and b="-?Nt x s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2250
    have px': "x < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2251
      by (auto simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2252
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2253
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2254
      assume y: "y < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2255
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2256
        by (simp add: pos_less_divide_eq[OF N, where a="y" and b="-?Nt x s", symmetric])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2257
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2258
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2259
      then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2260
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2261
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2262
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2263
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2264
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2265
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2266
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2267
      with lx px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2268
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2269
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2270
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2271
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2272
    case N: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2273
    from px neg_divide_less_eq[OF N, where a="x" and b="-?Nt x s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2274
    have px': "x > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2275
      by (auto simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2276
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2277
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2278
      assume y: "y > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2279
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2280
        by (simp add: neg_divide_less_eq[OF N, where a="y" and b="-?Nt x s", symmetric])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2281
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2282
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2283
      then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2284
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2285
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2286
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2287
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2288
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2289
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2290
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2291
      with xu px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2292
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2293
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2294
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2295
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2296
next
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2297
  case (6 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2298
  from "6.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2299
  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2300
    and px: "Ifm vs (x # bs) (Le (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2301
    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2302
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2303
  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2304
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2305
  then have ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2306
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2307
  have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2308
  consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2309
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2310
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2311
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2312
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2313
      using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2314
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2315
    case N: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2316
    from px pos_le_divide_eq[OF N, where a="x" and b="-?Nt x s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2317
    have px': "x \<le> - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2318
      by (simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2319
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2320
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2321
      assume y: "y < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2322
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2323
        by (simp add: pos_less_divide_eq[OF N, where a="y" and b="-?Nt x s", symmetric])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2324
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2325
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2326
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2327
        using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2328
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2329
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2330
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2331
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2332
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2333
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2334
      with lx px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2335
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2336
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2337
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2338
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2339
    case N: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2340
    from px neg_divide_le_eq[OF N, where a="x" and b="-?Nt x s"]
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2341
    have px': "x \<ge> - ?Nt x s / ?N c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2342
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2343
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2344
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2345
      assume y: "y > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2346
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2347
        by (simp add: neg_divide_less_eq[OF N, where a="y" and b="-?Nt x s", symmetric])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2348
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2349
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2350
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2351
        using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2352
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2353
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2354
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2355
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2356
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2357
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2358
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2359
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2360
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2361
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2362
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2363
  case (3 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2364
  from "3.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2365
  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2366
    and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2367
    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2368
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2369
  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2370
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2371
  then have ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2372
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2373
  consider "?N c = 0" | "?N c < 0" | "?N c > 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2374
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2375
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2376
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2377
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2378
      using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2379
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2380
    case 2
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2381
    then have cnz: "?N c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2382
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2383
    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2384
    have px': "x = - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2385
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2386
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2387
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2388
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2389
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2390
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2391
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2392
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2393
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2394
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2395
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2396
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2397
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2398
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2399
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2400
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2401
      with lx px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2402
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2403
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2404
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2405
    case 3
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2406
    then have cnz: "?N c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2407
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2408
    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2409
    have px': "x = - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2410
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2411
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2412
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2413
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2414
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2415
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2416
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2417
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2418
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2419
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2420
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2421
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2422
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2423
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2424
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2425
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2426
      with lx px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2427
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2428
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2429
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2430
next
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2431
  case (4 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2432
  from "4.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2433
  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2434
    and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2435
    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2436
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2437
  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2438
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2439
  then have ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2440
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2441
  show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2442
  proof (cases "?N c = 0")
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2443
    case True
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2444
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2445
      using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2446
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2447
    case False
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2448
    with yne eq_divide_eq[of "y" "- ?Nt x s" "?N c"]
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2449
    show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2450
      by (simp add: field_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric])
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2451
  qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2452
qed (auto simp add: tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"]
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2453
  bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2454
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2455
lemma inf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2456
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2457
    and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))" (is "\<not> (Ifm vs (x#bs) (?M p))")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2458
    and npi: "\<not> (Ifm vs (x#bs) (plusinf p))" (is "\<not> (Ifm vs (x#bs) (?P p))")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2459
    and ex: "\<exists>x.  Ifm vs (x#bs) p" (is "\<exists>x. ?I x p")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2460
  shows "\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2461
    ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / 2) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2462
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2463
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2464
  let ?N = "Ipoly vs"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2465
  let ?U = "set (uset p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2466
  from ex obtain a where pa: "?I a p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2467
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2468
  from bound0_I[OF minusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] nmi
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2469
  have nmi': "\<not> (?I a (?M p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2470
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2471
  from bound0_I[OF plusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] npi
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2472
  have npi': "\<not> (?I a (?P p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2473
    by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2474
  have "\<exists>(c,t) \<in> set (uset p). \<exists>(d,s) \<in> set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / 2) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2475
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2476
    let ?M = "(\<lambda>(c,t). - ?Nt a t / ?N c) ` ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2477
    have fM: "finite ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2478
      by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2479
    from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2480
    have "\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2481
        a \<le> - ?Nt x t / ?N c \<and> a \<ge> - ?Nt x s / ?N d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2482
      by blast
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2483
    then obtain c t d s
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2484
      where ctU: "(c, t) \<in> ?U"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  2485
        and dsU: "(d, s) \<in> ?U"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2486
        and xs1: "a \<le> - ?Nt x s / ?N d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2487
        and tx1: "a \<ge> - ?Nt x t / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2488
      by blast
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2489
    from uset_l[OF lp] ctU dsU tmbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2490
    have xs: "a \<le> - ?Nt a s / ?N d" and tx: "a \<ge> - ?Nt a t / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2491
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2492
    from ctU have Mne: "?M \<noteq> {}" by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2493
    then have Une: "?U \<noteq> {}" by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2494
    let ?l = "Min ?M"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2495
    let ?u = "Max ?M"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2496
    have linM: "?l \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2497
      using fM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2498
    have uinM: "?u \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2499
      using fM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2500
    have ctM: "- ?Nt a t / ?N c \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2501
      using ctU by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2502
    have dsM: "- ?Nt a s / ?N d \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2503
      using dsU by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2504
    have lM: "\<forall>t\<in> ?M. ?l \<le> t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2505
      using Mne fM by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2506
    have Mu: "\<forall>t\<in> ?M. t \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2507
      using Mne fM by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2508
    have "?l \<le> - ?Nt a t / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2509
      using ctM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2510
    then have lx: "?l \<le> a"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2511
      using tx by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2512
    have "- ?Nt a s / ?N d \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2513
      using dsM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2514
    then have xu: "a \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2515
      using xs by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2516
    from finite_set_intervals2[where P="\<lambda>x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2517
    consider u where "u \<in> ?M" "?I u p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2518
      | t1 t2 where "t1 \<in> ?M" "t2\<in> ?M" "\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" "t1 < a" "a < t2" "?I a p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2519
      by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2520
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2521
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2522
      case 1
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2523
      then have "\<exists>(nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2524
        by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2525
      then obtain tu nu where tuU: "(nu, tu) \<in> ?U" and tuu: "u = - ?Nt a tu / ?N nu"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2526
        by blast
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2527
      have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / 2) p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2528
        using \<open>?I u p\<close> tuu by simp
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2529
      with tuU show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2530
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2531
      case 2
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2532
      have "\<exists>(t1n, t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2533
        using \<open>t1 \<in> ?M\<close> by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2534
      then obtain t1u t1n where t1uU: "(t1n, t1u) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2535
        and t1u: "t1 = - ?Nt a t1u / ?N t1n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2536
        by blast
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2537
      have "\<exists>(t2n, t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2538
        using \<open>t2 \<in> ?M\<close> by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2539
      then obtain t2u t2n where t2uU: "(t2n, t2u) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2540
        and t2u: "t2 = - ?Nt a t2u / ?N t2n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2541
        by blast
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2542
      have "t1 < t2"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2543
        using \<open>t1 < a\<close> \<open>a < t2\<close> by simp
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  2544
      let ?u = "(t1 + t2) / 2"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2545
      have "t1 < ?u"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2546
        using less_half_sum [OF \<open>t1 < t2\<close>] by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2547
      have "?u < t2"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2548
        using gt_half_sum [OF \<open>t1 < t2\<close>] by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2549
      have "?I ?u p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2550
        using lp \<open>\<forall>y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M\<close> \<open>t1 < a\<close> \<open>a < t2\<close> \<open>?I a p\<close> \<open>t1 < ?u\<close> \<open>?u < t2\<close>
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2551
        by (rule lin_dense)
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2552
      with t1uU t2uU t1u t2u show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2553
    qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2554
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2555
  then obtain l n s  m
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2556
    where lnU: "(n, l) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2557
      and smU:"(m,s) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2558
      and pu: "?I ((- ?Nt a l / ?N n + - ?Nt a s / ?N m) / 2) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2559
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2560
  from lnU smU uset_l[OF lp] have nbl: "tmbound0 l" and nbs: "tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2561
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2562
  from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2563
    tmbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2564
  have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / 2) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2565
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2566
  with lnU smU show ?thesis by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2567
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2568
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2569
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2570
section \<open>The Ferrante - Rackoff Theorem\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2571
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2572
theorem fr_eq:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2573
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2574
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2575
    (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2576
     Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2577
     (\<exists>(n, t) \<in> set (uset p). \<exists>(m, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2578
       Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) / 2)#bs) p))"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2579
  (is "(\<exists>x. ?I x p) \<longleftrightarrow> ?M \<or> ?P \<or> ?F" is "?E \<longleftrightarrow> ?D")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2580
proof
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2581
  show ?D if ?E
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2582
  proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2583
    consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2584
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2585
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2586
      case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2587
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2588
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2589
      case 2
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2590
      from inf_uset[OF lp this] have ?F
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2591
        using \<open>?E\<close> by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2592
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2593
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2594
  qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2595
  show ?E if ?D
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2596
  proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2597
    from that consider ?M | ?P | ?F by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2598
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2599
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2600
      case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2601
      from minusinf_ex[OF lp this] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2602
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2603
      case 2
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2604
      from plusinf_ex[OF lp this] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2605
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2606
      case 3
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2607
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2608
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2609
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2610
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2611
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2612
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  2613
section \<open>First implementation : Naive by encoding all case splits locally\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2614
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2615
definition "msubsteq c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2616
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2617
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2618
    in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2619
   (conj (Eq (CP c)) (NEq (CP d)), Eq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2620
   (conj (NEq (CP c)) (Eq (CP d)), Eq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2621
   (conj (Eq (CP c)) (Eq (CP d)), Eq r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2622
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2623
lemma msubsteq_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2624
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2625
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2626
    and s: "tmbound0 s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2627
  shows "bound0 (msubsteq c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2628
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2629
  have th: "\<forall>x \<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2630
    [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2631
      in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2632
     (conj (Eq (CP c)) (NEq (CP d)), Eq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2633
     (conj (NEq (CP c)) (Eq (CP d)), Eq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2634
     (conj (Eq (CP c)) (Eq (CP d)), Eq r)]. bound0 (case_prod conj x)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2635
    using lp by (simp add: Let_def t s)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2636
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2637
    by (simp add: msubsteq_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2638
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2639
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2640
lemma msubsteq:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2641
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2642
  shows "Ifm vs (x#bs) (msubsteq c t d s a r) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2643
    Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / 2)#bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2644
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2645
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2646
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2647
  let ?N = "\<lambda>p. Ipoly vs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2648
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2649
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2650
  let ?t = "?Nt x t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2651
  let ?s = "?Nt x s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2652
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2653
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2654
  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2655
    by simp_all
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2656
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2657
  consider "?c = 0" "?d = 0" | "?c = 0" "?d \<noteq> 0" | "?c \<noteq> 0" "?d = 0" | "?c \<noteq> 0" "?d \<noteq> 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2658
    by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2659
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2660
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2661
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2662
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2663
      by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2664
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2665
    case cd: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2666
    then have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2*?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2667
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2668
    have "?rhs = Ifm vs (-?s / (2*?d) # bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2669
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2670
    also have "\<dots> \<longleftrightarrow> ?a * (-?s / (2*?d)) + ?r = 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2671
      by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2672
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (-?s / (2*?d)) + ?r) = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2673
      using cd(2) mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2674
    also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2 * ?d * ?r= 0"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  2675
      by (simp add: field_simps distrib_left [of "2*?d"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2676
    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2677
      using cd(2) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2678
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2679
      using cd
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  2680
      by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2681
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2682
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2683
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2684
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2685
    have "?rhs = Ifm vs (-?t / (2*?c) # bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2686
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2687
    also have "\<dots> \<longleftrightarrow> ?a * (-?t / (2*?c)) + ?r = 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2688
      by (simp add: r[of "- (?t/ (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2689
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (-?t / (2 * ?c)) + ?r) = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2690
      using cd(1) mult_cancel_left[of "2 * ?c" "(?a * (-?t / (2 * ?c)) + ?r)" 0] by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2691
    also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2 * ?c) / (2 * ?c) + 2 * ?c * ?r= 0"
66809
f6a30d48aab0 replaced recdef were easy to replace
haftmann
parents: 66453
diff changeset
  2692
      by (simp add: field_simps distrib_left [of "2 * ?c"])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2693
    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2 * ?c * ?r = 0"
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2694
      using cd(1) by simp
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2695
    finally show ?thesis using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2696
      by (simp add: r[of "- (?t/ (2 * ?c))"] msubsteq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2697
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2698
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2699
    then have cd2: "?c * ?d * 2 \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2700
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2701
    from add_frac_eq[OF cd, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2702
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  2703
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2704
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2705
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2706
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  2707
      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2708
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2709
      using cd mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2 * ?c * ?d)) + ?r" 0]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2710
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2711
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2 * ?c * ?d * ?r = 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2712
      using nonzero_mult_div_cancel_left [OF cd2] cd
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2713
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2714
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2715
      using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2716
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2717
          msubsteq_def Let_def evaldjf_ex field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2718
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2719
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2720
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2721
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2722
definition "msubstneq c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2723
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2724
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2725
    in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2726
   (conj (Eq (CP c)) (NEq (CP d)), NEq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2727
   (conj (NEq (CP c)) (Eq (CP d)), NEq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2728
   (conj (Eq (CP c)) (Eq (CP d)), NEq r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2729
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2730
lemma msubstneq_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2731
  assumes lp: "islin (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2732
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2733
    and s: "tmbound0 s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2734
  shows "bound0 (msubstneq c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2735
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2736
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2737
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2738
     in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2739
    (conj (Eq (CP c)) (NEq (CP d)), NEq (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2740
    (conj (NEq (CP c)) (Eq (CP d)), NEq (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2741
    (conj (Eq (CP c)) (Eq (CP d)), NEq r)]. bound0 (case_prod conj x)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2742
    using lp by (simp add: Let_def t s)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2743
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2744
    by (simp add: msubstneq_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2745
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2746
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2747
lemma msubstneq:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2748
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2749
  shows "Ifm vs (x#bs) (msubstneq c t d s a r) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2750
    Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2751
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2752
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2753
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2754
  let ?N = "\<lambda>p. Ipoly vs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2755
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2756
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2757
  let ?t = "?Nt x t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2758
  let ?s = "?Nt x s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2759
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2760
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2761
  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2762
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2763
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2764
  consider "?c = 0" "?d = 0" | "?c = 0" "?d \<noteq> 0" | "?c \<noteq> 0" "?d = 0" | "?c \<noteq> 0" "?d \<noteq> 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2765
    by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2766
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2767
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2768
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2769
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2770
      by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2771
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2772
    case cd: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2773
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2 * ?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2774
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2775
    have "?rhs = Ifm vs (-?s / (2*?d) # bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2776
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2777
    also have "\<dots> \<longleftrightarrow> ?a * (-?s / (2*?d)) + ?r \<noteq> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2778
      by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2779
    also have "\<dots> \<longleftrightarrow> 2*?d * (?a * (-?s / (2*?d)) + ?r) \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2780
      using cd(2) mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  2781
    also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2*?d*?r\<noteq> 0"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 48562
diff changeset
  2782
      by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2783
    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2784
      using cd(2) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2785
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2786
      using cd
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  2787
      by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2788
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2789
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2790
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2*?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2791
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2792
    have "?rhs = Ifm vs (-?t / (2*?c) # bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2793
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2794
    also have "\<dots> \<longleftrightarrow> ?a * (-?t / (2*?c)) + ?r \<noteq> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2795
      by (simp add: r[of "- (?t/ (2 * ?c))"])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2796
    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2797
      using cd(1) mult_cancel_left[of "2*?c" "(?a * (-?t / (2*?c)) + ?r)" 0] by simp
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  2798
    also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2*?c) / (2*?c) + 2*?c*?r \<noteq> 0"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 48562
diff changeset
  2799
      by (simp add: field_simps distrib_left[of "2*?c"] del: distrib_left)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2800
    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2801
      using cd(1) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2802
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2803
      using cd by (simp add: r[of "- (?t/ (2*?c))"] msubstneq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2804
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2805
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2806
    then have cd2: "?c * ?d * 2 \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2807
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2808
    from add_frac_eq[OF cd, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2809
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  2810
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2811
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2812
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2813
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  2814
      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2815
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2816
      using cd mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2817
      by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2818
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2819
      using nonzero_mult_div_cancel_left[OF cd2] cd
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2820
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2821
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2822
      using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2823
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2824
          msubstneq_def Let_def evaldjf_ex field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2825
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2826
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2827
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2828
definition "msubstlt c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2829
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2830
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2831
    in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2832
   (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2833
    in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2834
   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)), Lt (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2835
   (conj (lt (CP c)) (Eq (CP d)), Lt (Sub (Mul a t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2836
   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)), Lt (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2837
   (conj (lt (CP d)) (Eq (CP c)), Lt (Sub (Mul a s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2838
   (conj (Eq (CP c)) (Eq (CP d)), Lt r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2839
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2840
lemma msubstlt_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2841
  assumes lp: "islin (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2842
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2843
    and s: "tmbound0 s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2844
  shows "bound0 (msubstlt c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2845
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2846
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2847
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2848
    in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2849
   (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2850
    in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2851
   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)), Lt (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2852
   (conj (lt (CP c)) (Eq (CP d)), Lt (Sub (Mul a t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2853
   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)), Lt (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2854
   (conj (lt (CP d)) (Eq (CP c)), Lt (Sub (Mul a s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2855
   (conj (Eq (CP c)) (Eq (CP d)), Lt r)]. bound0 (case_prod conj x)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2856
    using lp by (simp add: Let_def t s lt_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2857
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2858
    by (simp add: msubstlt_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2859
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2860
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2861
lemma msubstlt:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2862
  assumes nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2863
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2864
    and lp: "islin (Lt (CNP 0 a r))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2865
  shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2866
    Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2867
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2868
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2869
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2870
  let ?N = "\<lambda>p. Ipoly vs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2871
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2872
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2873
  let ?t = "?Nt x t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2874
  let ?s = "?Nt x s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2875
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2876
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2877
  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2878
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2879
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2880
  consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2881
    | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2882
    by atomize_elim auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2883
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2884
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2885
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2886
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2887
      using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2888
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2889
    case cd: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2890
    then have cd2: "2 * ?c * ?d > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2891
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2892
    from cd have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2893
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2894
    from cd2 have cd2': "\<not> 2 * ?c * ?d < 0" by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2895
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2896
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  2897
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2898
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2899
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2900
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  2901
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  2902
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2903
      using cd2 cd2'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2904
        mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2905
      by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2906
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2907
      using nonzero_mult_div_cancel_left[of "2*?c*?d"] c d
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2908
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2909
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2910
      using cd c d nc nd cd2'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2911
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2912
          msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2913
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2914
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2915
    then have cd2: "2 * ?c * ?d < 0"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2916
      by (simp add: mult_less_0_iff field_simps)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2917
    from cd have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2918
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2919
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2920
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s) / (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  2921
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2922
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2 * ?c * ?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2923
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2924
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)) + ?r < 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  2925
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2926
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)) + ?r) > 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2927
      using cd2 order_less_not_sym[OF cd2]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2928
        mult_less_cancel_left_disj[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2929
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2930
    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2931
      using nonzero_mult_div_cancel_left[of "2 * ?c * ?d"] c d
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2932
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2933
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2934
      using cd c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2935
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2936
          msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2937
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2938
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2939
    from cd(1) have c'': "2 * ?c > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2940
      by (simp add: zero_less_mult_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2941
    from cd(1) have c': "2 * ?c \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2942
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2943
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2944
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2945
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2946
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2947
    also have "\<dots> \<longleftrightarrow> ?a * (- ?t / (2 * ?c))+ ?r < 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2948
      by (simp add: r[of "- (?t / (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2949
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2950
      using cd(1) mult_less_cancel_left_disj[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c''
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2951
        order_less_not_sym[OF c'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2952
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2953
    also have "\<dots> \<longleftrightarrow> - ?a * ?t + 2 * ?c * ?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2954
      using nonzero_mult_div_cancel_left[OF c'] \<open>?c > 0\<close>
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2955
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2956
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2957
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2958
      by (simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2959
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2960
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2961
    case cd: 5
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2962
    from cd(1) have c': "2 * ?c \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2963
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2964
    from cd(1) have c'': "2 * ?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2965
      by (simp add: mult_less_0_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2966
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2967
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2968
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2*?c) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2969
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2970
    also have "\<dots> \<longleftrightarrow> ?a * (- ?t / (2*?c))+ ?r < 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2971
      by (simp add: r[of "- (?t / (2*?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2972
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) > 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2973
      using cd(1) order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2974
        mult_less_cancel_left_disj[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2975
      by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2976
    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  2977
      using nonzero_mult_div_cancel_left[OF c'] cd(1) order_less_not_sym[OF c'']
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2978
          less_imp_neq[OF c''] c''
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2979
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2980
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2981
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2982
      by (simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2983
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2984
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2985
    case cd: 6
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2986
    from cd(2) have d'': "2 * ?d > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2987
      by (simp add: zero_less_mult_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2988
    from cd(2) have d': "2 * ?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2989
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2990
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2991
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2992
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2993
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2994
    also have "\<dots> \<longleftrightarrow> ?a * (- ?s / (2 * ?d))+ ?r < 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2995
      by (simp add: r[of "- (?s / (2 * ?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2996
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d))+ ?r) < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2997
      using cd(2) mult_less_cancel_left_disj[of "2 * ?d" "?a * (- ?s / (2 * ?d))+ ?r" 0] d' d''
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2998
        order_less_not_sym[OF d'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2999
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3000
    also have "\<dots> \<longleftrightarrow> - ?a * ?s+  2 * ?d * ?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3001
      using nonzero_mult_div_cancel_left[OF d'] cd(2)
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3002
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3003
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3004
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3005
      by (simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3006
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  3007
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3008
    case cd: 7
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3009
    from cd(2) have d': "2 * ?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3010
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3011
    from cd(2) have d'': "2 * ?d < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3012
      by (simp add: mult_less_0_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3013
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3014
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3015
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3016
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3017
    also have "\<dots> \<longleftrightarrow> ?a * (- ?s / (2 * ?d)) + ?r < 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3018
      by (simp add: r[of "- (?s / (2 * ?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3019
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) > 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3020
      using cd(2) order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3021
        mult_less_cancel_left_disj[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3022
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3023
    also have "\<dots> \<longleftrightarrow> ?a * ?s -  2 * ?d * ?r < 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3024
      using nonzero_mult_div_cancel_left[OF d'] cd(2) order_less_not_sym[OF d'']
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3025
          less_imp_neq[OF d''] d''
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3026
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3027
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3028
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3029
      by (simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3030
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  3031
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3032
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3033
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3034
definition "msubstle c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  3035
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3036
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3037
     in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3038
    (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3039
     in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3040
    (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)), Le (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3041
    (conj (lt (CP c)) (Eq (CP d)), Le (Sub (Mul a t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3042
    (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)), Le (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3043
    (conj (lt (CP d)) (Eq (CP c)), Le (Sub (Mul a s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3044
    (conj (Eq (CP c)) (Eq (CP d)), Le r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3045
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3046
lemma msubstle_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3047
  assumes lp: "islin (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3048
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3049
    and s: "tmbound0 s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3050
  shows "bound0 (msubstle c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3051
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3052
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3053
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3054
     in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3055
    (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3056
     in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul ((2)\<^sub>p *\<^sub>p cd) r)))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3057
    (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3058
    (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul ((2)\<^sub>p *\<^sub>p c) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3059
    (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3060
    (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul ((2)\<^sub>p *\<^sub>p d) r))),
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  3061
    (conj (Eq (CP c)) (Eq (CP d)) , Le r)]. bound0 (case_prod conj x)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3062
    using lp by (simp add: Let_def t s lt_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3063
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3064
    by (simp add: msubstle_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3065
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3066
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3067
lemma msubstle:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3068
  assumes nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3069
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3070
    and lp: "islin (Le (CNP 0 a r))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3071
  shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3072
    Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3073
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3074
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3075
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3076
  let ?N = "\<lambda>p. Ipoly vs p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3077
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3078
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3079
  let ?t = "?Nt x t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3080
  let ?s = "?Nt x s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3081
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3082
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3083
  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3084
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3085
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3086
  have "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3087
    by auto
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3088
  then consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3089
    | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3090
    by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3091
  then show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3092
  proof cases
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3093
    case 1
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3094
    with nc nd show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3095
      by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3096
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3097
    case dc: 2
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3098
    from dc have dc': "2 * ?c * ?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3099
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3100
    then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3101
      by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3102
    from dc' have dc'': "\<not> 2 * ?c * ?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3103
      by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3104
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3105
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  3106
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3107
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3108
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3109
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<le> 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3110
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3111
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3112
      using dc' dc''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3113
        mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3114
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3115
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3116
      using nonzero_mult_div_cancel_left[of "2*?c*?d"] c d
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3117
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3118
    finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3119
      using dc c d  nc nd dc'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3120
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3121
          Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3122
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3123
    case dc: 3
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3124
    from dc have dc': "2 * ?c * ?d < 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3125
      by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3126
    then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3127
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3128
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3129
    have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)"
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  3130
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3131
    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3132
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3133
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<le> 0"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3134
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3135
    also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<ge> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3136
      using dc' order_less_not_sym[OF dc']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3137
        mult_le_cancel_left[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3138
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3139
    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3140
      using nonzero_mult_div_cancel_left[of "2 * ?c * ?d"] c d
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3141
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3142
    finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3143
      using dc c d  nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3144
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3145
          Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3146
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3147
    case 4
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3148
    have c: "?c > 0" and d: "?d = 0" by fact+
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3149
    from c have c'': "2 * ?c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3150
      by (simp add: zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3151
    from c have c': "2 * ?c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3152
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3153
    from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3154
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3155
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3156
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3157
    also have "\<dots> \<longleftrightarrow> ?a * (- ?t / (2 * ?c))+ ?r \<le> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3158
      by (simp add: r[of "- (?t / (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3159
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3160
      using c mult_le_cancel_left[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3161
        order_less_not_sym[OF c'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3162
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3163
    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3164
      using nonzero_mult_div_cancel_left[OF c'] c
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3165
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3166
    finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3167
      using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3168
      by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3169
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3170
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3171
    case 5
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3172
    have c: "?c < 0" and d: "?d = 0" by fact+
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3173
    then have c': "2 * ?c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3174
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3175
    from c have c'': "2 * ?c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3176
      by (simp add: mult_less_0_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3177
    from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3178
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3179
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3180
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3181
    also have "\<dots> \<longleftrightarrow> ?a * (- ?t / (2*?c))+ ?r \<le> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3182
      by (simp add: r[of "- (?t / (2*?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3183
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \<ge> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3184
      using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3185
        mult_le_cancel_left[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3186
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3187
    also have "\<dots> \<longleftrightarrow> ?a * ?t - 2 * ?c * ?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3188
      using nonzero_mult_div_cancel_left[OF c'] c order_less_not_sym[OF c'']
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3189
          less_imp_neq[OF c''] c''
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3190
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3191
    finally show ?thesis using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3192
      by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3193
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3194
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3195
    case 6
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3196
    have c: "?c = 0" and d: "?d > 0" by fact+
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3197
    from d have d'': "2 * ?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3198
      by (simp add: zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3199
    from d have d': "2 * ?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3200
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3201
    from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3202
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3203
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3204
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3205
    also have "\<dots> \<longleftrightarrow> ?a * (- ?s / (2 * ?d))+ ?r \<le> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3206
      by (simp add: r[of "- (?s / (2*?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3207
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3208
      using d mult_le_cancel_left[of "2 * ?d" "?a* (- ?s / (2*?d))+ ?r" 0] d' d''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3209
        order_less_not_sym[OF d'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3210
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3211
    also have "\<dots> \<longleftrightarrow> - ?a * ?s + 2 * ?d * ?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3212
      using nonzero_mult_div_cancel_left[OF d'] d
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3213
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3214
    finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3215
      using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3216
      by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3217
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3218
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3219
    case 7
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3220
    have c: "?c = 0" and d: "?d < 0" by fact+
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3221
    then have d': "2 * ?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3222
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3223
    from d have d'': "2 * ?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3224
      by (simp add: mult_less_0_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3225
    from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3226
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3227
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2*?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3228
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3229
    also have "\<dots> \<longleftrightarrow> ?a* (- ?s / (2*?d))+ ?r \<le> 0"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3230
      by (simp add: r[of "- (?s / (2*?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3231
    also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) \<ge> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3232
      using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3233
        mult_le_cancel_left[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3234
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3235
    also have "\<dots> \<longleftrightarrow> ?a * ?s -  2 * ?d * ?r \<le> 0"
64240
eabf80376aab more standardized names
haftmann
parents: 61586
diff changeset
  3236
      using nonzero_mult_div_cancel_left[OF d'] d order_less_not_sym[OF d'']
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3237
        less_imp_neq[OF d''] d''
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3238
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3239
    finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3240
      using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3241
      by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3242
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3243
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3244
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3245
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3246
fun msubst :: "fm \<Rightarrow> (poly \<times> tm) \<times> (poly \<times> tm) \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3247
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3248
    "msubst (And p q) ((c, t), (d, s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3249
  | "msubst (Or p q) ((c, t), (d, s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3250
  | "msubst (Eq (CNP 0 a r)) ((c, t), (d, s)) = msubsteq c t d s a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3251
  | "msubst (NEq (CNP 0 a r)) ((c, t), (d, s)) = msubstneq c t d s a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3252
  | "msubst (Lt (CNP 0 a r)) ((c, t), (d, s)) = msubstlt c t d s a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3253
  | "msubst (Le (CNP 0 a r)) ((c, t), (d, s)) = msubstle c t d s a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3254
  | "msubst p ((c, t), (d, s)) = p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3255
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3256
lemma msubst_I:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3257
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3258
    and nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3259
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3260
  shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3261
    Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3262
  using lp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3263
  by (induct p rule: islin.induct)
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3264
    (auto simp add: tmbound0_I
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3265
      [where b = "(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3266
        and b' = x and bs = bs and vs = vs]
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3267
      msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3268
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3269
lemma msubst_nb:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3270
  assumes "islin p"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3271
    and "tmbound0 t"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3272
    and "tmbound0 s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3273
  shows "bound0 (msubst p ((c,t),(d,s)))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3274
  using assms
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3275
  by (induct p rule: islin.induct) (auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3276
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3277
lemma fr_eq_msubst:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3278
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3279
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3280
    (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3281
     Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3282
     (\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3283
      Ifm vs (x#bs) (msubst p ((c, t), (d, s)))))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3284
  (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3285
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3286
  from uset_l[OF lp] have *: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3287
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3288
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3289
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3290
    assume ctU: "(c, t) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3291
      and dsU: "(d,s) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3292
      and pts: "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3293
    from *[rule_format, OF ctU] *[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3294
      by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3295
    from msubst_I[OF lp norm, of vs x bs t s] pts
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3296
    have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3297
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3298
  moreover
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3299
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3300
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3301
    assume ctU: "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3302
      and dsU: "(d,s) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3303
      and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3304
    from *[rule_format, OF ctU] *[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3305
      by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3306
    from msubst_I[OF lp norm, of vs x bs t s] pts
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3307
    have "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p" ..
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3308
  }
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3309
  ultimately have **: "(\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3310
      Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p) \<longleftrightarrow> ?F"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3311
    by blast
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3312
  from fr_eq[OF lp, of vs bs x, simplified **] show ?thesis .
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3313
qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3314
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3315
lemma simpfm_lin:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  3316
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3317
  shows "qfree p \<Longrightarrow> islin (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3318
  by (induct p rule: simpfm.induct) (auto simp add: conj_lin disj_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3319
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3320
definition "ferrack p \<equiv>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3321
  let
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3322
    q = simpfm p;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3323
    mp = minusinf q;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3324
    pp = plusinf q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3325
  in
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3326
    if (mp = T \<or> pp = T) then T
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3327
    else
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3328
     (let U = alluopairs (remdups (uset  q))
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3329
      in decr0 (disj mp (disj pp (evaldjf (simpfm o (msubst q)) U ))))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3330
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3331
lemma ferrack:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3332
  assumes qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3333
  shows "qfree (ferrack p) \<and> Ifm vs bs (ferrack p) = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3334
  (is "_ \<and> ?rhs = ?lhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3335
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3336
  let ?I = "\<lambda>x p. Ifm vs (x#bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3337
  let ?N = "\<lambda>t. Ipoly vs t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3338
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3339
  let ?q = "simpfm p"
41823
81d64ec48427 eliminated remdps in favor of List.remdups
krauss
parents: 41822
diff changeset
  3340
  let ?U = "remdups(uset ?q)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3341
  let ?Up = "alluopairs ?U"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3342
  let ?mp = "minusinf ?q"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3343
  let ?pp = "plusinf ?q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3344
  fix x
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3345
  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3346
  from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3347
  from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3348
  from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3349
  from uset_l[OF lq] have U_l: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3350
    by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3351
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3352
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3353
    assume ctU: "(c, t) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3354
      and dsU: "(d,s) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3355
    from U_l ctU dsU have norm: "isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3356
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3357
    from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3358
    have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3359
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3360
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3361
  then have th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3362
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3363
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3364
    fix x
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3365
    assume xUp: "x \<in> set ?Up"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3366
    then obtain c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3367
      where ctU: "(c, t) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3368
        and dsU: "(d,s) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3369
        and x: "x = ((c, t),(d, s))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3370
      using alluopairs_set1[of ?U] by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3371
    from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3372
    have nbs: "tmbound0 t" "tmbound0 s" by simp_all
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3373
    from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3374
    have "bound0 ((simpfm o (msubst (simpfm p))) x)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3375
      using x by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3376
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3377
  with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3378
  have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3379
  with mp_nb pp_nb
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3380
  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3381
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3382
  from decr0_qf[OF th1] have thqf: "qfree (ferrack p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3383
    by (simp add: ferrack_def Let_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3384
  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3385
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3386
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3387
      (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3388
    using fr_eq_msubst[OF lq, of vs bs x] by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3389
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3390
      (\<exists>(x, y) \<in> set ?Up. ?I ((simpfm \<circ> msubst ?q) (x, y)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3391
    using alluopairs_bex[OF th0] by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3392
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (evaldjf (simpfm \<circ> msubst ?q) ?Up)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3393
    by (simp add: evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3394
  also have "\<dots> \<longleftrightarrow> ?I (disj ?mp (disj ?pp (evaldjf (simpfm \<circ> msubst ?q) ?Up)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3395
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3396
  also have "\<dots> \<longleftrightarrow> ?rhs"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3397
    using decr0[OF th1, of vs x bs]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3398
    apply (simp add: ferrack_def Let_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3399
    apply (cases "?mp = T \<or> ?pp = T")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3400
     apply auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3401
    done
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3402
  finally show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3403
    using thqf by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3404
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3405
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3406
definition "frpar p = simpfm (qelim p ferrack)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3407
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3408
lemma frpar: "qfree (frpar p) \<and> (Ifm vs bs (frpar p) \<longleftrightarrow> Ifm vs bs p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3409
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3410
  from ferrack
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3411
  have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack p) \<and> Ifm vs bs (ferrack p) = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3412
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3413
  from qelim[OF th, of p bs] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3414
    unfolding frpar_def by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3415
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3416
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3417
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3418
section \<open>Second implementation: case splits not local\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3419
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3420
lemma fr_eq2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3421
  assumes lp: "islin p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3422
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3423
   (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3424
    Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3425
    Ifm vs (0#bs) p \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3426
    (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3427
      Ipoly vs n \<noteq> 0 \<and> Ifm vs ((- Itm vs (x#bs) t /  (Ipoly vs n * 2))#bs) p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3428
      (\<exists>(n, t) \<in> set (uset p). \<exists>(m, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3429
        Ipoly vs n \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3430
        Ipoly vs m \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3431
        Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /2)#bs) p))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3432
  (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?Z \<or> ?U \<or> ?F)" is "?E = ?D")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3433
proof
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3434
  assume px: "\<exists>x. ?I x p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3435
  consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3436
  then show ?D
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3437
  proof cases
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3438
    case 1
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3439
    then show ?thesis by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3440
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3441
    case 2
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3442
    have nmi: "\<not> ?M" and npi: "\<not> ?P" by fact+
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3443
    from inf_uset[OF lp nmi npi, OF px]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3444
    obtain c t d s where ct:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3445
      "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3446
      "(d, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3447
      "?I ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1)) p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3448
      by auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3449
    let ?c = "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3450
    let ?d = "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3451
    let ?s = "Itm vs (x # bs) s"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3452
    let ?t = "Itm vs (x # bs) t"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3453
    have eq2: "\<And>(x::'a). x + x = 2 * x"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3454
      by (simp add: field_simps)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3455
    consider "?c = 0" "?d = 0" | "?c = 0" "?d \<noteq> 0" | "?c \<noteq> 0" "?d = 0" | "?c \<noteq> 0" "?d \<noteq> 0"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3456
      by auto
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3457
    then show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3458
    proof cases
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3459
      case 1
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3460
      with ct show ?thesis by simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3461
    next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3462
      case 2
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3463
      with ct show ?thesis by auto
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3464
    next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3465
      case 3
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3466
      with ct show ?thesis by auto
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3467
    next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3468
      case z: 4
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3469
      from z have ?F
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3470
        using ct
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3471
        apply -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3472
        apply (rule bexI[where x = "(c,t)"])
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3473
         apply simp_all
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3474
        apply (rule bexI[where x = "(d,s)"])
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3475
         apply simp_all
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3476
        done
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3477
      then show ?thesis by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3478
    qed
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3479
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3480
next
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3481
  assume ?D
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3482
  then consider ?M | ?P | ?Z | ?U | ?F by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3483
  then show ?E
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3484
  proof cases
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3485
    case 1
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3486
    show ?thesis by (rule minusinf_ex[OF lp \<open>?M\<close>])
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3487
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3488
    case 2
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3489
    show ?thesis by (rule plusinf_ex[OF lp \<open>?P\<close>])
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3490
  qed blast+
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3491
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3492
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3493
definition "msubsteq2 c t a b = Eq (Add (Mul a t) (Mul c b))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3494
definition "msubstltpos c t a b = Lt (Add (Mul a t) (Mul c b))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3495
definition "msubstlepos c t a b = Le (Add (Mul a t) (Mul c b))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3496
definition "msubstltneg c t a b = Lt (Neg (Add (Mul a t) (Mul c b)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3497
definition "msubstleneg c t a b = Le (Neg (Add (Mul a t) (Mul c b)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3498
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3499
lemma msubsteq2:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3500
  assumes nz: "Ipoly vs c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3501
    and l: "islin (Eq (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3502
  shows "Ifm vs (x#bs) (msubsteq2 c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3503
    Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Eq (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3504
  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>", symmetric]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3505
  by (simp add: msubsteq2_def field_simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3506
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3507
lemma msubstltpos:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3508
  assumes nz: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3509
    and l: "islin (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3510
  shows "Ifm vs (x#bs) (msubstltpos c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3511
    Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3512
  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>", symmetric]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3513
  by (simp add: msubstltpos_def field_simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3514
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3515
lemma msubstlepos:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3516
  assumes nz: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3517
    and l: "islin (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3518
  shows "Ifm vs (x#bs) (msubstlepos c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3519
    Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3520
  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>", symmetric]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3521
  by (simp add: msubstlepos_def field_simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3522
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3523
lemma msubstltneg:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3524
  assumes nz: "Ipoly vs c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3525
    and l: "islin (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3526
  shows "Ifm vs (x#bs) (msubstltneg c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3527
    Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3528
  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>", symmetric]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3529
  by (simp add: msubstltneg_def field_simps del: minus_add_distrib)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3530
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3531
lemma msubstleneg:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3532
  assumes nz: "Ipoly vs c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3533
    and l: "islin (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3534
  shows "Ifm vs (x#bs) (msubstleneg c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3535
    Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3536
  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>", symmetric]
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3537
  by (simp add: msubstleneg_def field_simps del: minus_add_distrib)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3538
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3539
fun msubstpos :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3540
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3541
    "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3542
  | "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3543
  | "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3544
  | "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3545
  | "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3546
  | "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3547
  | "msubstpos p c t = p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3548
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3549
lemma msubstpos_I:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3550
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3551
    and pos: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3552
  shows "Ifm vs (x#bs) (msubstpos p c t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3553
    Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3554
  using lp pos
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3555
  by (induct p rule: islin.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3556
    (auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3557
      tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3558
      bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3559
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3560
fun msubstneg :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3561
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3562
    "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3563
  | "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3564
  | "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3565
  | "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3566
  | "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3567
  | "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3568
  | "msubstneg p c t = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3569
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3570
lemma msubstneg_I:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3571
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3572
    and pos: "Ipoly vs c < 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3573
  shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3574
  using lp pos
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3575
  by (induct p rule: islin.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3576
    (auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3577
      tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3578
      bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3579
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3580
definition "msubst2 p c t =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3581
  disj (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t)))
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3582
    (conj (lt (CP c)) (simpfm (msubstneg p c t)))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3583
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3584
lemma msubst2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3585
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3586
    and nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3587
    and nz: "Ipoly vs c \<noteq> 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3588
  shows "Ifm vs (x#bs) (msubst2 p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3589
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3590
  let ?c = "Ipoly vs c"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3591
  from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p c))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3592
    by (simp_all add: polyneg_norm)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3593
  from nz consider "?c < 0" | "?c > 0" by arith
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3594
  then show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3595
  proof cases
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3596
    case c: 1
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3597
    from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3598
    show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3599
      by (auto simp add: msubst2_def)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3600
  next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3601
    case c: 2
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3602
    from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3603
    show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3604
      by (auto simp add: msubst2_def)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3605
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3606
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3607
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3608
lemma msubsteq2_nb: "tmbound0 t \<Longrightarrow> islin (Eq (CNP 0 a r)) \<Longrightarrow> bound0 (msubsteq2 c t a r)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3609
  by (simp add: msubsteq2_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3610
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3611
lemma msubstltpos_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltpos c t a r)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3612
  by (simp add: msubstltpos_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3613
lemma msubstltneg_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltneg c t a r)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3614
  by (simp add: msubstltneg_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3615
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3616
lemma msubstlepos_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstlepos c t a r)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3617
  by (simp add: msubstlepos_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3618
lemma msubstleneg_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstleneg c t a r)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3619
  by (simp add: msubstleneg_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3620
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3621
lemma msubstpos_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3622
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3623
    and tnb: "tmbound0 t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3624
  shows "bound0 (msubstpos p c t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3625
  using lp tnb
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3626
  by (induct p c t rule: msubstpos.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3627
    (auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3628
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3629
lemma msubstneg_nb:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  3630
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3631
    and lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3632
    and tnb: "tmbound0 t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3633
  shows "bound0 (msubstneg p c t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3634
  using lp tnb
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3635
  by (induct p c t rule: msubstneg.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3636
    (auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3637
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3638
lemma msubst2_nb:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  3639
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3640
    and lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3641
    and tnb: "tmbound0 t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3642
  shows "bound0 (msubst2 p c t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3643
  using lp tnb
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3644
  by (simp add: msubst2_def msubstneg_nb msubstpos_nb lt_nb simpfm_bound0)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3645
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3646
lemma mult_minus2_left: "-2 * x = - (2 * x)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3647
  for x :: "'a::comm_ring_1"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3648
  by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3649
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3650
lemma mult_minus2_right: "x * -2 = - (x * 2)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3651
  for x :: "'a::comm_ring_1"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3652
  by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3653
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3654
lemma islin_qf: "islin p \<Longrightarrow> qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3655
  by (induct p rule: islin.induct) (auto simp add: bound0_qf)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3656
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3657
lemma fr_eq_msubst2:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3658
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3659
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3660
    ((Ifm vs (x#bs) (minusinf p)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3661
     (Ifm vs (x#bs) (plusinf p)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3662
     Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3663
     (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3664
      Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3665
      (\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3666
        Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3667
  (is "(\<exists>x. ?I x p) = (?M \<or> ?P \<or> ?Pz \<or> ?PU \<or> ?F)" is "?E = ?D")
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3668
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3669
  from uset_l[OF lp] have *: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3670
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3671
  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3672
  have n2: "isnpoly (C (-2,1))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3673
    by (simp add: isnpoly_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3674
  note eq0 = subst0[OF islin_qf[OF lp], of vs x bs "CP 0\<^sub>p", simplified]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3675
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3676
  have eq1: "(\<exists>(n, t) \<in> set (uset p). ?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3677
    (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3678
      \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3679
      Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3680
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3681
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3682
      fix n t
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3683
      assume H: "(n, t) \<in> set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3684
      from H(1) * have "isnpoly n"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3685
        by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3686
      then have nn: "isnpoly (n *\<^sub>p (C (-2,1)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3687
        by (simp_all add: polymul_norm n2)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3688
      have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3689
        by (simp add: polyneg_norm nn)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3690
      then have nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3691
        using H(2) nn' nn
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3692
        by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3693
      from msubst2[OF lp nn nn2(1), of x bs t]
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3694
      have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3695
        using H(2) nn2 by (simp add: mult_minus2_right)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3696
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3697
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3698
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3699
      fix n t
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3700
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3701
        "(n, t) \<in> set (uset p)" "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3702
        "Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * 2) # bs) p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3703
      from H(1) * have "isnpoly n"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3704
        by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3705
      then have nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3706
        using H(2) by (simp_all add: polymul_norm n2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3707
      from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  3708
        using H(2,3) by (simp add: mult_minus2_right)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3709
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3710
    ultimately show ?thesis by blast
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3711
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3712
  have eq2: "(\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3713
      Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3714
    (\<exists>(n, t)\<in>set (uset p). \<exists>(m, s)\<in>set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3715
      \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3716
      \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3717
      Ifm vs ((- Itm vs (x # bs) t / \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3718
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3719
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3720
      fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3721
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3722
        "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3723
        "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3724
      from H(1,2) * have "isnpoly c" "isnpoly d"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3725
        by blast+
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3726
      then have nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3727
        by (simp_all add: polymul_norm n2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3728
      have stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3729
          "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3730
          "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3731
        by (simp_all add: polyneg_norm nn)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3732
      have nn': "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3733
        using H(3)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3734
        by (auto simp add: msubst2_def lt[OF stupid(1)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3735
            lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3736
      from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3737
      have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3738
          Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  3739
        by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult.commute)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3740
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3741
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3742
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3743
      fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3744
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3745
        "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3746
        "(d, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3747
        "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3748
        "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3749
        "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3750
      from H(1,2) * have "isnpoly c" "isnpoly d"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3751
        by blast+
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3752
      then have nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3753
        using H(3,4) by (simp_all add: polymul_norm n2)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3754
      from msubst2[OF lp nn, of x bs ] H(3,4,5)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3755
      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
  3756
        by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult.commute)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3757
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3758
    ultimately show ?thesis by blast
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3759
  qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3760
  from fr_eq2[OF lp, of vs bs x] show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3761
    unfolding eq0 eq1 eq2 by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3762
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3763
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3764
definition "ferrack2 p \<equiv>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3765
  let
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3766
    q = simpfm p;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3767
    mp = minusinf q;
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3768
    pp = plusinf q
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3769
  in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3770
    if (mp = T \<or> pp = T) then T
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3771
    else
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3772
     (let U = remdups (uset  q)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3773
      in
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3774
        decr0
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3775
          (list_disj
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3776
            [mp,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3777
             pp,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3778
             simpfm (subst0 (CP 0\<^sub>p) q),
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3779
             evaldjf (\<lambda>(c, t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U,
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3780
             evaldjf (\<lambda>((b, a),(d, c)).
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3781
              msubst2 q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs U)]))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3782
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3783
definition "frpar2 p = simpfm (qelim (prep p) ferrack2)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3784
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3785
lemma ferrack2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3786
  assumes qf: "qfree p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3787
  shows "qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3788
    (is "_ \<and> (?rhs = ?lhs)")
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3789
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3790
  let ?J = "\<lambda>x p. Ifm vs (x#bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3791
  let ?N = "\<lambda>t. Ipoly vs t"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3792
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3793
  let ?q = "simpfm p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3794
  let ?qz = "subst0 (CP 0\<^sub>p) ?q"
41823
81d64ec48427 eliminated remdps in favor of List.remdups
krauss
parents: 41822
diff changeset
  3795
  let ?U = "remdups(uset ?q)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3796
  let ?Up = "alluopairs ?U"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3797
  let ?mp = "minusinf ?q"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3798
  let ?pp = "plusinf ?q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3799
  fix x
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3800
  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3801
  from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3802
  from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3803
  from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3804
  from uset_l[OF lq]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3805
  have U_l: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3806
    by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3807
  have bnd0: "\<forall>x \<in> set ?U. bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3808
  proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3809
    have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3810
      if "(c, t) \<in> set ?U" for c t
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3811
    proof -
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3812
      from U_l that have "tmbound0 t" by blast
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3813
      from msubst2_nb[OF lq this] show ?thesis by simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3814
    qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3815
    then show ?thesis by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3816
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3817
  have bnd1: "\<forall>x \<in> set ?Up. bound0 ((\<lambda>((b, a), (d, c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3818
    msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) x)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3819
  proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3820
    have "bound0 ((\<lambda>((b, a),(d, c)).
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3821
      msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3822
      if  "((b,a),(d,c)) \<in> set ?Up" for b a d c
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3823
    proof -
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3824
      from U_l alluopairs_set1[of ?U] that have this: "tmbound0 (Add (Mul d a) (Mul b c))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3825
        by auto
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3826
      from msubst2_nb[OF lq this] show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3827
        by simp
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3828
    qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3829
    then show ?thesis by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3830
  qed
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3831
  have stupid: "bound0 F" by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3832
  let ?R =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3833
    "list_disj
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3834
     [?mp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3835
      ?pp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3836
      simpfm (subst0 (CP 0\<^sub>p) ?q),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3837
      evaldjf (\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) ?U,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3838
      evaldjf (\<lambda>((b,a),(d,c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3839
        msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs ?U)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3840
  from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3841
    evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3842
  have nb: "bound0 ?R"
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3843
    by (simp add: list_disj_def simpfm_bound0)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3844
  let ?s = "\<lambda>((b, a),(d, c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3845
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3846
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3847
    fix b a d c
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3848
    assume baU: "(b,a) \<in> set ?U" and dcU: "(d,c) \<in> set ?U"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3849
    from U_l baU dcU have norm: "isnpoly b" "isnpoly d" "isnpoly (C (-2, 1))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3850
      by auto (simp add: isnpoly_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3851
    have norm2: "isnpoly (C (-2, 1) *\<^sub>p b*\<^sub>p d)" "isnpoly (C (-2, 1) *\<^sub>p d*\<^sub>p b)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3852
      using norm by (simp_all add: polymul_norm)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3853
    have stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3854
        "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p b *\<^sub>p d))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3855
        "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p d *\<^sub>p b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3856
        "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p b *\<^sub>p d)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3857
        "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p d*\<^sub>p b)))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3858
      by (simp_all add: polyneg_norm norm2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3859
    have "?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3860
        ?I (msubst2 ?q (C (-2, 1) *\<^sub>p d*\<^sub>p b) (Add (Mul b c) (Mul d a)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3861
      (is "?lhs \<longleftrightarrow> ?rhs")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3862
    proof
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3863
      assume H: ?lhs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3864
      then have z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3865
        by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3866
            mult_less_0_iff zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3867
      from msubst2[OF lq norm2(1) z(1), of x bs] msubst2[OF lq norm2(2) z(2), of x bs] H
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3868
      show ?rhs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3869
        by (simp add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3870
    next
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3871
      assume H: ?rhs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3872
      then have z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3873
        by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3874
            mult_less_0_iff zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3875
      from msubst2[OF lq norm2(1) z(1), of x bs] msubst2[OF lq norm2(2) z(2), of x bs] H
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3876
      show ?lhs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3877
        by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3878
    qed
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3879
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3880
  then have th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3881
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3882
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3883
  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3884
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3885
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3886
      (\<exists>(n, t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3887
      (\<exists>(b, a) \<in> set ?U. \<exists>(d, c) \<in> set ?U.
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3888
        ?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3889
    using fr_eq_msubst2[OF lq, of vs bs x] by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3890
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3891
      (\<exists>(n, t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3892
      (\<exists>x \<in> set ?U. \<exists>y \<in>set ?U. ?I (?s (x, y)))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3893
    by (simp add: split_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3894
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3895
      (\<exists>(n, t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3896
      (\<exists>(x, y) \<in> set ?Up. ?I (?s (x, y)))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3897
    using alluopairs_bex[OF th0] by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3898
  also have "\<dots> \<longleftrightarrow> ?I ?R"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3899
    by (simp add: list_disj_def evaldjf_ex split_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3900
  also have "\<dots> \<longleftrightarrow> ?rhs"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3901
    unfolding ferrack2_def
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3902
    apply (cases "?mp = T")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3903
     apply (simp add: list_disj_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3904
    apply (cases "?pp = T")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3905
     apply (simp add: list_disj_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3906
    apply (simp_all add: Let_def decr0[OF nb])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3907
    done
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3908
  finally show ?thesis using decr0_qf[OF nb]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3909
    by (simp add: ferrack2_def Let_def)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3910
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3911
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3912
lemma frpar2: "qfree (frpar2 p) \<and> (Ifm vs bs (frpar2 p) \<longleftrightarrow> Ifm vs bs p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3913
proof -
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3914
  from ferrack2
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3915
  have this: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3916
    by blast
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3917
  from qelim[OF this, of "prep p" bs] show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3918
    unfolding frpar2_def by (auto simp add: prep)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3919
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3920
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3921
oracle frpar_oracle =
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66809
diff changeset
  3922
\<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3923
let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3924
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3925
fun binopT T = T --> T --> T;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3926
fun relT T = T --> T --> @{typ bool};
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3927
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58889
diff changeset
  3928
val mk_C = @{code C} o apply2 @{code int_of_integer};
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  3929
val mk_poly_Bound = @{code poly.Bound} o @{code nat_of_integer};
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  3930
val mk_Bound = @{code Bound} o @{code nat_of_integer};
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  3931
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3932
val dest_num = snd o HOLogic.dest_number;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3933
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3934
fun try_dest_num t = SOME ((snd o HOLogic.dest_number) t)
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3935
  handle TERM _ => NONE;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3936
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3937
fun dest_nat (t as Const (@{const_name Suc}, _)) = HOLogic.dest_nat t
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3938
  | dest_nat t = dest_num t;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3939
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3940
fun the_index ts t =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3941
  let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3942
    val k = find_index (fn t' => t aconv t') ts;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3943
  in if k < 0 then raise General.Subscript else k end;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3944
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3945
fun num_of_term ps (Const (@{const_name Groups.uminus}, _) $ t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3946
      @{code poly.Neg} (num_of_term ps t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3947
  | num_of_term ps (Const (@{const_name Groups.plus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3948
      @{code poly.Add} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3949
  | num_of_term ps (Const (@{const_name Groups.minus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3950
      @{code poly.Sub} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3951
  | num_of_term ps (Const (@{const_name Groups.times}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3952
      @{code poly.Mul} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3953
  | num_of_term ps (Const (@{const_name Power.power}, _) $ a $ n) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3954
      @{code poly.Pw} (num_of_term ps a, @{code nat_of_integer} (dest_nat n))
60352
d46de31a50c4 separate class for division operator, with particular syntax added in more specific classes
haftmann
parents: 59867
diff changeset
  3955
  | num_of_term ps (Const (@{const_name Rings.divide}, _) $ a $ b) =
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3956
      mk_C (dest_num a, dest_num b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3957
  | num_of_term ps t =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3958
      (case try_dest_num t of
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3959
        SOME k => mk_C (k, 1)
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  3960
      | NONE => mk_poly_Bound (the_index ps t));
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3961
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3962
fun tm_of_term fs ps (Const(@{const_name Groups.uminus}, _) $ t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3963
      @{code Neg} (tm_of_term fs ps t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3964
  | tm_of_term fs ps (Const(@{const_name Groups.plus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3965
      @{code Add} (tm_of_term fs ps a, tm_of_term fs ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3966
  | tm_of_term fs ps (Const(@{const_name Groups.minus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3967
      @{code Sub} (tm_of_term fs ps a, tm_of_term fs ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3968
  | tm_of_term fs ps (Const(@{const_name Groups.times}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3969
      @{code Mul} (num_of_term ps a, tm_of_term fs ps b)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3970
  | tm_of_term fs ps t = (@{code CP} (num_of_term ps t)
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  3971
      handle TERM _ => mk_Bound (the_index fs t)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3972
        | General.Subscript => mk_Bound (the_index fs t));
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3973
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3974
fun fm_of_term fs ps @{term True} = @{code T}
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3975
  | fm_of_term fs ps @{term False} = @{code F}
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3976
  | fm_of_term fs ps (Const (@{const_name Not}, _) $ p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3977
      @{code NOT} (fm_of_term fs ps p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3978
  | fm_of_term fs ps (Const (@{const_name HOL.conj}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3979
      @{code And} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3980
  | fm_of_term fs ps (Const (@{const_name HOL.disj}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3981
      @{code Or} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3982
  | fm_of_term fs ps (Const (@{const_name HOL.implies}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3983
      @{code Imp} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3984
  | fm_of_term fs ps (@{term HOL.iff} $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3985
      @{code Iff} (fm_of_term fs ps p, fm_of_term fs ps q)
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3986
  | fm_of_term fs ps (Const (@{const_name HOL.eq}, T) $ p $ q) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3987
      @{code Eq} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q))
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3988
  | fm_of_term fs ps (Const (@{const_name Orderings.less}, _) $ p $ q) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3989
      @{code Lt} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q))
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3990
  | fm_of_term fs ps (Const (@{const_name Orderings.less_eq}, _) $ p $ q) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3991
      @{code Le} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q))
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3992
  | fm_of_term fs ps (Const (@{const_name Ex}, _) $ Abs (abs as (_, xT, _))) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3993
      let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3994
        val (xn', p') = Syntax_Trans.variant_abs abs;  (* FIXME !? *)
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3995
      in @{code E} (fm_of_term (Free (xn', xT) :: fs) ps p') end
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3996
  | fm_of_term fs ps (Const (@{const_name All},_) $ Abs (abs as (_, xT, _))) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3997
      let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3998
        val (xn', p') = Syntax_Trans.variant_abs abs;  (* FIXME !? *)
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3999
      in @{code A} (fm_of_term (Free (xn', xT) :: fs) ps p') end
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4000
  | fm_of_term fs ps _ = error "fm_of_term";
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4001
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4002
fun term_of_num T ps (@{code poly.C} (a, b)) =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4003
      let
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58889
diff changeset
  4004
        val (c, d) = apply2 (@{code integer_of_int}) (a, b)
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4005
      in
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4006
        (if d = 1 then HOLogic.mk_number T c
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4007
        else if d = 0 then Const (@{const_name Groups.zero}, T)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4008
        else
60352
d46de31a50c4 separate class for division operator, with particular syntax added in more specific classes
haftmann
parents: 59867
diff changeset
  4009
          Const (@{const_name Rings.divide}, binopT T) $
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4010
            HOLogic.mk_number T c $ HOLogic.mk_number T d)
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4011
      end
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4012
  | term_of_num T ps (@{code poly.Bound} i) = nth ps (@{code integer_of_nat} i)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4013
  | term_of_num T ps (@{code poly.Add} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4014
      Const (@{const_name Groups.plus}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4015
  | term_of_num T ps (@{code poly.Mul} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4016
      Const (@{const_name Groups.times}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4017
  | term_of_num T ps (@{code poly.Sub} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4018
      Const (@{const_name Groups.minus}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4019
  | term_of_num T ps (@{code poly.Neg} a) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4020
      Const (@{const_name Groups.uminus}, T --> T) $ term_of_num T ps a
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4021
  | term_of_num T ps (@{code poly.Pw} (a, n)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4022
      Const (@{const_name Power.power}, T --> @{typ nat} --> T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4023
        term_of_num T ps a $ HOLogic.mk_number HOLogic.natT (@{code integer_of_nat} n)
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4024
  | term_of_num T ps (@{code poly.CN} (c, n, p)) =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4025
      term_of_num T ps (@{code poly.Add} (c, @{code poly.Mul} (@{code poly.Bound} n, p)));
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4026
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4027
fun term_of_tm T fs ps (@{code CP} p) = term_of_num T ps p
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4028
  | term_of_tm T fs ps (@{code Bound} i) = nth fs (@{code integer_of_nat} i)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4029
  | term_of_tm T fs ps (@{code Add} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4030
      Const (@{const_name Groups.plus}, binopT T) $ term_of_tm T fs ps a $ term_of_tm T fs ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4031
  | term_of_tm T fs ps (@{code Mul} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4032
      Const (@{const_name Groups.times}, binopT T) $ term_of_num T ps a $ term_of_tm T fs ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4033
  | term_of_tm T fs ps (@{code Sub} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4034
      Const (@{const_name Groups.minus}, binopT T) $ term_of_tm T fs ps a $ term_of_tm T fs ps b
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4035
  | term_of_tm T fs ps (@{code Neg} a) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4036
      Const (@{const_name Groups.uminus}, T --> T) $ term_of_tm T fs ps a
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4037
  | term_of_tm T fs ps (@{code CNP} (n, c, p)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4038
      term_of_tm T fs ps (@{code Add} (@{code Mul} (c, @{code Bound} n), p));
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4039
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4040
fun term_of_fm T fs ps @{code T} = @{term True}
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4041
  | term_of_fm T fs ps @{code F} = @{term False}
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4042
  | term_of_fm T fs ps (@{code NOT} p) = @{term Not} $ term_of_fm T fs ps p
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4043
  | term_of_fm T fs ps (@{code And} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4044
      @{term HOL.conj} $ term_of_fm T fs ps p $ term_of_fm T fs ps q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4045
  | term_of_fm T fs ps (@{code Or} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4046
      @{term HOL.disj} $ term_of_fm T fs ps p $ term_of_fm T fs ps q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4047
  | term_of_fm T fs ps (@{code Imp} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4048
      @{term HOL.implies} $ term_of_fm T fs ps p $ term_of_fm T fs ps q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4049
  | term_of_fm T fs ps (@{code Iff} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4050
      @{term HOL.iff} $ term_of_fm T fs ps p $ term_of_fm T fs ps q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4051
  | term_of_fm T fs ps (@{code Lt} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4052
      Const (@{const_name Orderings.less}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4053
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4054
  | term_of_fm T fs ps (@{code Le} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4055
      Const (@{const_name Orderings.less_eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4056
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4057
  | term_of_fm T fs ps (@{code Eq} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4058
      Const (@{const_name HOL.eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4059
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4060
  | term_of_fm T fs ps (@{code NEq} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4061
      @{term Not} $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4062
        (Const (@{const_name HOL.eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4063
          term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T))
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4064
  | term_of_fm T fs ps _ = error "term_of_fm: quantifiers";
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4065
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4066
fun frpar_procedure alternative T ps fm =
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4067
  let
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4068
    val frpar = if alternative then @{code frpar2} else @{code frpar};
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4069
    val fs = subtract (op aconv) (map Free (Term.add_frees fm [])) ps;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4070
    val eval = term_of_fm T fs ps o frpar o fm_of_term fs ps;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4071
    val t = HOLogic.dest_Trueprop fm;
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4072
  in HOLogic.mk_Trueprop (HOLogic.mk_eq (t, eval t)) end;
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4073
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4074
in
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4075
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4076
  fn (ctxt, alternative, ty, ps, ct) =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  4077
    Thm.cterm_of ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  4078
      (frpar_procedure alternative ty ps (Thm.term_of ct))
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4079
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4080
end
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4081
\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4082
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4083
ML \<open>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4084
structure Parametric_Ferrante_Rackoff =
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4085
struct
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4086
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4087
fun tactic ctxt alternative T ps =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54230
diff changeset
  4088
  Object_Logic.full_atomize_tac ctxt
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4089
  THEN' CSUBGOAL (fn (g, i) =>
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4090
    let
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4091
      val th = frpar_oracle (ctxt, alternative, T, ps, g);
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60567
diff changeset
  4092
    in resolve_tac ctxt [th RS iffD2] i end);
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4093
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4094
fun method alternative =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4095
  let
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4096
    fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4097
    val parsN = "pars";
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4098
    val typN = "type";
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4099
    val any_keyword = keyword parsN || keyword typN;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4100
    val terms = Scan.repeat (Scan.unless any_keyword Args.term);
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4101
    val typ = Scan.unless any_keyword Args.typ;
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4102
  in
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4103
    (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4104
      (fn (T, ps) => fn ctxt => SIMPLE_METHOD' (tactic ctxt alternative T ps))
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4105
  end;
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4106
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4107
end;
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4108
\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4109
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4110
method_setup frpar = \<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4111
  Parametric_Ferrante_Rackoff.method false
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4112
\<close> "parametric QE for linear Arithmetic over fields"
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4113
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4114
method_setup frpar2 = \<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4115
  Parametric_Ferrante_Rackoff.method true
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4116
\<close> "parametric QE for linear Arithmetic over fields, Version 2"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4117
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4118
lemma "\<exists>(x::'a::linordered_field). y \<noteq> -1 \<longrightarrow> (y + 1) * x < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4119
  apply (frpar type: 'a pars: y)
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  4120
  apply (simp add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4121
  apply (rule spec[where x=y])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4122
  apply (frpar type: 'a pars: "z::'a")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4123
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4124
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4125
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4126
lemma "\<exists>(x::'a::linordered_field). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4127
  apply (frpar2 type: 'a pars: y)
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4128
  apply (simp add: field_simps)
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4129
  apply (rule spec[where x=y])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4130
  apply (frpar2 type: 'a pars: "z::'a")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4131
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4132
  done
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4133
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4134
text \<open>Collins/Jones Problem\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4135
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4136
lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4137
proof -
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4138
  have "(\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  4139
by (simp add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4140
have "?rhs"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4141
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4142
  apply (frpar type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}")
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  4143
  apply (simp add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4144
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4145
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4146
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4147
lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4148
apply (frpar type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4149
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4150
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4151
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4152
text \<open>Collins/Jones Problem\<close>
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4153
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4154
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4155
lemma "\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4156
proof -
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4157
  have "(\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
36348
89c54f51f55a dropped group_simps, ring_simps, field_eq_simps; classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
haftmann
parents: 35625
diff changeset
  4158
by (simp add: field_simps)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4159
have "?rhs"
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4160
  apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4161
  apply simp
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4162
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4163
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4164
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4165
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4166
lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4167
apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4168
apply (simp add: field_simps linorder_neq_iff[symmetric])
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4169
apply ferrack
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4170
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4171
*)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  4172
end