src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
author haftmann
Tue, 13 Oct 2015 09:21:15 +0200
changeset 61424 c3658c18b7bc
parent 60754 02924903a6fd
child 61586 5197a2ecb658
permissions -rw-r--r--
prod_case as canonical name for product type eliminator
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
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    12
  "~~/src/HOL/Library/Code_Target_Numeral"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    13
  "~~/src/HOL/Library/Old_Recdef"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    14
begin
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    15
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
    16
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
    17
58310
91ea607a34d8 updated news
blanchet
parents: 58259
diff changeset
    18
datatype 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
    19
  | 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
    20
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    21
text \<open>A size for poly to make inductive proofs simpler.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    22
primrec tmsize :: "tm \<Rightarrow> nat"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    23
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    24
  "tmsize (CP c) = polysize c"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    25
| "tmsize (Bound n) = 1"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    26
| "tmsize (Neg a) = 1 + tmsize a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    27
| "tmsize (Add a b) = 1 + tmsize a + tmsize b"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    28
| "tmsize (Sub a b) = 3 + tmsize a + tmsize b"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    29
| "tmsize (Mul c a) = 1 + polysize c + tmsize a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    30
| "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    31
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    32
text \<open>Semantics of terms tm.\<close>
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    33
primrec Itm :: "'a::{field_char_0,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    34
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    35
  "Itm vs bs (CP c) = (Ipoly vs c)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    36
| "Itm vs bs (Bound n) = bs!n"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    37
| "Itm vs bs (Neg a) = -(Itm vs bs a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    38
| "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    39
| "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    40
| "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    41
| "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
    42
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    43
fun allpolys :: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    44
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    45
  "allpolys P (CP c) = P c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    46
| "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    47
| "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    48
| "allpolys P (Neg p) = allpolys P p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    49
| "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    50
| "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    51
| "allpolys P p = True"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    52
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    53
primrec tmboundslt :: "nat \<Rightarrow> tm \<Rightarrow> bool"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    54
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    55
  "tmboundslt n (CP c) = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    56
| "tmboundslt n (Bound m) = (m < n)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    57
| "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    58
| "tmboundslt n (Neg a) = tmboundslt n a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    59
| "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    60
| "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    61
| "tmboundslt n (Mul i a) = tmboundslt n a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    62
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    63
primrec tmbound0 :: "tm \<Rightarrow> bool"  -- \<open>a tm is INDEPENDENT of Bound 0\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    64
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    65
  "tmbound0 (CP c) = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    66
| "tmbound0 (Bound n) = (n>0)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    67
| "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    68
| "tmbound0 (Neg a) = tmbound0 a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    69
| "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    70
| "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    71
| "tmbound0 (Mul i a) = tmbound0 a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    72
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    73
lemma tmbound0_I:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    74
  assumes nb: "tmbound0 a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    75
  shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    76
  using nb
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    77
  by (induct a rule: tm.induct) auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    78
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
    79
primrec tmbound :: "nat \<Rightarrow> tm \<Rightarrow> bool"  -- \<open>a tm is INDEPENDENT of Bound n\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    80
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
    81
  "tmbound n (CP c) = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    82
| "tmbound n (Bound m) = (n \<noteq> m)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    83
| "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    84
| "tmbound n (Neg a) = tmbound n a"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    85
| "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    86
| "tmbound n (Sub a b) = (tmbound n a \<and> tmbound n b)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    87
| "tmbound n (Mul i a) = tmbound n a"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    88
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    89
lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    90
  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
    91
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    92
lemma tmbound_I:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    93
  assumes bnd: "tmboundslt (length bs) t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    94
    and nb: "tmbound n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    95
    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
    96
  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
    97
  using nb le bnd
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
    98
  by (induct t rule: tm.induct) auto
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
    99
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   100
fun decrtm0 :: "tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   101
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   102
  "decrtm0 (Bound n) = Bound (n - 1)"
41821
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   103
| "decrtm0 (Neg a) = Neg (decrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   104
| "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   105
| "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   106
| "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   107
| "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   108
| "decrtm0 a = a"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   109
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   110
fun incrtm0 :: "tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   111
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   112
  "incrtm0 (Bound n) = Bound (n + 1)"
41821
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   113
| "incrtm0 (Neg a) = Neg (incrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   114
| "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   115
| "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   116
| "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   117
| "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   118
| "incrtm0 a = a"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   119
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   120
lemma decrtm0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   121
  assumes nb: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   122
  shows "Itm vs (x # bs) t = Itm vs bs (decrtm0 t)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   123
  using nb by (induct t rule: decrtm0.induct) simp_all
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   124
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   125
lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   126
  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
   127
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   128
primrec decrtm :: "nat \<Rightarrow> tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   129
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   130
  "decrtm m (CP c) = (CP c)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   131
| "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   132
| "decrtm m (Neg a) = Neg (decrtm m a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   133
| "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   134
| "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   135
| "decrtm m (Mul c a) = Mul c (decrtm m a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   136
| "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
   137
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   138
primrec removen :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   139
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   140
  "removen n [] = []"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   141
| "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
   142
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   143
lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   144
  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
   145
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   146
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
   147
  by (induct xs arbitrary: n) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   148
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   149
lemma removen_length: "length (removen n xs) = (if n \<ge> length xs then length xs else length xs - 1)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   150
  by (induct xs arbitrary: n, auto)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   151
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   152
lemma removen_nth:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   153
  "(removen n xs)!m =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   154
    (if n \<ge> length xs then xs!m
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   155
     else if m < n then xs!m
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   156
     else if m \<le> length xs then xs!(Suc m)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   157
     else []!(m - (length xs - 1)))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   158
proof (induct xs arbitrary: n m)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   159
  case Nil
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   160
  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
   161
next
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   162
  case (Cons x xs)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   163
  let ?l = "length (x # xs)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   164
  consider "n \<ge> ?l" | "n < ?l" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   165
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   166
  proof cases
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   167
    case 1
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   168
    with removen_same[OF this] show ?thesis by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   169
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   170
    case nl: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   171
    consider "m < n" | "m \<ge> n" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   172
    then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   173
    proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   174
      case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   175
      then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   176
        using Cons by (cases m) auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   177
    next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   178
      case 2
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   179
      consider "m \<le> ?l" | "m > ?l" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   180
      then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   181
      proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   182
        case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   183
        then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   184
          using Cons by (cases m) auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   185
      next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   186
        case ml: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   187
        have th: "length (removen n (x # xs)) = length xs"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   188
          using removen_length[where n = n and xs= "x # xs"] nl by simp
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   189
        with ml have "m \<ge> length (removen n (x # xs))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   190
          by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   191
        from th nth_length_exceeds[OF this] have "(removen n (x # xs))!m = [] ! (m - length xs)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   192
           by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   193
        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
   194
          by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   195
        then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
   196
          using ml nl by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   197
      qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   198
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   199
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   200
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   201
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   202
lemma decrtm:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   203
  assumes bnd: "tmboundslt (length bs) t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   204
    and nb: "tmbound m t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   205
    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
   206
  shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41413
diff changeset
   207
  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
   208
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   209
primrec tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   210
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   211
  "tmsubst0 t (CP c) = CP c"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   212
| "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   213
| "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   214
| "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   215
| "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   216
| "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   217
| "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   218
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   219
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
   220
  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
   221
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   222
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
   223
  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
   224
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   225
primrec tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   226
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   227
  "tmsubst n t (CP c) = CP c"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   228
| "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   229
| "tmsubst n t (CNP m c a) =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   230
    (if n = m then Add (Mul c t) (tmsubst n t a) else CNP m c (tmsubst n t a))"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   231
| "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   232
| "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   233
| "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   234
| "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
   235
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   236
lemma tmsubst:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   237
  assumes nb: "tmboundslt (length bs) a"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   238
    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
   239
  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
   240
  using nb nlt
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   241
  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
   242
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   243
lemma tmsubst_nb0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   244
  assumes tnb: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   245
  shows "tmbound0 (tmsubst 0 t a)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   246
  using tnb
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   247
  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
   248
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   249
lemma tmsubst_nb:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   250
  assumes tnb: "tmbound m t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   251
  shows "tmbound m (tmsubst m t a)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   252
  using tnb
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   253
  by (induct a rule: tm.induct) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   254
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   255
lemma incrtm0_tmbound: "tmbound n t \<Longrightarrow> tmbound (Suc n) (incrtm0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   256
  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
   257
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   258
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   259
text \<open>Simplification.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   260
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   261
consts tmadd:: "tm \<times> tm \<Rightarrow> tm"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   262
recdef tmadd "measure (\<lambda>(t,s). size t + size s)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   263
  "tmadd (CNP n1 c1 r1,CNP n2 c2 r2) =
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   264
    (if n1 = n2 then
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   265
      let c = c1 +\<^sub>p c2
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   266
      in if c = 0\<^sub>p then tmadd(r1,r2) else CNP n1 c (tmadd (r1, r2))
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   267
    else if n1 \<le> n2 then (CNP n1 c1 (tmadd (r1,CNP n2 c2 r2)))
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   268
    else (CNP n2 c2 (tmadd (CNP n1 c1 r1, r2))))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   269
  "tmadd (CNP n1 c1 r1, t) = CNP n1 c1 (tmadd (r1, t))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   270
  "tmadd (t, CNP n2 c2 r2) = CNP n2 c2 (tmadd (t, r2))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   271
  "tmadd (CP b1, CP b2) = CP (b1 +\<^sub>p b2)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   272
  "tmadd (a, b) = Add a b"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   273
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   274
lemma tmadd[simp]: "Itm vs bs (tmadd (t, s)) = Itm vs bs (Add t s)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   275
  apply (induct t s rule: tmadd.induct)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   276
  apply (simp_all add: Let_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   277
  apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   278
  apply (case_tac "n1 \<le> n2")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   279
  apply simp_all
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   280
  apply (case_tac "n1 = n2")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   281
  apply (simp_all add: field_simps)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   282
  apply (simp only: distrib_left[symmetric])
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   283
  apply (auto simp del: polyadd simp add: polyadd[symmetric])
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   284
  done
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   285
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   286
lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd (t, s))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   287
  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
   288
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   289
lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd (t, s))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   290
  by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   291
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   292
lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd (t, s))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   293
  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
   294
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   295
lemma tmadd_allpolys_npoly[simp]:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   296
  "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t, s))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   297
  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
   298
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   299
fun tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   300
where
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   301
  "tmmul (CP j) = (\<lambda>i. CP (i *\<^sub>p j))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   302
| "tmmul (CNP n c a) = (\<lambda>i. CNP n (i *\<^sub>p c) (tmmul a i))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   303
| "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
   304
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   305
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
   306
  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
   307
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   308
lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   309
  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
   310
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   311
lemma tmmul_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   312
  by (induct t arbitrary: n rule: tmmul.induct) auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   313
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   314
lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   315
  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
   316
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   317
lemma tmmul_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   318
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   319
  shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   320
  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
   321
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   322
definition tmneg :: "tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   323
  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
   324
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   325
definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   326
  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
   327
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   328
lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   329
  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
   330
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   331
lemma tmneg_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   332
  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
   333
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   334
lemma tmneg_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   335
  using tmneg_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   336
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   337
lemma tmneg_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmneg t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   338
  using tmneg_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   339
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   340
lemma [simp]: "isnpoly (C (-1, 1))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   341
  unfolding isnpoly_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   342
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   343
lemma tmneg_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   344
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   345
  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   346
  unfolding tmneg_def by auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   347
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   348
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
   349
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   350
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   351
lemma tmsub_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmsub t s)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   352
  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
   353
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   354
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
   355
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   356
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   357
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
   358
  using tmsub_def by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   359
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   360
lemma tmsub_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   361
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   362
  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   363
  unfolding tmsub_def by (simp add: isnpoly_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   364
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   365
fun simptm :: "tm \<Rightarrow> tm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   366
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   367
  "simptm (CP j) = CP (polynate j)"
50282
fe4d4bb9f4c2 more robust syntax that survives collapse of \<^isub> and \<^sub>;
wenzelm
parents: 50045
diff changeset
   368
| "simptm (Bound n) = CNP n (1)\<^sub>p (CP 0\<^sub>p)"
41821
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   369
| "simptm (Neg t) = tmneg (simptm t)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   370
| "simptm (Add t s) = tmadd (simptm t,simptm s)"
c118ae98dfbf recdef -> fun
krauss
parents: 41816
diff changeset
   371
| "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   372
| "simptm (Mul i t) =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   373
    (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   374
| "simptm (CNP n c t) =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   375
    (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
   376
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   377
lemma polynate_stupid:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   378
  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
   379
  shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   380
  apply (subst polynate[symmetric])
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   381
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   382
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   383
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   384
lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   385
  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
   386
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   387
lemma simptm_tmbound0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   388
  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
   389
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   390
lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   391
  by (induct t rule: simptm.induct) (auto simp add: Let_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   392
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   393
lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   394
  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
   395
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   396
lemma [simp]: "isnpoly 0\<^sub>p"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   397
  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
   398
  by (simp_all add: isnpoly_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   399
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   400
lemma simptm_allpolys_npoly[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   401
  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
   402
  shows "allpolys isnpoly (simptm p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   403
  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
   404
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   405
declare let_cong[fundef_cong del]
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   406
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   407
fun split0 :: "tm \<Rightarrow> poly \<times> tm"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   408
where
50282
fe4d4bb9f4c2 more robust syntax that survives collapse of \<^isub> and \<^sub>;
wenzelm
parents: 50045
diff changeset
   409
  "split0 (Bound 0) = ((1)\<^sub>p, CP 0\<^sub>p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   410
| "split0 (CNP 0 c t) = (let (c', t') = split0 t in (c +\<^sub>p c', t'))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   411
| "split0 (Neg t) = (let (c, t') = split0 t in (~\<^sub>p c, Neg t'))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   412
| "split0 (CNP n c t) = (let (c', t') = split0 t in (c', CNP n c t'))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   413
| "split0 (Add s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   414
| "split0 (Sub s t) = (let (c1, s') = split0 s; (c2, t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   415
| "split0 (Mul c t) = (let (c', t') = split0 t in (c *\<^sub>p c', Mul c t'))"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   416
| "split0 t = (0\<^sub>p, t)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   417
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   418
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
   419
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   420
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
   421
  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
   422
  apply (rule exI[where x="snd (split0 p)"])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   423
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   424
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   425
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   426
lemma split0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   427
  "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
   428
  apply (induct t rule: split0.induct)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   429
  apply simp
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
   430
  apply (simp add: Let_def split_def field_simps)
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
   431
  apply (simp add: Let_def split_def field_simps)
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
   432
  apply (simp add: Let_def split_def field_simps)
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
   433
  apply (simp add: Let_def split_def field_simps)
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
   434
  apply (simp add: Let_def split_def field_simps)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56479
diff changeset
   435
  apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric])
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
   436
  apply (simp add: Let_def split_def field_simps)
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
   437
  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
   438
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   439
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   440
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
   441
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   442
  fix c' t'
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   443
  assume "split0 t = (c', t')"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   444
  then have "c' = fst (split0 t)" and "t' = snd (split0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   445
    by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   446
  with split0[where t="t" and bs="bs"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   447
  show "Itm vs bs t = Itm vs bs (CNP 0 c' t')"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   448
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   449
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   450
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   451
lemma split0_nb0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   452
  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
   453
  shows "split0 t = (c',t') \<Longrightarrow>  tmbound 0 t'"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   454
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   455
  fix c' t'
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   456
  assume "split0 t = (c', t')"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   457
  then have "c' = fst (split0 t)" and "t' = snd (split0 t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   458
    by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   459
  with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   460
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   461
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   462
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   463
lemma split0_nb0'[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   464
  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
   465
  shows "tmbound0 (snd (split0 t))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   466
  using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   467
  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
   468
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   469
lemma split0_nb:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   470
  assumes nb: "tmbound n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   471
  shows "tmbound n (snd (split0 t))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   472
  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
   473
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   474
lemma split0_blt:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   475
  assumes nb: "tmboundslt n t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   476
  shows "tmboundslt n (snd (split0 t))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   477
  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
   478
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   479
lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   480
  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 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
   483
  by (induct t rule: split0.induct) (auto simp add: Let_def split_def)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   484
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   485
lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst (split0 t)) = 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)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   487
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   488
lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   489
  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
   490
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   491
lemma isnpoly_fst_split0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   492
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   493
  shows "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   494
  by (induct p rule: split0.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   495
    (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
   496
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   497
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   498
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
   499
58310
91ea607a34d8 updated news
blanchet
parents: 58259
diff changeset
   500
datatype fm  =  T| F| Le tm | Lt tm | Eq tm | NEq tm|
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   501
  NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   502
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   503
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   504
text \<open>A size for fm.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   505
fun fmsize :: "fm \<Rightarrow> nat"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   506
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   507
  "fmsize (NOT p) = 1 + fmsize p"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   508
| "fmsize (And p q) = 1 + fmsize p + fmsize q"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   509
| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   510
| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   511
| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   512
| "fmsize (E p) = 1 + fmsize p"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   513
| "fmsize (A p) = 4+ fmsize p"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   514
| "fmsize p = 1"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   515
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   516
lemma fmsize_pos[termination_simp]: "fmsize p > 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   517
  by (induct p rule: fmsize.induct) simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   518
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   519
text \<open>Semantics of formulae (fm).\<close>
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   520
primrec Ifm ::"'a::linordered_field list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   521
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   522
  "Ifm vs bs T = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   523
| "Ifm vs bs F = False"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   524
| "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   525
| "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   526
| "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   527
| "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   528
| "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   529
| "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   530
| "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   531
| "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   532
| "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   533
| "Ifm vs bs (E p) = (\<exists>x. Ifm vs (x#bs) p)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   534
| "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
   535
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   536
fun not:: "fm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   537
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   538
  "not (NOT (NOT p)) = not p"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   539
| "not (NOT p) = p"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   540
| "not T = F"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   541
| "not F = T"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   542
| "not (Lt t) = Le (tmneg t)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   543
| "not (Le t) = Lt (tmneg t)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   544
| "not (Eq t) = NEq t"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   545
| "not (NEq t) = Eq t"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   546
| "not p = NOT p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   547
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   548
lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   549
  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
   550
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   551
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   552
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   553
  "conj p q \<equiv>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   554
    (if p = F \<or> q = F then F
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   555
     else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   556
     else if q = T then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   557
     else if p = q then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   558
     else And p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   559
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   560
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
   561
  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
   562
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   563
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   564
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   565
  "disj p q \<equiv>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   566
    (if (p = T \<or> q = T) then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   567
     else if p = F then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   568
     else if q = F 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 Or p q)"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   571
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   572
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
   573
  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
   574
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   575
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   576
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   577
  "imp p q \<equiv>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   578
    (if p = F \<or> q = T \<or> p = q then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   579
     else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   580
     else if q = F then not p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   581
     else Imp p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   582
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   583
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
   584
  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
   585
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   586
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   587
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   588
  "iff p q \<equiv>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   589
   (if p = q then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   590
    else if p = NOT q \<or> NOT p = q then F
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   591
    else if p = F then not q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   592
    else if q = F then not p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   593
    else if p = T then q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   594
    else if q = T then p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   595
    else Iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   596
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   597
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
   598
  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
   599
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   600
text \<open>Quantifier freeness.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   601
fun qfree:: "fm \<Rightarrow> bool"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   602
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   603
  "qfree (E p) = False"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   604
| "qfree (A p) = False"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   605
| "qfree (NOT p) = qfree p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   606
| "qfree (And p q) = (qfree p \<and> qfree q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   607
| "qfree (Or  p q) = (qfree p \<and> qfree q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   608
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   609
| "qfree (Iff p q) = (qfree p \<and> qfree q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   610
| "qfree p = True"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   611
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   612
text \<open>Boundedness and substitution.\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   613
primrec boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   614
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   615
  "boundslt n T = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   616
| "boundslt n F = True"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   617
| "boundslt n (Lt t) = tmboundslt n t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   618
| "boundslt n (Le t) = tmboundslt n t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   619
| "boundslt n (Eq t) = tmboundslt n t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   620
| "boundslt n (NEq t) = tmboundslt n t"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   621
| "boundslt n (NOT p) = boundslt n p"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   622
| "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   623
| "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   624
| "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   625
| "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   626
| "boundslt n (E p) = boundslt (Suc n) p"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   627
| "boundslt n (A p) = boundslt (Suc n) p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   628
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   629
fun bound0:: "fm \<Rightarrow> bool"  -- \<open>a Formula is independent of Bound 0\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   630
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   631
  "bound0 T = True"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   632
| "bound0 F = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   633
| "bound0 (Lt a) = tmbound0 a"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   634
| "bound0 (Le a) = tmbound0 a"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   635
| "bound0 (Eq a) = tmbound0 a"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   636
| "bound0 (NEq a) = tmbound0 a"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   637
| "bound0 (NOT p) = bound0 p"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   638
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   639
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   640
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   641
| "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   642
| "bound0 p = False"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   643
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   644
lemma bound0_I:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   645
  assumes bp: "bound0 p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   646
  shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   647
  using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   648
  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
   649
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   650
primrec bound:: "nat \<Rightarrow> fm \<Rightarrow> bool"  -- \<open>a Formula is independent of Bound n\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   651
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   652
  "bound m T = True"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   653
| "bound m F = True"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   654
| "bound m (Lt t) = tmbound m t"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   655
| "bound m (Le t) = tmbound m t"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   656
| "bound m (Eq t) = tmbound m t"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   657
| "bound m (NEq t) = tmbound m t"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   658
| "bound m (NOT p) = bound m p"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   659
| "bound m (And p q) = (bound m p \<and> bound m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   660
| "bound m (Or p q) = (bound m p \<and> bound m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   661
| "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   662
| "bound m (Iff p q) = (bound m p \<and> bound m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   663
| "bound m (E p) = bound (Suc m) p"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   664
| "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
   665
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   666
lemma bound_I:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   667
  assumes bnd: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   668
    and nb: "bound n p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   669
    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
   670
  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
   671
  using bnd nb le tmbound_I[where bs=bs and vs = vs]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   672
proof (induct p arbitrary: bs n rule: fm.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   673
  case (E p bs n)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   674
  have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" for y
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   675
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   676
    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
   677
      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
   678
    from E.hyps[OF bnd nb le tmbound_I] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   679
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   680
  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
   681
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   682
  case (A 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 A have bnd: "boundslt (length (y#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   686
      and nb: "bound (Suc n) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   687
      and le: "Suc n \<le> length (y#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   688
      by simp_all
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   689
    from A.hyps[OF bnd nb le tmbound_I] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
   690
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   691
  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
   692
qed auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   693
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   694
fun decr0 :: "fm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   695
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   696
  "decr0 (Lt a) = Lt (decrtm0 a)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   697
| "decr0 (Le a) = Le (decrtm0 a)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   698
| "decr0 (Eq a) = Eq (decrtm0 a)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   699
| "decr0 (NEq a) = NEq (decrtm0 a)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   700
| "decr0 (NOT p) = NOT (decr0 p)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   701
| "decr0 (And p q) = conj (decr0 p) (decr0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   702
| "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   703
| "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   704
| "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   705
| "decr0 p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   706
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   707
lemma decr0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   708
  assumes nb: "bound0 p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   709
  shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   710
  using nb
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   711
  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
   712
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   713
primrec decr :: "nat \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   714
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   715
  "decr m T = T"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   716
| "decr m F = F"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   717
| "decr m (Lt t) = (Lt (decrtm m t))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   718
| "decr m (Le t) = (Le (decrtm m t))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   719
| "decr m (Eq t) = (Eq (decrtm m t))"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   720
| "decr m (NEq t) = (NEq (decrtm m t))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   721
| "decr m (NOT p) = NOT (decr m p)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   722
| "decr m (And p q) = conj (decr m p) (decr m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   723
| "decr m (Or p q) = disj (decr m p) (decr m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   724
| "decr m (Imp p q) = imp (decr m p) (decr m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   725
| "decr m (Iff p q) = iff (decr m p) (decr m q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   726
| "decr m (E p) = E (decr (Suc m) p)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   727
| "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
   728
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   729
lemma decr:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   730
  assumes bnd: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   731
    and nb: "bound m p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   732
    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
   733
  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
   734
  using bnd nb nle
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   735
proof (induct p arbitrary: bs m rule: fm.induct)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   736
  case (E p bs m)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   737
  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
   738
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   739
    from E
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   740
    have bnd: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   741
      and nb: "bound (Suc m) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   742
      and nle: "Suc m < length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   743
      by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   744
    from E(1)[OF bnd nb nle] show ?thesis .
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   745
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   746
  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
   747
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   748
  case (A p bs m)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   749
  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
   750
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   751
    from A
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   752
    have bnd: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   753
      and nb: "bound (Suc m) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   754
      and nle: "Suc m < length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   755
      by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   756
    from A(1)[OF bnd nb nle] show ?thesis .
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   757
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   758
  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
   759
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
   760
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   761
primrec subst0 :: "tm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   762
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   763
  "subst0 t T = T"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   764
| "subst0 t F = F"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   765
| "subst0 t (Lt a) = Lt (tmsubst0 t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   766
| "subst0 t (Le a) = Le (tmsubst0 t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   767
| "subst0 t (Eq a) = Eq (tmsubst0 t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   768
| "subst0 t (NEq a) = NEq (tmsubst0 t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   769
| "subst0 t (NOT p) = NOT (subst0 t p)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   770
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   771
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   772
| "subst0 t (Imp p q) = Imp (subst0 t p)  (subst0 t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   773
| "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   774
| "subst0 t (E p) = E p"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   775
| "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
   776
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   777
lemma subst0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   778
  assumes qf: "qfree p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   779
  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
   780
  using qf tmsubst0[where x="x" and bs="bs" and t="t"]
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   781
  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
   782
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   783
lemma subst0_nb:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   784
  assumes bp: "tmbound0 t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   785
    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
   786
  shows "bound0 (subst0 t p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   787
  using qf tmsubst0_nb[OF bp] bp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   788
  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
   789
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   790
primrec subst:: "nat \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   791
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   792
  "subst n t T = T"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   793
| "subst n t F = F"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   794
| "subst n t (Lt a) = Lt (tmsubst n t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   795
| "subst n t (Le a) = Le (tmsubst n t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   796
| "subst n t (Eq a) = Eq (tmsubst n t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   797
| "subst n t (NEq a) = NEq (tmsubst n t a)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   798
| "subst n t (NOT p) = NOT (subst n t p)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   799
| "subst n t (And p q) = And (subst n t p) (subst n t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   800
| "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   801
| "subst n t (Imp p q) = Imp (subst n t p)  (subst n t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   802
| "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   803
| "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   804
| "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
   805
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   806
lemma subst:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   807
  assumes nb: "boundslt (length bs) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   808
    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
   809
  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
   810
  using nb nlm
39246
9e58f0499f57 modernized primrec
haftmann
parents: 38864
diff changeset
   811
proof (induct p arbitrary: bs n t rule: fm.induct)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   812
  case (E p bs n)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   813
  have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   814
        Ifm vs (x#bs[n:= Itm vs bs t]) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   815
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   816
    from E have bn: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   817
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   818
    from E have nlm: "Suc n \<le> length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   819
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   820
    from E(1)[OF bn nlm]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   821
    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   822
        Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   823
      by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   824
    then show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   825
      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   826
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   827
  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
   828
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   829
  case (A p bs n)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   830
  have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   831
        Ifm vs (x#bs[n:= Itm vs bs t]) p" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   832
  proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   833
    from A have bn: "boundslt (length (x#bs)) p"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   834
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   835
    from A have nlm: "Suc n \<le> length (x#bs)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   836
      by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   837
    from A(1)[OF bn nlm]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   838
    have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   839
        Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   840
      by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   841
    then show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   842
      by (simp add: incrtm0[where x="x" and bs="bs" and t="t"])
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   843
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   844
  then show ?case by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   845
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
   846
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   847
lemma subst_nb:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   848
  assumes tnb: "tmbound m t"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   849
  shows "bound m (subst m t p)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   850
  using tnb tmsubst_nb incrtm0_tmbound
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   851
  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
   852
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   853
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   854
  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
   855
lemma not_bn0[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   856
  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
   857
lemma not_nb[simp]: "bound n p \<Longrightarrow> bound n (not p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   858
  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
   859
lemma not_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n (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
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   862
lemma conj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (conj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   863
  using conj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   864
lemma conj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (conj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   865
  using conj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   866
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
   867
  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
   868
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
   869
  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
   870
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   871
lemma disj_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (disj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   872
  using disj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   873
lemma disj_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (disj p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   874
  using disj_def by auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   875
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
   876
  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
   877
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
   878
  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
   879
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   880
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
   881
  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
   882
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
   883
  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
   884
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
   885
  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
   886
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
   887
  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
   888
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   889
lemma iff_qf[simp]: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   890
  unfolding iff_def by (cases "p = q") auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   891
lemma iff_nb0[simp]: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   892
  using iff_def unfolding iff_def by (cases "p = q") auto
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   893
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
   894
  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
   895
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
   896
  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
   897
lemma decr0_qf: "bound0 p \<Longrightarrow> qfree (decr0 p)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   898
  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
   899
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   900
fun isatom :: "fm \<Rightarrow> bool"  -- \<open>test for atomicity\<close>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   901
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   902
  "isatom T = True"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   903
| "isatom F = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   904
| "isatom (Lt a) = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   905
| "isatom (Le a) = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   906
| "isatom (Eq a) = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   907
| "isatom (NEq a) = True"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   908
| "isatom p = False"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   909
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   910
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   911
  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
   912
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   913
definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   914
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   915
  "djf f p q \<equiv>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   916
    (if q = T then T
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   917
     else if q = F then f p
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   918
     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
   919
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   920
definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   921
  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
   922
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   923
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
   924
  apply (cases "q = T")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   925
  apply (simp add: djf_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   926
  apply (cases "q = F")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   927
  apply (simp add: djf_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   928
  apply (cases "f p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   929
  apply (simp_all add: Let_def djf_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   930
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   931
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   932
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
   933
  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
   934
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   935
lemma evaldjf_bound0:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   936
  assumes nb: "\<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
   937
  shows "bound0 (evaldjf f xs)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   938
  using nb
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   939
  apply (induct xs)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   940
  apply (auto simp add: evaldjf_def djf_def Let_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   941
  apply (case_tac "f a")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   942
  apply auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   943
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   944
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   945
lemma evaldjf_qf:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   946
  assumes nb: "\<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
   947
  shows "qfree (evaldjf f xs)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   948
  using nb
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   949
  apply (induct xs)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   950
  apply (auto simp add: evaldjf_def djf_def Let_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   951
  apply (case_tac "f a")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   952
  apply auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   953
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   954
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   955
fun disjuncts :: "fm \<Rightarrow> fm list"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   956
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   957
  "disjuncts (Or p q) = disjuncts p @ disjuncts q"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   958
| "disjuncts F = []"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
   959
| "disjuncts p = [p]"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   960
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   961
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
   962
  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
   963
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   964
lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q \<in> set (disjuncts p). bound0 q"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   965
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   966
  assume nb: "bound0 p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   967
  then have "list_all bound0 (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   968
    by (induct p rule:disjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   969
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   970
    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
   971
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   972
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   973
lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall>q \<in> set (disjuncts p). qfree q"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
   974
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   975
  assume qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   976
  then have "list_all qfree (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   977
    by (induct p rule: disjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   978
  then show ?thesis 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
   979
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   980
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   981
definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   982
  where "DJ f p \<equiv> evaldjf f (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   983
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   984
lemma DJ:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   985
  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
   986
    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
   987
  shows "Ifm vs bs (DJ f p) = Ifm vs bs (f p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   988
proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
   989
  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
   990
    by (simp add: DJ_def evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   991
  also have "\<dots> = Ifm vs bs (f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   992
    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
   993
  finally show ?thesis .
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   994
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
   995
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   996
lemma DJ_qf:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   997
  assumes fqf: "\<forall>p. qfree p \<longrightarrow> qfree (f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   998
  shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
   999
proof clarify
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1000
  fix  p
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1001
  assume qf: "qfree p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1002
  have th: "DJ f p = evaldjf f (disjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1003
    by (simp add: DJ_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1004
  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
  1005
  with fqf have th':"\<forall>q\<in> set (disjuncts p). qfree (f q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1006
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1007
  from evaldjf_qf[OF th'] th show "qfree (DJ f p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1008
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1009
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1010
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1011
lemma DJ_qe:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1012
  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
  1013
  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
  1014
proof clarify
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1015
  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
  1016
  assume qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1017
  from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1018
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1019
  from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1020
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1021
  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
  1022
    by (simp add: DJ_def evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1023
  also have "\<dots> = (\<exists>q \<in> set(disjuncts p). Ifm vs bs (E q))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1024
    using qe disjuncts_qf[OF qf] by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1025
  also have "\<dots> = Ifm vs bs (E p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1026
    by (induct p rule: disjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1027
  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
  1028
    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
  1029
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1030
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1031
fun conjuncts :: "fm \<Rightarrow> fm list"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1032
where
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1033
  "conjuncts (And p q) = conjuncts p @ conjuncts q"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1034
| "conjuncts T = []"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1035
| "conjuncts p = [p]"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1036
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1037
definition list_conj :: "fm list \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1038
  where "list_conj ps \<equiv> foldr conj ps T"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1039
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1040
definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1041
where
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1042
  "CJNB f p \<equiv>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1043
    (let cjs = conjuncts p;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1044
      (yes, no) = partition bound0 cjs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1045
     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
  1046
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1047
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
  1048
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1049
  assume qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1050
  then have "list_all qfree (conjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1051
    by (induct p rule: conjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1052
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1053
    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
  1054
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1055
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1056
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
  1057
  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
  1058
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1059
lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall>q\<in> set (conjuncts p). bound0 q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1060
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1061
  assume nb: "bound0 p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1062
  then have "list_all bound0 (conjuncts p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1063
    by (induct p rule:conjuncts.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1064
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1065
    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
  1066
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1067
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1068
fun islin :: "fm \<Rightarrow> bool"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1069
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1070
  "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)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1071
| "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)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1072
| "islin (Eq (CNP 0 c s)) = (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
  1073
| "islin (NEq (CNP 0 c s)) = (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
  1074
| "islin (Lt (CNP 0 c s)) = (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
  1075
| "islin (Le (CNP 0 c s)) = (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
  1076
| "islin (NOT p) = False"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1077
| "islin (Imp p q) = False"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1078
| "islin (Iff p q) = False"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1079
| "islin p = bound0 p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1080
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1081
lemma islin_stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1082
  assumes nb: "tmbound0 p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1083
  shows "islin (Lt p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1084
    and "islin (Le p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1085
    and "islin (Eq p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1086
    and "islin (NEq p)"
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1087
  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
  1088
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1089
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
  1090
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
  1091
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
  1092
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
  1093
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1094
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
  1095
  apply (simp add: lt_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1096
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1097
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1098
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1099
  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
  1100
  done
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 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
  1103
  apply (simp add: le_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1104
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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 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
  1111
  apply (simp add: eq_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1112
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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 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
  1119
  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
  1120
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1121
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
  1122
  apply (simp add: lt_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1123
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1124
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1125
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1126
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1127
  apply (rename_tac nat a b, case_tac nat)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1128
  apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1129
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1130
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1131
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
  1132
  apply (simp add: le_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1133
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1134
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1135
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1136
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1137
  apply (rename_tac nat a b, case_tac nat)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1138
  apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1139
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1140
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1141
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
  1142
  apply (simp add: eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1143
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1144
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1145
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1146
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1147
  apply (rename_tac nat a b, case_tac nat)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1148
  apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1149
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1150
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1151
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
  1152
  apply (simp add: neq_def eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1153
  apply (cases p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1154
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1155
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1156
  apply simp_all
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1157
  apply (rename_tac nat a b, case_tac nat)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1158
  apply simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1159
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1160
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1161
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
  1162
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
  1163
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
  1164
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
  1165
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1166
lemma simplt_islin[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1167
  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
  1168
  shows "islin (simplt t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1169
  unfolding simplt_def
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1170
  using split0_nb0'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1171
  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
  1172
      islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1173
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1174
lemma simple_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 (simple t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1177
  unfolding simple_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: 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] le_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1181
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1182
lemma simpeq_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 (simpeq t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1185
  unfolding simpeq_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] eq_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1189
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1190
lemma simpneq_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 (simpneq t)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1193
  unfolding simpneq_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] neq_lin)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1197
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1198
lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1199
  by (cases "split0 s") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1200
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1201
lemma split0_npoly:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1202
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1203
    and n: "allpolys isnpoly t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1204
  shows "isnpoly (fst (split0 t))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1205
    and "allpolys isnpoly (snd (split0 t))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1206
  using n
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1207
  by (induct t rule: split0.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1208
    (auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1209
      polysub_norm really_stupid)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1210
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1211
lemma simplt[simp]: "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1212
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1213
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1214
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1215
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1216
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1217
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1218
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1219
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1220
      using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF n], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1221
      by (simp add: simplt_def Let_def split_def lt)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1222
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1223
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1224
    then show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1225
      using split0[of "simptm t" vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1226
      by (simp add: simplt_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1227
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1228
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1229
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1230
lemma simple[simp]: "Ifm vs bs (simple t) = Ifm vs bs (Le t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1231
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1232
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1233
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1234
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1235
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1236
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1237
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1238
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1239
      using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF n], of vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1240
      by (simp add: simple_def Let_def split_def le)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1241
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1242
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1243
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1244
      using split0[of "simptm t" vs bs]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1245
      by (simp add: simple_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1246
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1247
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1248
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1249
lemma simpeq[simp]: "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1250
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1251
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1252
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1253
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1254
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1255
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1256
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1257
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1258
      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
  1259
      by (simp add: simpeq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1260
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1261
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1262
    then show ?thesis using  split0[of "simptm t" vs bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1263
      by (simp add: simpeq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1264
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1265
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1266
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1267
lemma simpneq[simp]: "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1268
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1269
  have n: "allpolys isnpoly (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1270
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1271
  let ?t = "simptm t"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1272
  show ?thesis
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1273
  proof (cases "fst (split0 ?t) = 0\<^sub>p")
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1274
    case True
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1275
    then show ?thesis
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1276
      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
  1277
      by (simp add: simpneq_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1278
  next
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1279
    case False
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1280
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1281
      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
  1282
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1283
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1284
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1285
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
  1286
  apply (simp add: lt_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1287
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1288
  apply auto
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1289
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1290
  apply auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1291
  done
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 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
  1294
  apply (simp add: le_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1295
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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 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
  1302
  apply (simp add: eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1303
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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 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
  1310
  apply (simp add: neq_def eq_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1311
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
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
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1317
lemma simplt_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1318
  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
  1319
  shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1320
proof (simp add: simplt_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
  1321
  assume nb: "tmbound0 t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1322
  then have nb': "tmbound0 (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1323
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1324
  let ?c = "fst (split0 (simptm t))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1325
  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1326
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1327
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1328
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1329
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1330
    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
  1331
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1332
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1333
  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
  1334
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1335
    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
  1336
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1337
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1338
lemma simple_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1339
  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
  1340
  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
  1341
proof(simp add: simple_def Let_def split_def)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1342
  assume nb: "tmbound0 t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1343
  then have nb': "tmbound0 (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1344
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1345
  let ?c = "fst (split0 (simptm t))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1346
  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1347
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1348
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1349
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1350
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1351
    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
  1352
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1353
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1354
  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
  1355
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1356
    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
  1357
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1358
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1359
lemma simpeq_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1360
  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
  1361
  shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1362
proof (simp add: simpeq_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
  1363
  assume nb: "tmbound0 t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1364
  then have nb': "tmbound0 (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1365
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1366
  let ?c = "fst (split0 (simptm t))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1367
  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1368
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1369
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1370
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1371
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1372
    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
  1373
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1374
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1375
  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
  1376
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1377
    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
  1378
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1379
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1380
lemma simpneq_nb[simp]:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1381
  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
  1382
  shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1383
proof (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
  1384
  assume nb: "tmbound0 t"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1385
  then have nb': "tmbound0 (simptm t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1386
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1387
  let ?c = "fst (split0 (simptm t))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1388
  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1389
  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1390
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1391
  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1392
  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1393
    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
  1394
  from iffD1[OF isnpolyh_unique[OF ths] th]
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1395
  have "fst (split0 (simptm t)) = 0\<^sub>p" .
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1396
  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
  1397
      fst (split0 (simptm t)) = 0\<^sub>p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1398
    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
  1399
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1400
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1401
fun conjs :: "fm \<Rightarrow> fm list"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1402
where
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1403
  "conjs (And p q) = conjs p @ conjs q"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1404
| "conjs T = []"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1405
| "conjs p = [p]"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1406
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1407
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
  1408
  by (induct p rule: conjs.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1409
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1410
definition list_disj :: "fm list \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1411
  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
  1412
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1413
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
  1414
  by (induct ps) (auto simp add: list_conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1415
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1416
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
  1417
  by (induct ps) (auto simp add: list_conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1418
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1419
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
  1420
  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
  1421
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1422
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
  1423
  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
  1424
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1425
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
  1426
  apply (induct p rule: conjs.induct)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1427
  apply (unfold conjs.simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1428
  apply (unfold set_append)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1429
  apply (unfold ball_Un)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1430
  apply (unfold bound.simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1431
  apply auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1432
  done
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1433
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1434
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
  1435
  apply (induct p rule: conjs.induct)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1436
  apply (unfold conjs.simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1437
  apply (unfold set_append)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1438
  apply (unfold ball_Un)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1439
  apply (unfold boundslt.simps)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1440
  apply blast
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1441
  apply simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1442
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1443
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1444
lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1445
  unfolding list_conj_def
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1446
  by (induct ps) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1447
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1448
lemma list_conj_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1449
  assumes bnd: "\<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
  1450
  shows "bound n (list_conj ps)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1451
  using bnd
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1452
  unfolding list_conj_def
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1453
  by (induct ps) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1454
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1455
lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1456
  unfolding list_conj_def by (induct ps) auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1457
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1458
lemma CJNB_qe:
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1459
  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
  1460
  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
  1461
proof clarify
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1462
  fix bs p
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1463
  assume qfp: "qfree p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1464
  let ?cjs = "conjuncts p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1465
  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
  1466
  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
  1467
  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
  1468
  let ?cyes = "list_conj ?yes"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1469
  have part: "partition bound0 ?cjs = (?yes,?no)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1470
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1471
  from partition_P[OF part] have "\<forall>q\<in> set ?yes. bound0 q"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1472
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1473
  then have yes_nb: "bound0 ?cyes"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1474
    by (simp add: list_conj_nb')
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1475
  then have yes_qf: "qfree (decr0 ?cyes)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1476
    by (simp add: decr0_qf)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1477
  from conjuncts_qf[OF qfp] partition_set[OF part]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1478
  have " \<forall>q\<in> set ?no. qfree q"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1479
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1480
  then have no_qf: "qfree ?cno"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1481
    by (simp add: list_conj_qf)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1482
  with qe have cno_qf:"qfree (qe ?cno)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1483
    and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1484
    by blast+
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1485
  from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1486
    by (simp add: CJNB_def Let_def split_def)
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1487
  have "Ifm vs bs p = ((Ifm vs bs ?cyes) \<and> (Ifm vs bs ?cno))" for bs
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1488
  proof -
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1489
    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
  1490
      by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1491
    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
  1492
      using partition_set[OF part] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1493
    finally show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1494
      using list_conj[of vs bs] by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1495
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1496
  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
  1497
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1498
  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
  1499
    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
  1500
  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
  1501
    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
  1502
  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
  1503
    using qe[rule_format, OF no_qf] by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1504
  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
  1505
    by (simp add: Let_def CJNB_def split_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1506
  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
  1507
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1508
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1509
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1510
consts simpfm :: "fm \<Rightarrow> fm"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1511
recdef simpfm "measure fmsize"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1512
  "simpfm (Lt t) = simplt (simptm t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1513
  "simpfm (Le t) = simple (simptm t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1514
  "simpfm (Eq t) = simpeq(simptm t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1515
  "simpfm (NEq t) = simpneq(simptm t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1516
  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1517
  "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1518
  "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1519
  "simpfm (Iff p q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1520
    disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1521
  "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1522
  "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1523
  "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1524
  "simpfm (NOT (Iff p q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1525
    disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1526
  "simpfm (NOT (Eq t)) = simpneq t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1527
  "simpfm (NOT (NEq t)) = simpeq t"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1528
  "simpfm (NOT (Le t)) = simplt (Neg t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1529
  "simpfm (NOT (Lt t)) = simple (Neg t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1530
  "simpfm (NOT (NOT p)) = simpfm p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1531
  "simpfm (NOT T) = F"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1532
  "simpfm (NOT F) = T"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1533
  "simpfm p = p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1534
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1535
lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1536
  by (induct p arbitrary: bs rule: simpfm.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1537
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1538
lemma simpfm_bound0:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1539
  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
  1540
  shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1541
  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
  1542
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1543
lemma lt_qf[simp]: "qfree (lt t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1544
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1545
  apply (auto simp add: lt_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1546
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1547
  apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1548
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1549
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1550
lemma le_qf[simp]: "qfree (le t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1551
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1552
  apply (auto simp add: le_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1553
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1554
  apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1555
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1556
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1557
lemma eq_qf[simp]: "qfree (eq t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1558
  apply (cases t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1559
  apply (auto simp add: eq_def)
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
  1560
  apply (rename_tac poly, case_tac poly)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1561
  apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1562
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1563
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1564
lemma neq_qf[simp]: "qfree (neq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1565
  by (simp add: neq_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1566
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1567
lemma simplt_qf[simp]: "qfree (simplt t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1568
  by (simp add: simplt_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1569
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1570
lemma simple_qf[simp]: "qfree (simple t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1571
  by (simp add: simple_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1572
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1573
lemma simpeq_qf[simp]: "qfree (simpeq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1574
  by (simp add: simpeq_def Let_def split_def)
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1575
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1576
lemma simpneq_qf[simp]: "qfree (simpneq t)"
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1577
  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
  1578
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1579
lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1580
  by (induct p rule: simpfm.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1581
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1582
lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1583
  by (simp add: disj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1584
lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1585
  by (simp add: conj_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1586
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1587
lemma
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1588
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1589
  shows "qfree p \<Longrightarrow> islin (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1590
  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
  1591
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1592
consts prep :: "fm \<Rightarrow> fm"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1593
recdef prep "measure fmsize"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1594
  "prep (E T) = T"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1595
  "prep (E F) = F"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1596
  "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1597
  "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1598
  "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1599
  "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1600
  "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1601
  "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1602
  "prep (E p) = E (prep p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1603
  "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1604
  "prep (A p) = prep (NOT (E (NOT p)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1605
  "prep (NOT (NOT p)) = prep p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1606
  "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1607
  "prep (NOT (A p)) = prep (E (NOT p))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1608
  "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1609
  "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1610
  "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1611
  "prep (NOT p) = not (prep p)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1612
  "prep (Or p q) = disj (prep p) (prep q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1613
  "prep (And p q) = conj (prep p) (prep q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1614
  "prep (Imp p q) = prep (Or (NOT p) q)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1615
  "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1616
  "prep p = p"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1617
  (hints simp add: fmsize_pos)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1618
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1619
lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1620
  by (induct p arbitrary: bs rule: prep.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1621
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1622
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1623
text \<open>Generic quantifier elimination.\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1624
function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1625
where
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1626
  "qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1627
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1628
| "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1629
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1630
| "qelim (Or  p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1631
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1632
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1633
| "qelim p = (\<lambda>y. simpfm p)"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1634
  by pat_completeness simp_all
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1635
termination by (relation "measure fmsize") auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1636
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1637
lemma qelim:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1638
  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
  1639
  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
  1640
  using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1641
  by (induct p rule: qelim.induct) auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1642
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1643
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  1644
subsection \<open>Core Procedure\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1645
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1646
fun minusinf:: "fm \<Rightarrow> fm"  -- \<open>Virtual substitution of -\<infinity>\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1647
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1648
  "minusinf (And p q) = conj (minusinf p) (minusinf q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1649
| "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1650
| "minusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1651
| "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1652
| "minusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1653
| "minusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1654
| "minusinf p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1655
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1656
fun plusinf:: "fm \<Rightarrow> fm"  -- \<open>Virtual substitution of +\<infinity>\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1657
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1658
  "plusinf (And p q) = conj (plusinf p) (plusinf q)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  1659
| "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
41822
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1660
| "plusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1661
| "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1662
| "plusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1663
| "plusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
27afef7d6c37 recdef -> fun
krauss
parents: 41821
diff changeset
  1664
| "plusinf p = p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1665
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1666
lemma minusinf_inf:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1667
  assumes lp: "islin p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1668
  shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> 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
  1669
  using lp
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1670
proof (induct p rule: minusinf.induct)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1671
  case 1
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1672
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1673
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1674
    apply (rule_tac x="min z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1675
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1676
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1677
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1678
  case 2
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1679
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1680
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1681
    apply (rule_tac x="min z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1682
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1683
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1684
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1685
  case (3 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1686
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1687
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1688
  from 3 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1689
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1690
  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
  1691
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1692
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1693
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1694
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1695
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1696
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1697
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1698
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1699
      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1700
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1701
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1702
    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
  1703
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1704
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1705
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1706
        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
  1707
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1708
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1709
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1710
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1711
        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
  1712
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1713
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1714
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1715
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1716
    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
  1717
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1718
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1719
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1720
        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
  1721
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1722
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1723
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1724
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1725
        using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1726
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1727
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1728
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1729
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1730
  case (4 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1731
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1732
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1733
  from 4 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1734
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1735
  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
  1736
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1737
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1738
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1739
  consider "?c = 0" | "?c > 0" | "?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1740
    by arith
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1741
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1742
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1743
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1744
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1745
      using eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1746
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1747
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1748
    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
  1749
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1750
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1751
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1752
        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
  1753
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1754
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1755
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1756
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1757
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1758
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1759
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1760
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1761
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1762
    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
  1763
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1764
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1765
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1766
        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
  1767
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1768
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1769
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1770
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1771
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1772
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1773
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1774
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1775
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1776
  case (5 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1777
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1778
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1779
  from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1780
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1781
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1782
    by (simp add: polyneg_norm)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1783
  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
  1784
  let ?c = "Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1785
  fix y
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1786
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1787
  consider "?c = 0" | "?c > 0" | "?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1788
    by arith
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1789
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1790
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1791
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1792
    then show ?thesis using eqs by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1793
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1794
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1795
    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
  1796
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1797
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1798
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1799
        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
  1800
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1801
      then have "?c * x + ?e < 0" by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1802
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1803
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1804
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1805
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1806
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1807
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1808
    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
  1809
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1810
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1811
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1812
        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
  1813
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1814
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1815
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1816
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1817
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1818
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1819
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1820
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1821
next
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1822
  case (6 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1823
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1824
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1825
  from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1826
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1827
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1828
    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
  1829
  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
  1830
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1831
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1832
  let ?e = "Itm vs (y#bs) e"
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1833
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1834
  then show ?case
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1835
  proof cases
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1836
    case 1
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1837
    then show ?thesis using eqs by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1838
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1839
    case c: 2
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1840
    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
  1841
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1842
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1843
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1844
        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
  1845
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1846
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1847
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1848
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1849
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1850
        by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1851
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1852
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1853
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1854
    case c: 3
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1855
    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
  1856
      if "x < -?e / ?c" for x
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1857
    proof -
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1858
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1859
        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
  1860
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1861
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1862
        by simp
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1863
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1864
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1865
        by auto
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1866
    qed
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1867
    then show ?thesis by auto
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  1868
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1869
qed auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1870
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1871
lemma plusinf_inf:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1872
  assumes lp: "islin p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1873
  shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> 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
  1874
  using lp
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1875
proof (induct p rule: plusinf.induct)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1876
  case 1
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1877
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1878
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1879
    apply (rule_tac x="max z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1880
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1881
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1882
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1883
  case 2
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1884
  then show ?case
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1885
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1886
    apply (rule_tac x="max z za" in exI)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1887
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1888
    done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1889
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1890
  case (3 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1891
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1892
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1893
  from 3 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1894
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1895
  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
  1896
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1897
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1898
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1899
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1900
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1901
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1902
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1903
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1904
      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1905
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1906
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1907
    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
  1908
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1909
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1910
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1911
        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
  1912
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1913
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1914
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1915
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1916
        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
  1917
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1918
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1919
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1920
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1921
    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
  1922
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1923
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1924
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1925
        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
  1926
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1927
      then have "?c * x + ?e < 0" by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1928
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1929
        using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1930
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1931
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1932
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1933
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1934
  case (4 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1935
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1936
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1937
  from 4 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1938
    by simp_all
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1939
  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
  1940
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1941
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1942
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1943
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1944
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1945
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1946
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1947
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1948
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1949
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1950
    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
  1951
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1952
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1953
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1954
        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
  1955
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1956
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1957
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1958
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1959
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1960
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1961
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1962
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1963
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1964
    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
  1965
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1966
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1967
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1968
        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
  1969
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1970
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1971
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1972
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1973
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1974
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1975
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1976
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  1977
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1978
  case (5 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1979
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1980
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1981
  from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1982
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1983
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1984
    by (simp add: polyneg_norm)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1985
  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
  1986
  let ?c = "Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1987
  fix y
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  1988
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1989
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1990
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1991
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1992
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1993
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1994
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  1995
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1996
    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
  1997
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1998
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  1999
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2000
        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
  2001
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2002
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2003
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2004
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2005
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2006
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2007
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2008
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2009
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2010
    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
  2011
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2012
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2013
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2014
        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
  2015
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2016
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2017
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2018
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2019
        using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2020
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2021
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2022
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2023
next
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2024
  case (6 c e)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2025
  then have nbe: "tmbound0 e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2026
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2027
  from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2028
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2029
  then have nc': "allpolys isnpoly (CP (~\<^sub>p c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2030
    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
  2031
  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
  2032
  let ?c = "Ipoly vs c"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2033
  fix y
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2034
  let ?e = "Itm vs (y#bs) e"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2035
  consider "?c = 0" | "?c > 0" | "?c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2036
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2037
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2038
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2039
    then show ?thesis using eqs by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2040
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2041
    case c: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2042
    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
  2043
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2044
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2045
      from that have "?c * x > - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2046
        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
  2047
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2048
      then have "?c * x + ?e > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2049
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2050
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2051
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2052
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2053
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2054
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2055
    case c: 3
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2056
    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
  2057
      if "x > -?e / ?c" for x
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2058
    proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2059
      from that have "?c * x < - ?e"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2060
        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
  2061
        by (simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2062
      then have "?c * x + ?e < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2063
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2064
      then show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2065
        using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2066
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2067
    then show ?thesis by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2068
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2069
qed auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2070
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2071
lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2072
  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
  2073
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2074
lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf 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
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2077
lemma minusinf_ex:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2078
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2079
    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
  2080
  shows "\<exists>x. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2081
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2082
  from bound0_I [OF minusinf_nb[OF lp], where bs ="bs"] ex
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2083
  have th: "\<forall>x. Ifm vs (x#bs) (minusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2084
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2085
  from minusinf_inf[OF lp, where bs="bs"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2086
  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
  2087
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2088
  from th have "Ifm vs ((z - 1)#bs) (minusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2089
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2090
  moreover have "z - 1 < z"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2091
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2092
  ultimately show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2093
    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
  2094
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2095
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2096
lemma plusinf_ex:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2097
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2098
    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
  2099
  shows "\<exists>x. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2100
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2101
  from bound0_I [OF plusinf_nb[OF lp], where bs ="bs"] ex
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2102
  have th: "\<forall>x. Ifm vs (x#bs) (plusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2103
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2104
  from plusinf_inf[OF lp, where bs="bs"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2105
  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
  2106
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2107
  from th have "Ifm vs ((z + 1)#bs) (plusinf p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2108
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2109
  moreover have "z + 1 > z"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2110
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2111
  ultimately show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2112
    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
  2113
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2114
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2115
fun uset :: "fm \<Rightarrow> (poly \<times> tm) list"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2116
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2117
  "uset (And p q) = uset p @ uset q"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2118
| "uset (Or p q) = uset p @ uset q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2119
| "uset (Eq (CNP 0 a e)) = [(a, e)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2120
| "uset (Le (CNP 0 a e)) = [(a, e)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2121
| "uset (Lt (CNP 0 a e)) = [(a, e)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2122
| "uset (NEq (CNP 0 a e)) = [(a, e)]"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2123
| "uset p = []"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2124
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2125
lemma uset_l:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2126
  assumes lp: "islin p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2127
  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
  2128
  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
  2129
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2130
lemma minusinf_uset0:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2131
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2132
    and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2133
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2134
  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
  2135
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2136
  have "\<exists>(c, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2137
      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
  2138
      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
  2139
    using lp nmi ex
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2140
    apply (induct p rule: minusinf.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2141
    apply (auto simp add: eq le lt polyneg_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2142
    apply (auto simp add: linorder_not_less order_le_less)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2143
    done
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2144
  then obtain c s where csU: "(c, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2145
    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
  2146
      (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
  2147
  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
  2148
    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
  2149
    by (auto simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2150
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2151
    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
  2152
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2153
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2154
lemma minusinf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2155
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2156
    and nmi: "\<not> (Ifm vs (a#bs) (minusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2157
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2158
  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
  2159
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2160
  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
  2161
    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
  2162
  from minusinf_uset0[OF lp nmi' ex]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2163
  obtain c s where csU: "(c,s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2164
    and th: "x \<ge> - Itm vs (x#bs) s / Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2165
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2166
  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2167
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2168
  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2169
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2170
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2171
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2172
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2173
lemma plusinf_uset0:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2174
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2175
    and nmi: "\<not> (Ifm vs (x#bs) (plusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2176
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2177
  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
  2178
proof -
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2179
  have "\<exists>(c, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2180
      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
  2181
      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
  2182
    using lp nmi ex
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2183
    apply (induct p rule: minusinf.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2184
    apply (auto simp add: eq le lt polyneg_norm)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2185
    apply (auto simp add: linorder_not_less order_le_less)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2186
    done
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2187
  then obtain c s where csU: "(c, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2188
    and x: "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
  2189
      Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2190
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2191
  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
  2192
    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
  2193
    by (auto simp add: mult.commute)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2194
  then show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2195
    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
  2196
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2197
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2198
lemma plusinf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2199
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2200
    and nmi: "\<not> (Ifm vs (a#bs) (plusinf p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2201
    and ex: "Ifm vs (x#bs) p" (is "?I x p")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2202
  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
  2203
proof -
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2204
  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
  2205
    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
  2206
  from plusinf_uset0[OF lp nmi' ex]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2207
  obtain c s where csU: "(c,s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2208
    and th: "x \<le> - Itm vs (x#bs) s / Ipoly vs c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2209
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2210
  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2211
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2212
  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2213
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2214
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2215
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2216
lemma lin_dense:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2217
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2218
    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
  2219
      (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
  2220
    and lx: "l < x" and xu: "x < u"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2221
    and px: "Ifm vs (x # bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2222
    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
  2223
  shows "Ifm vs (y#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2224
  using lp px noS
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2225
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
  2226
  case (5 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2227
  from "5.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2228
  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
  2229
    and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2230
    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
  2231
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2232
  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
  2233
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2234
  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
  2235
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2236
  consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2237
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2238
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2239
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2240
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2241
      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
  2242
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2243
    case N: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2244
    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
  2245
    have px': "x < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2246
      by (auto simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2247
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2248
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2249
      assume y: "y < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2250
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2251
        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
  2252
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2253
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2254
      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
  2255
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2256
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2257
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2258
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2259
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2260
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2261
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2262
      with lx px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2263
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2264
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2265
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2266
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2267
    case N: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2268
    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
  2269
    have px': "x > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2270
      by (auto simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2271
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2272
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2273
      assume y: "y > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2274
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2275
        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
  2276
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2277
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2278
      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
  2279
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2280
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2281
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2282
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2283
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2284
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2285
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2286
      with xu px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2287
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2288
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2289
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2290
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2291
next
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2292
  case (6 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2293
  from "6.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2294
  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
  2295
    and px: "Ifm vs (x # bs) (Le (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2296
    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
  2297
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2298
  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
  2299
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2300
  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
  2301
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2302
  have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2303
  consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2304
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2305
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2306
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2307
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2308
      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
  2309
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2310
    case N: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2311
    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
  2312
    have px': "x \<le> - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2313
      by (simp add: not_less field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2314
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2315
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2316
      assume y: "y < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2317
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2318
        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
  2319
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2320
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2321
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2322
        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
  2323
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2324
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2325
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2326
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2327
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2328
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2329
      with lx px' have False
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2330
        by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2331
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2332
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2333
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2334
    case N: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2335
    from px neg_divide_le_eq[OF N, where a="x" and b="-?Nt x s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2336
    have px': "x >= - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2337
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2338
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2339
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2340
      assume y: "y > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2341
      then have "y * ?N c < - ?Nt x s"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2342
        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
  2343
      then have "?N c * y + ?Nt x s < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2344
        by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2345
      then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2346
        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
  2347
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2348
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2349
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2350
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2351
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2352
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2353
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2354
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2355
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2356
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2357
next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2358
  case (3 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2359
  from "3.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2360
  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
  2361
    and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2362
    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
  2363
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2364
  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
  2365
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2366
  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
  2367
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2368
  consider "?N c = 0" | "?N c < 0" | "?N c > 0" by arith
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2369
  then show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2370
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2371
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2372
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2373
      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
  2374
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2375
    case 2
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2376
    then have cnz: "?N c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2377
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2378
    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2379
    have px': "x = - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2380
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2381
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2382
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2383
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2384
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2385
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2386
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2387
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2388
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2389
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2390
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2391
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2392
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2393
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2394
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2395
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2396
      with lx px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2397
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2398
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2399
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2400
    case 3
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2401
    then have cnz: "?N c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2402
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2403
    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2404
    have px': "x = - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2405
      by (simp add: field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2406
    from ycs show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2407
    proof
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2408
      assume y: "y < -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2409
      with ly have eu: "l < - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2410
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2411
      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2412
        by (cases "- ?Nt x s / ?N c < u") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2413
      with xu px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2414
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2415
    next
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2416
      assume y: "y > -?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2417
      with yu have eu: "u > - ?Nt x s / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2418
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2419
      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2420
        by (cases "- ?Nt x s / ?N c > l") auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2421
      with lx px' have False by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2422
      then show ?thesis ..
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2423
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2424
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2425
next
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2426
  case (4 c s)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2427
  from "4.prems"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2428
  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
  2429
    and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2430
    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
  2431
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2432
  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
  2433
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2434
  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
  2435
    by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2436
  show ?case
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2437
  proof (cases "?N c = 0")
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2438
    case True
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2439
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2440
      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
  2441
  next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2442
    case False
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2443
    with yne eq_divide_eq[of "y" "- ?Nt x s" "?N c"]
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2444
    show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2445
      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
  2446
  qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2447
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
  2448
  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
  2449
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2450
lemma inf_uset:
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2451
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2452
    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
  2453
    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
  2454
    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
  2455
  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
  2456
    ?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
  2457
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2458
  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
  2459
  let ?N = "Ipoly vs"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2460
  let ?U = "set (uset p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2461
  from ex obtain a where pa: "?I a p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2462
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2463
  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
  2464
  have nmi': "\<not> (?I a (?M p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2465
    by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2466
  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
  2467
  have npi': "\<not> (?I a (?P p))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2468
    by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2469
  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
  2470
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2471
    let ?M = "(\<lambda>(c,t). - ?Nt a t / ?N c) ` ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2472
    have fM: "finite ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2473
      by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2474
    from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2475
    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
  2476
        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
  2477
      by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2478
    then obtain "c" "t" "d" "s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2479
      where ctU: "(c,t) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2480
        and dsU: "(d,s) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2481
        and xs1: "a \<le> - ?Nt x s / ?N d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2482
        and tx1: "a \<ge> - ?Nt x t / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2483
      by blast
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2484
    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
  2485
    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
  2486
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2487
    from ctU have Mne: "?M \<noteq> {}" by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2488
    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
  2489
    let ?l = "Min ?M"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2490
    let ?u = "Max ?M"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2491
    have linM: "?l \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2492
      using fM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2493
    have uinM: "?u \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2494
      using fM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2495
    have ctM: "- ?Nt a t / ?N c \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2496
      using ctU by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2497
    have dsM: "- ?Nt a s / ?N d \<in> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2498
      using dsU by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2499
    have lM: "\<forall>t\<in> ?M. ?l \<le> t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2500
      using Mne fM by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2501
    have Mu: "\<forall>t\<in> ?M. t \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2502
      using Mne fM by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2503
    have "?l \<le> - ?Nt a t / ?N c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2504
      using ctM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2505
    then have lx: "?l \<le> a"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2506
      using tx by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2507
    have "- ?Nt a s / ?N d \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2508
      using dsM Mne by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2509
    then have xu: "a \<le> ?u"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2510
      using xs by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2511
    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
  2512
    consider u where "u \<in> ?M" "?I u p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2513
      | 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
  2514
      by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2515
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2516
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2517
      case 1
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2518
      then have "\<exists>(nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2519
        by auto
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2520
      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
  2521
        by blast
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2522
      have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / 2) p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2523
        using \<open>?I u p\<close> tuu by simp
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2524
      with tuU show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2525
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2526
      case 2
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2527
      have "\<exists>(t1n, t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2528
        using \<open>t1 \<in> ?M\<close> by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2529
      then obtain t1u t1n where t1uU: "(t1n, t1u) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2530
        and t1u: "t1 = - ?Nt a t1u / ?N t1n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2531
        by blast
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2532
      have "\<exists>(t2n, t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2533
        using \<open>t2 \<in> ?M\<close> by auto
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2534
      then obtain t2u t2n where t2uU: "(t2n, t2u) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2535
        and t2u: "t2 = - ?Nt a t2u / ?N t2n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2536
        by blast
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2537
      have "t1 < t2"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2538
        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
  2539
      let ?u = "(t1 + t2) / 2"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2540
      have "t1 < ?u"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2541
        using less_half_sum [OF \<open>t1 < t2\<close>] by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2542
      have "?u < t2"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2543
        using gt_half_sum [OF \<open>t1 < t2\<close>] by auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2544
      have "?I ?u p"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2545
        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
  2546
        by (rule lin_dense)
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2547
      with t1uU t2uU t1u t2u show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2548
    qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2549
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2550
  then obtain l n s  m
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2551
    where lnU: "(n, l) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2552
      and smU:"(m,s) \<in> ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2553
      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
  2554
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2555
  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
  2556
    by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2557
  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
  2558
    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
  2559
  have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / 2) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2560
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2561
  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
  2562
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2563
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2564
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2565
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
  2566
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2567
theorem fr_eq:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2568
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2569
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2570
    (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2571
     Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2572
     (\<exists>(n, t) \<in> set (uset p). \<exists>(m, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2573
       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
  2574
  (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
  2575
proof
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2576
  show ?D if ?E
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2577
  proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2578
    consider "?M \<or> ?P" | "\<not> ?M" "\<not> ?P" by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2579
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2580
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2581
      case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2582
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2583
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2584
      case 2
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2585
      from inf_uset[OF lp this] have ?F
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2586
        using \<open>?E\<close> by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2587
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2588
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2589
  qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2590
  show ?E if ?D
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2591
  proof -
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2592
    from that consider ?M | ?P | ?F by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2593
    then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2594
    proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2595
      case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2596
      from minusinf_ex[OF lp this] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2597
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2598
      case 2
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2599
      from plusinf_ex[OF lp this] show ?thesis .
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2600
    next
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2601
      case 3
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2602
      then show ?thesis by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2603
    qed
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2604
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2605
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2606
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2607
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  2608
section \<open>First implementation : Naive by encoding all case splits locally\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2609
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2610
definition "msubsteq c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2611
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2612
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2613
    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
  2614
   (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
  2615
   (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
  2616
   (conj (Eq (CP c)) (Eq (CP d)), Eq r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2617
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2618
lemma msubsteq_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2619
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2620
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2621
    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
  2622
  shows "bound0 (msubsteq c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2623
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2624
  have th: "\<forall>x \<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2625
    [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2626
      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
  2627
     (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
  2628
     (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
  2629
     (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
  2630
    using lp by (simp add: Let_def t s)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2631
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2632
    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
  2633
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2634
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2635
lemma msubsteq:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2636
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2637
  shows "Ifm vs (x#bs) (msubsteq c t d s a r) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2638
    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
  2639
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2640
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2641
  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
  2642
  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
  2643
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2644
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2645
  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
  2646
  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
  2647
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2648
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2649
  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
  2650
    by simp_all
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2651
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2652
  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
  2653
    by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2654
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2655
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2656
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2657
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2658
      by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2659
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2660
    case cd: 2
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2661
    then have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2*?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2662
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2663
    have "?rhs = Ifm vs (-?s / (2*?d) # bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2664
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2665
    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
  2666
      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
  2667
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (-?s / (2*?d)) + ?r) = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2668
      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
  2669
    also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2 * ?d * ?r= 0"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 48562
diff changeset
  2670
      by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2671
    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2672
      using cd(2) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2673
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2674
      using cd
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  2675
      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
  2676
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2677
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2678
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2679
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2680
    have "?rhs = Ifm vs (-?t / (2*?c) # bs) (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2681
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2682
    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
  2683
      by (simp add: r[of "- (?t/ (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2684
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (-?t / (2 * ?c)) + ?r) = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2685
      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
  2686
    also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2 * ?c) / (2 * ?c) + 2 * ?c * ?r= 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2687
      by (simp add: field_simps distrib_left[of "2 * ?c"] del: distrib_left)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2688
    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2 * ?c * ?r = 0"
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2689
      using cd(1) by simp
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2690
    finally show ?thesis using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2691
      by (simp add: r[of "- (?t/ (2 * ?c))"] msubsteq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2692
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2693
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2694
    then have cd2: "?c * ?d * 2 \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2695
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2696
    from add_frac_eq[OF cd, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2697
    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
  2698
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2699
    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
  2700
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2701
    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
  2702
      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2703
    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
  2704
      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
  2705
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2706
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2 * ?c * ?d * ?r = 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2707
      using nonzero_mult_divide_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
  2708
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2709
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2710
      using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2711
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2712
          msubsteq_def Let_def evaldjf_ex field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2713
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2714
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2715
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2716
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2717
definition "msubstneq c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2718
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2719
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2720
    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
  2721
   (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
  2722
   (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
  2723
   (conj (Eq (CP c)) (Eq (CP d)), NEq r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2724
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2725
lemma msubstneq_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2726
  assumes lp: "islin (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2727
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2728
    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
  2729
  shows "bound0 (msubstneq c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2730
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2731
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2732
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2733
     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
  2734
    (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
  2735
    (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
  2736
    (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
  2737
    using lp by (simp add: Let_def t s)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2738
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2739
    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
  2740
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2741
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2742
lemma msubstneq:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2743
  assumes lp: "islin (Eq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2744
  shows "Ifm vs (x#bs) (msubstneq c t d s a r) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2745
    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
  2746
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2747
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2748
  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
  2749
  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
  2750
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2751
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2752
  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
  2753
  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
  2754
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2755
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2756
  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
  2757
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2758
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2759
  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
  2760
    by blast
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2761
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2762
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2763
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2764
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2765
      by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2766
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2767
    case cd: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2768
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2 * ?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2769
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2770
    have "?rhs = Ifm vs (-?s / (2*?d) # bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2771
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2772
    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
  2773
      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
  2774
    also have "\<dots> \<longleftrightarrow> 2*?d * (?a * (-?s / (2*?d)) + ?r) \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2775
      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
  2776
    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
  2777
      by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2778
    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2779
      using cd(2) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2780
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2781
      using cd
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  2782
      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
  2783
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2784
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2785
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2*?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2786
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2787
    have "?rhs = Ifm vs (-?t / (2*?c) # bs) (NEq (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2788
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2789
    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
  2790
      by (simp add: r[of "- (?t/ (2 * ?c))"])
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2791
    also have "\<dots> \<longleftrightarrow> 2*?c * (?a * (-?t / (2*?c)) + ?r) \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2792
      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
  2793
    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
  2794
      by (simp add: field_simps distrib_left[of "2*?c"] del: distrib_left)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2795
    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2*?c*?r \<noteq> 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2796
      using cd(1) by simp
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2797
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2798
      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
  2799
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2800
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2801
    then have cd2: "?c * ?d * 2 \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2802
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2803
    from add_frac_eq[OF cd, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2804
    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
  2805
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2806
    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
  2807
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2808
    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
  2809
      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2810
    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
  2811
      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
  2812
      by simp
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"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2814
      using nonzero_mult_divide_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
  2815
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2816
    finally show ?thesis
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2817
      using cd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2818
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2819
          msubstneq_def Let_def evaldjf_ex field_simps)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2820
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2821
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2822
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2823
definition "msubstlt c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  2824
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2825
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2826
    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
  2827
   (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2828
    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
  2829
   (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
  2830
   (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
  2831
   (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
  2832
   (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
  2833
   (conj (Eq (CP c)) (Eq (CP d)), Lt r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2834
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2835
lemma msubstlt_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2836
  assumes lp: "islin (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2837
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2838
    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
  2839
  shows "bound0 (msubstlt c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2840
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2841
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2842
  [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2843
    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
  2844
   (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2845
    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
  2846
   (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
  2847
   (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
  2848
   (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
  2849
   (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
  2850
   (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
  2851
    using lp by (simp add: Let_def t s lt_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2852
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2853
    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
  2854
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2855
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2856
lemma msubstlt:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2857
  assumes nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2858
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2859
    and lp: "islin (Lt (CNP 0 a r))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2860
  shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2861
    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
  2862
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2863
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2864
  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
  2865
  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
  2866
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2867
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2868
  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
  2869
  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
  2870
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2871
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2872
  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
  2873
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2874
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2875
  consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2876
    | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0"
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2877
    by atomize_elim auto
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2878
  then show ?thesis
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2879
  proof cases
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2880
    case 1
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2881
    then show ?thesis
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2882
      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
  2883
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2884
    case cd: 2
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2885
    then have cd2: "2 * ?c * ?d > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2886
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2887
    from cd have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2888
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2889
    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
  2890
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2891
    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
  2892
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2893
    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
  2894
      by (simp only: th)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2895
    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
  2896
      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
  2897
    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
  2898
      using cd2 cd2'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2899
        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
  2900
      by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2901
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r < 0"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  2902
      using nonzero_mult_divide_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
  2903
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2904
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2905
      using cd c d nc nd cd2'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2906
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2907
          msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2908
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2909
    case cd: 3
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2910
    then have cd2: "2 * ?c * ?d < 0"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2911
      by (simp add: mult_less_0_iff field_simps)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2912
    from cd have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2913
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  2914
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2915
    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
  2916
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2917
    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
  2918
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2919
    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
  2920
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2921
    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
  2922
      using cd2 order_less_not_sym[OF cd2]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2923
        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
  2924
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2925
    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2926
      using nonzero_mult_divide_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
  2927
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2928
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2929
      using cd c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2930
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2931
          msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2932
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2933
    case cd: 4
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2934
    from cd(1) have c'': "2 * ?c > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2935
      by (simp add: zero_less_mult_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2936
    from cd(1) have c': "2 * ?c \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2937
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2938
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2939
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2940
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2941
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2942
    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
  2943
      by (simp add: r[of "- (?t / (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2944
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2945
      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
  2946
        order_less_not_sym[OF c'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2947
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2948
    also have "\<dots> \<longleftrightarrow> - ?a * ?t + 2 * ?c * ?r < 0"
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2949
      using nonzero_mult_divide_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
  2950
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2951
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2952
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2953
      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
  2954
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2955
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2956
    case cd: 5
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2957
    from cd(1) have c': "2 * ?c \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2958
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2959
    from cd(1) have c'': "2 * ?c < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2960
      by (simp add: mult_less_0_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2961
    from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2962
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2963
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2*?c) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2964
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2965
    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
  2966
      by (simp add: r[of "- (?t / (2*?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2967
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) > 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2968
      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
  2969
        mult_less_cancel_left_disj[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2970
      by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  2971
    also have "\<dots> \<longleftrightarrow> ?a*?t -  2*?c *?r < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2972
      using nonzero_mult_divide_cancel_left[OF c'] cd(1) order_less_not_sym[OF c'']
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2973
          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
  2974
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2975
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2976
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2977
      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
  2978
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  2979
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2980
    case cd: 6
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2981
    from cd(2) have d'': "2 * ?d > 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2982
      by (simp add: zero_less_mult_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2983
    from cd(2) have d': "2 * ?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2984
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2985
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2986
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2987
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2988
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2989
    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
  2990
      by (simp add: r[of "- (?s / (2 * ?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2991
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d))+ ?r) < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2992
      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
  2993
        order_less_not_sym[OF d'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2994
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  2995
    also have "\<dots> \<longleftrightarrow> - ?a * ?s+  2 * ?d * ?r < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2996
      using nonzero_mult_divide_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
  2997
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2998
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  2999
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3000
      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
  3001
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  3002
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3003
    case cd: 7
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3004
    from cd(2) have d': "2 * ?d \<noteq> 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3005
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3006
    from cd(2) have d'': "2 * ?d < 0"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3007
      by (simp add: mult_less_0_iff)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3008
    from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3009
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3010
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3011
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3012
    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
  3013
      by (simp add: r[of "- (?s / (2 * ?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3014
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) > 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3015
      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
  3016
        mult_less_cancel_left_disj[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3017
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3018
    also have "\<dots> \<longleftrightarrow> ?a * ?s -  2 * ?d * ?r < 0"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3019
      using nonzero_mult_divide_cancel_left[OF d'] cd(2) order_less_not_sym[OF d'']
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3020
          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
  3021
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3022
    finally show ?thesis
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60561
diff changeset
  3023
      using cd nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3024
      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
  3025
          lt polyneg_norm polymul_norm)
60561
85aed595feb4 tuned proofs;
wenzelm
parents: 60560
diff changeset
  3026
  qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3027
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3028
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3029
definition "msubstle c t d s a r =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 60754
diff changeset
  3030
  evaldjf (case_prod conj)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3031
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3032
     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
  3033
    (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3034
     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
  3035
    (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
  3036
    (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
  3037
    (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
  3038
    (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
  3039
    (conj (Eq (CP c)) (Eq (CP d)), Le r)]"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3040
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3041
lemma msubstle_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3042
  assumes lp: "islin (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3043
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3044
    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
  3045
  shows "bound0 (msubstle c t d s a r)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3046
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3047
  have th: "\<forall>x\<in> set
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3048
   [(let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3049
     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
  3050
    (let cd = c *\<^sub>p d
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3051
     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
  3052
    (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
  3053
    (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
  3054
    (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
  3055
    (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
  3056
    (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
  3057
    using lp by (simp add: Let_def t s lt_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3058
  from evaldjf_bound0[OF th] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3059
    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
  3060
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3061
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3062
lemma msubstle:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3063
  assumes nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3064
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3065
    and lp: "islin (Le (CNP 0 a r))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3066
  shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3067
    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
  3068
  (is "?lhs = ?rhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3069
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3070
  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
  3071
  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
  3072
  let ?c = "?N c"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3073
  let ?d = "?N d"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3074
  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
  3075
  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
  3076
  let ?a = "?N a"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3077
  let ?r = "?Nt x r"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3078
  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
  3079
    by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3080
  note r = tmbound0_I[OF lin(3), of vs _ bs x]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3081
  have cd_cs: "?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)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3082
    by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3083
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3084
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3085
    assume c: "?c = 0" and d: "?d = 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3086
    then have ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3087
      using nc nd
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3088
      by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3089
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3090
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3091
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3092
    assume dc: "?c * ?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3093
    from dc have dc': "2 * ?c * ?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3094
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3095
    then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3096
      by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3097
    from dc' have dc'': "\<not> 2 * ?c * ?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3098
      by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3099
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3100
    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
  3101
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3102
    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
  3103
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3104
    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
  3105
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3106
    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
  3107
      using dc' dc''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3108
        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
  3109
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3110
    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<le> 0"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3111
      using nonzero_mult_divide_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
  3112
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3113
    finally  have ?thesis using dc c d  nc nd dc'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3114
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3115
          Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3116
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3117
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3118
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3119
    assume dc: "?c * ?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3120
    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
  3121
      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
  3122
    then have c: "?c \<noteq> 0" and d: "?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3123
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3124
    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3125
    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
  3126
      by (simp add: field_simps)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3127
    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
  3128
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3129
    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
  3130
      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3131
    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
  3132
      using dc' order_less_not_sym[OF dc']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3133
        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
  3134
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3135
    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3136
      using nonzero_mult_divide_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
  3137
      by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3138
    finally  have ?thesis using dc c d  nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3139
      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3140
          Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3141
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3142
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3143
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3144
    assume c: "?c > 0" and d: "?d = 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3145
    from c have c'': "2 * ?c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3146
      by (simp add: zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3147
    from c have c': "2 * ?c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3148
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3149
    from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3150
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3151
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3152
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3153
    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
  3154
      by (simp add: r[of "- (?t / (2 * ?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3155
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3156
      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
  3157
        order_less_not_sym[OF c'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3158
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3159
    also have "\<dots> \<longleftrightarrow> - ?a*?t+  2*?c *?r \<le> 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3160
      using nonzero_mult_divide_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
  3161
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3162
    finally have ?thesis using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3163
      by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3164
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3165
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3166
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3167
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3168
    assume c: "?c < 0" and d: "?d = 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3169
    then have c': "2 * ?c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3170
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3171
    from c have c'': "2 * ?c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3172
      by (simp add: mult_less_0_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3173
    from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3174
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3175
    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3176
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3177
    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
  3178
      by (simp add: r[of "- (?t / (2*?c))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3179
    also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \<ge> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3180
      using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3181
        mult_le_cancel_left[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3182
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3183
    also have "\<dots> \<longleftrightarrow> ?a * ?t - 2 * ?c * ?r \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3184
      using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3185
          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
  3186
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3187
    finally have ?thesis using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3188
      by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3189
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3190
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3191
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3192
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3193
    assume c: "?c = 0" and d: "?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3194
    from d have d'': "2 * ?d > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3195
      by (simp add: zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3196
    from d have d': "2 * ?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3197
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3198
    from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3199
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3200
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2 * ?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3201
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3202
    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
  3203
      by (simp add: r[of "- (?s / (2*?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3204
    also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3205
      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
  3206
        order_less_not_sym[OF d'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3207
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3208
    also have "\<dots> \<longleftrightarrow> - ?a * ?s + 2 * ?d * ?r \<le> 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3209
      using nonzero_mult_divide_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
  3210
      by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3211
    finally have ?thesis using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3212
      by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3213
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3214
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3215
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3216
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3217
    assume c: "?c = 0" and d: "?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3218
    then have d': "2 * ?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3219
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3220
    from d have d'': "2 * ?d < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3221
      by (simp add: mult_less_0_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3222
    from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3223
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3224
    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / (2*?d) # bs) (Le (CNP 0 a r))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3225
      by (simp only: th)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3226
    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
  3227
      by (simp add: r[of "- (?s / (2*?d))"])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3228
    also have "\<dots> \<longleftrightarrow> 2*?d * (?a* (- ?s / (2*?d))+ ?r) \<ge> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3229
      using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3230
        mult_le_cancel_left[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3231
      by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3232
    also have "\<dots> \<longleftrightarrow> ?a * ?s -  2 * ?d * ?r \<le> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3233
      using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d'']
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3234
          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
  3235
        by (simp add: algebra_simps diff_divide_distrib del: distrib_right)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3236
    finally have ?thesis using c d nc nd
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3237
      by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3238
          evaldjf_ex field_simps lt polyneg_norm polymul_norm)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3239
  }
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3240
  ultimately show ?thesis by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3241
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3242
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3243
fun msubst :: "fm \<Rightarrow> (poly \<times> tm) \<times> (poly \<times> tm) \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3244
where
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3245
  "msubst (And p q) ((c, t), (d, s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3246
| "msubst (Or p q) ((c, t), (d, s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c, t), (d, s)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3247
| "msubst (Eq (CNP 0 a r)) ((c, t), (d, s)) = msubsteq c t d s a r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3248
| "msubst (NEq (CNP 0 a r)) ((c, t), (d, s)) = msubstneq c t d s a r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3249
| "msubst (Lt (CNP 0 a r)) ((c, t), (d, s)) = msubstlt c t d s a r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3250
| "msubst (Le (CNP 0 a r)) ((c, t), (d, s)) = msubstle c t d s a r"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3251
| "msubst p ((c, t), (d, s)) = p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3252
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3253
lemma msubst_I:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3254
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3255
    and nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3256
    and nd: "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3257
  shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3258
    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
  3259
  using lp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3260
  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
  3261
    (auto simp add: tmbound0_I
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3262
      [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
  3263
        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
  3264
      msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3265
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3266
lemma msubst_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3267
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3268
    and t: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3269
    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
  3270
  shows "bound0 (msubst p ((c,t),(d,s)))"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3271
  using lp t s
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3272
  by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3273
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3274
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
  3275
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3276
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3277
    (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3278
     Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3279
     (\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3280
      Ifm vs (x#bs) (msubst p ((c, t), (d, s)))))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3281
  (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
  3282
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3283
  from uset_l[OF lp]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3284
  have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3285
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3286
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3287
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3288
    assume ctU: "(c, t) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3289
      and dsU: "(d,s) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3290
      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"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3291
    from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3292
      by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3293
    from msubst_I[OF lp norm, of vs x bs t s] pts
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3294
    have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3295
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3296
  moreover
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3297
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3298
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3299
    assume ctU: "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3300
      and dsU: "(d,s) \<in>set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3301
      and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3302
    from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3303
      by simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3304
    from msubst_I[OF lp norm, of vs x bs t s] pts
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3305
    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
  3306
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3307
  ultimately have th': "(\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3308
      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
  3309
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3310
  from fr_eq[OF lp, of vs bs x, simplified th'] show ?thesis .
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3311
qed
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3312
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3313
lemma simpfm_lin:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  3314
  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
  3315
  shows "qfree p \<Longrightarrow> islin (simpfm p)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3316
  by (induct p rule: simpfm.induct) (auto simp add: conj_lin disj_lin)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3317
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3318
definition "ferrack p \<equiv>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3319
  let
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3320
    q = simpfm p;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3321
    mp = minusinf q;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3322
    pp = plusinf q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3323
  in
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3324
    if (mp = T \<or> pp = T) then T
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3325
    else
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3326
     (let U = alluopairs (remdups (uset  q))
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3327
      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
  3328
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3329
lemma ferrack:
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3330
  assumes qf: "qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3331
  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
  3332
  (is "_ \<and> ?rhs = ?lhs")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3333
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3334
  let ?I = "\<lambda>x p. Ifm vs (x#bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3335
  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
  3336
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3337
  let ?q = "simpfm p"
41823
81d64ec48427 eliminated remdps in favor of List.remdups
krauss
parents: 41822
diff changeset
  3338
  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
  3339
  let ?Up = "alluopairs ?U"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3340
  let ?mp = "minusinf ?q"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3341
  let ?pp = "plusinf ?q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3342
  fix x
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3343
  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
  3344
  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
  3345
  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
  3346
  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
  3347
  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
  3348
    by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3349
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3350
    fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3351
    assume ctU: "(c, t) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3352
      and dsU: "(d,s) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3353
    from U_l ctU dsU have norm: "isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3354
      by auto
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3355
    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
  3356
    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
  3357
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3358
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3359
  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
  3360
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3361
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3362
    fix x
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3363
    assume xUp: "x \<in> set ?Up"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3364
    then obtain c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3365
      where ctU: "(c, t) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3366
        and dsU: "(d,s) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3367
        and x: "x = ((c, t),(d, s))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3368
      using alluopairs_set1[of ?U] by auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3369
    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
  3370
    have nbs: "tmbound0 t" "tmbound0 s" by simp_all
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3371
    from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3372
    have "bound0 ((simpfm o (msubst (simpfm p))) x)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3373
      using x by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3374
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3375
  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
  3376
  have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3377
  with mp_nb pp_nb
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3378
  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3379
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3380
  from decr0_qf[OF th1] have thqf: "qfree (ferrack p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3381
    by (simp add: ferrack_def Let_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3382
  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3383
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3384
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3385
      (\<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
  3386
    using fr_eq_msubst[OF lq, of vs bs x] by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3387
  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3388
      (\<exists>(x, y) \<in> set ?Up. ?I ((simpfm \<circ> msubst ?q) (x, y)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3389
    using alluopairs_bex[OF th0] by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3390
  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
  3391
    by (simp add: evaldjf_ex)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3392
  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
  3393
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3394
  also have "\<dots> \<longleftrightarrow> ?rhs"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3395
    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
  3396
    apply (simp add: ferrack_def Let_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3397
    apply (cases "?mp = T \<or> ?pp = T")
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3398
    apply auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3399
    done
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3400
  finally show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3401
    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
  3402
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3403
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3404
definition "frpar p = simpfm (qelim p ferrack)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3405
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3406
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
  3407
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3408
  from ferrack
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3409
  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
  3410
    by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3411
  from qelim[OF th, of p bs] show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3412
    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
  3413
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3414
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3415
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  3416
section \<open>Second implemenation: Case splits not local\<close>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3417
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3418
lemma fr_eq2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3419
  assumes lp: "islin p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3420
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3421
   (Ifm vs (x#bs) (minusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3422
    Ifm vs (x#bs) (plusinf p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3423
    Ifm vs (0#bs) p \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3424
    (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3425
      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
  3426
      (\<exists>(n, t) \<in> set (uset p). \<exists>(m, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3427
        Ipoly vs n \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3428
        Ipoly vs m \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3429
        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
  3430
  (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
  3431
proof
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3432
  assume px: "\<exists>x. ?I x p"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3433
  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3434
  moreover {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3435
    assume "?M \<or> ?P"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3436
    then have "?D" by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3437
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3438
  moreover {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3439
    assume nmi: "\<not> ?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3440
      and npi: "\<not> ?P"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3441
    from inf_uset[OF lp nmi npi, OF px]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3442
    obtain c t d s where ct:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3443
      "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3444
      "(d, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3445
      "?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
  3446
      by auto
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3447
    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
  3448
    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
  3449
    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
  3450
    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
  3451
    have eq2: "\<And>(x::'a). x + x = 2 * x"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3452
      by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3453
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3454
      assume "?c = 0 \<and> ?d = 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3455
      with ct have ?D by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3456
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3457
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3458
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3459
      assume z: "?c = 0" "?d \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3460
      from z have ?D using ct by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3461
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3462
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3463
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3464
      assume z: "?c \<noteq> 0" "?d = 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3465
      with ct have ?D by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3466
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3467
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3468
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3469
      assume z: "?c \<noteq> 0" "?d \<noteq> 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3470
      from z have ?F 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)"])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3473
        apply simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3474
        apply (rule bexI[where x = "(d,s)"])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3475
        apply simp_all
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3476
        done
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3477
      then have ?D by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3478
    }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3479
    ultimately have ?D by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3480
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3481
  ultimately show "?D" by blast
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3482
next
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3483
  assume "?D"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3484
  moreover {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3485
    assume m: "?M"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3486
    from minusinf_ex[OF lp m] have "?E" .
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3487
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3488
  moreover {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3489
    assume p: "?P"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3490
    from plusinf_ex[OF lp p] have "?E" .
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3491
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3492
  moreover {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3493
    assume f:"?F"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3494
    then have "?E" by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3495
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3496
  ultimately show "?E" by blast
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3497
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3498
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3499
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
  3500
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
  3501
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
  3502
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
  3503
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
  3504
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3505
lemma msubsteq2:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3506
  assumes nz: "Ipoly vs c \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3507
    and l: "islin (Eq (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3508
  shows "Ifm vs (x#bs) (msubsteq2 c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3509
    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
  3510
  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
  3511
  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
  3512
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3513
lemma msubstltpos:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3514
  assumes nz: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3515
    and l: "islin (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3516
  shows "Ifm vs (x#bs) (msubstltpos c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3517
    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
  3518
  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
  3519
  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
  3520
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3521
lemma msubstlepos:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3522
  assumes nz: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3523
    and l: "islin (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3524
  shows "Ifm vs (x#bs) (msubstlepos c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3525
    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
  3526
  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
  3527
  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
  3528
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3529
lemma msubstltneg:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3530
  assumes nz: "Ipoly vs c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3531
    and l: "islin (Lt (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3532
  shows "Ifm vs (x#bs) (msubstltneg c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3533
    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
  3534
  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
  3535
  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
  3536
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3537
lemma msubstleneg:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3538
  assumes nz: "Ipoly vs c < 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3539
    and l: "islin (Le (CNP 0 a b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3540
  shows "Ifm vs (x#bs) (msubstleneg c t a b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3541
    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
  3542
  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
  3543
  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
  3544
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3545
fun msubstpos :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3546
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3547
  "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3548
| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3549
| "msubstpos (Eq (CNP 0 a r)) c t = 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
  3550
| "msubstpos (NEq (CNP 0 a r)) c t = NOT (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
  3551
| "msubstpos (Lt (CNP 0 a r)) c t = 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
  3552
| "msubstpos (Le (CNP 0 a r)) c t = 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
  3553
| "msubstpos p c t = p"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3554
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3555
lemma msubstpos_I:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3556
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3557
    and pos: "Ipoly vs c > 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3558
  shows "Ifm vs (x#bs) (msubstpos p c t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3559
    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
  3560
  using lp pos
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3561
  by (induct p rule: islin.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3562
    (auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3563
      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
  3564
      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
  3565
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3566
fun msubstneg :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3567
where
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3568
  "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3569
| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3570
| "msubstneg (Eq (CNP 0 a r)) c t = 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
  3571
| "msubstneg (NEq (CNP 0 a r)) c t = NOT (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
  3572
| "msubstneg (Lt (CNP 0 a r)) c t = 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
  3573
| "msubstneg (Le (CNP 0 a r)) c t = 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
  3574
| "msubstneg p c t = p"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3575
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3576
lemma msubstneg_I:
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3577
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3578
    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
  3579
  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
  3580
  using lp pos
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3581
  by (induct p rule: islin.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3582
    (auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3583
      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
  3584
      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
  3585
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3586
definition
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3587
  "msubst2 p c t =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3588
    disj
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3589
      (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t)))
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3590
      (conj (lt (CP c)) (simpfm (msubstneg p c t)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3591
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3592
lemma msubst2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3593
  assumes lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3594
    and nc: "isnpoly c"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3595
    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
  3596
  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
  3597
proof -
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3598
  let ?c = "Ipoly vs c"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3599
  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
  3600
    by (simp_all add: polyneg_norm)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3601
  from nz have "?c > 0 \<or> ?c < 0" by arith
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3602
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3603
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3604
    assume c: "?c < 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3605
    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"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3606
    have ?thesis by (auto simp add: msubst2_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3607
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3608
  moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3609
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3610
    assume c: "?c > 0"
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3611
    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"]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3612
    have ?thesis by (auto simp add: msubst2_def)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3613
  }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3614
  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
  3615
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3616
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3617
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
  3618
  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
  3619
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3620
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
  3621
  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
  3622
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
  3623
  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
  3624
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3625
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
  3626
  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
  3627
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
  3628
  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
  3629
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3630
lemma msubstpos_nb:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3631
  assumes 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 (msubstpos 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: msubstpos.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3636
    (auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3637
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3638
lemma msubstneg_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 (msubstneg 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 (induct p c t rule: msubstneg.induct)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3645
    (auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3646
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3647
lemma msubst2_nb:
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  3648
  assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3649
    and lp: "islin p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3650
    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
  3651
  shows "bound0 (msubst2 p c t)"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3652
  using lp tnb
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3653
  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
  3654
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  3655
lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3656
  by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3657
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45499
diff changeset
  3658
lemma mult_minus2_right: "(x::'a::comm_ring_1) * -2 = - (x * 2)"
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  3659
  by simp
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3660
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3661
lemma islin_qf: "islin p \<Longrightarrow> qfree p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3662
  by (induct p rule: islin.induct) (auto simp add: bound0_qf)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3663
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3664
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
  3665
  assumes lp: "islin p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3666
  shows "(\<exists>x. Ifm vs (x#bs) p) \<longleftrightarrow>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3667
    ((Ifm vs (x#bs) (minusinf p)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3668
     (Ifm vs (x#bs) (plusinf p)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3669
     Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3670
     (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3671
      Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3672
      (\<exists>(c, t) \<in> set (uset p). \<exists>(d, s) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3673
        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
  3674
  (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
  3675
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3676
  from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3677
    by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3678
  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3679
  have n2: "isnpoly (C (-2,1))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3680
    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
  3681
  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
  3682
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3683
  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
  3684
    (\<exists>(n, t) \<in> set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3685
      \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3686
      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
  3687
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3688
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3689
      fix n t
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3690
      assume H: "(n, t) \<in> set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3691
      from H(1) th have "isnpoly n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3692
        by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3693
      then have nn: "isnpoly (n *\<^sub>p (C (-2,1)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3694
        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
  3695
      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
  3696
        by (simp add: polyneg_norm nn)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3697
      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
  3698
        using H(2) nn' nn
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33212
diff changeset
  3699
        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
  3700
      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
  3701
      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
  3702
        using H(2) nn2 by (simp add: mult_minus2_right)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3703
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3704
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3705
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3706
      fix n t
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3707
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3708
        "(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
  3709
        "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
  3710
      from H(1) th have "isnpoly n"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3711
        by blast
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3712
      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
  3713
        using H(2) by (simp_all add: polymul_norm n2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3714
      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
  3715
        using H(2,3) by (simp add: mult_minus2_right)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3716
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3717
    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
  3718
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3719
  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
  3720
      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
  3721
    (\<exists>(n, t)\<in>set (uset p). \<exists>(m, s)\<in>set (uset p).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3722
      \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3723
      \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3724
      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
  3725
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3726
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3727
      fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3728
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3729
        "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3730
        "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3731
      from H(1,2) th have "isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3732
        by blast+
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3733
      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
  3734
        by (simp_all add: polymul_norm n2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3735
      have stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3736
          "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3737
          "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
  3738
        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
  3739
      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
  3740
        using H(3)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3741
        by (auto simp add: msubst2_def lt[OF stupid(1)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3742
            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
  3743
      from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3744
      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
  3745
          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
  3746
        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
  3747
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3748
    moreover
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3749
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3750
      fix c t d s
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3751
      assume H:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3752
        "(c, t) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3753
        "(d, s) \<in> set (uset p)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3754
        "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3755
        "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3756
        "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
  3757
      from H(1,2) th have "isnpoly c" "isnpoly d"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3758
        by blast+
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3759
      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
  3760
        using H(3,4) by (simp_all add: polymul_norm n2)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3761
      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
  3762
      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
  3763
        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
  3764
    }
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3765
    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
  3766
  qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3767
  from fr_eq2[OF lp, of vs bs x] show ?thesis
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3768
    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
  3769
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3770
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3771
definition
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3772
  "ferrack2 p \<equiv>
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3773
    let
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3774
      q = simpfm p;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3775
      mp = minusinf q;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3776
      pp = plusinf q
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3777
    in
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3778
      if (mp = T \<or> pp = T) then T
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3779
      else
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3780
       (let U = remdups (uset  q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3781
        in
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3782
          decr0
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3783
            (list_disj
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3784
              [mp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3785
               pp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3786
               simpfm (subst0 (CP 0\<^sub>p) q),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3787
               evaldjf (\<lambda>(c, t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3788
               evaldjf (\<lambda>((b, a),(d, c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3789
                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
  3790
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3791
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
  3792
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3793
lemma ferrack2:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3794
  assumes qf: "qfree p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3795
  shows "qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 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
  3796
  (is "_ \<and> (?rhs = ?lhs)")
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3797
proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3798
  let ?J = "\<lambda>x p. Ifm vs (x#bs) p"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3799
  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
  3800
  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3801
  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
  3802
  let ?qz = "subst0 (CP 0\<^sub>p) ?q"
41823
81d64ec48427 eliminated remdps in favor of List.remdups
krauss
parents: 41822
diff changeset
  3803
  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
  3804
  let ?Up = "alluopairs ?U"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3805
  let ?mp = "minusinf ?q"
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3806
  let ?pp = "plusinf ?q"
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3807
  fix x
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3808
  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
  3809
  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
  3810
  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
  3811
  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
  3812
  from uset_l[OF lq]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3813
  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
  3814
    by simp
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3815
  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
  3816
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3817
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3818
      fix c t
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3819
      assume ct: "(c, t) \<in> set ?U"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3820
      then have tnb: "tmbound0 t"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3821
        using U_l by blast
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3822
      from msubst2_nb[OF lq tnb]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3823
      have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3824
    }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3825
    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
  3826
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3827
  have bnd1: "\<forall>x \<in> set ?Up. bound0 ((\<lambda>((b, a), (d, c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3828
    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
  3829
  proof -
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3830
    {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3831
      fix b a d c
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3832
      assume badc: "((b,a),(d,c)) \<in> set ?Up"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3833
      from badc U_l alluopairs_set1[of ?U]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3834
      have nb: "tmbound0 (Add (Mul d a) (Mul b c))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3835
        by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3836
      from msubst2_nb[OF lq nb]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3837
      have "bound0 ((\<lambda>((b, a),(d, c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3838
          msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3839
        by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3840
    }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3841
    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
  3842
  qed
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3843
  have stupid: "bound0 F"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3844
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3845
  let ?R =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3846
    "list_disj
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3847
     [?mp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3848
      ?pp,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3849
      simpfm (subst0 (CP 0\<^sub>p) ?q),
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3850
      evaldjf (\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) ?U,
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3851
      evaldjf (\<lambda>((b,a),(d,c)).
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3852
        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
  3853
  from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3854
    evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3855
  have nb: "bound0 ?R"
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3856
    by (simp add: list_disj_def simpfm_bound0)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3857
  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
  3858
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3859
  {
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3860
    fix b a d c
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3861
    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
  3862
    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
  3863
      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
  3864
    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
  3865
      using norm by (simp_all add: polymul_norm)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3866
    have stupid:
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3867
        "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p b *\<^sub>p d))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3868
        "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p d *\<^sub>p b))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3869
        "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p b *\<^sub>p d)))"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3870
        "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
  3871
      by (simp_all add: polyneg_norm norm2)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3872
    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
  3873
        ?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
  3874
      (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
  3875
    proof
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3876
      assume H: ?lhs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3877
      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
  3878
        by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3879
            mult_less_0_iff zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3880
      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
  3881
      show ?rhs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3882
        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
  3883
    next
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3884
      assume H: ?rhs
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3885
      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
  3886
        by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)]
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3887
            mult_less_0_iff zero_less_mult_iff)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3888
      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
  3889
      show ?lhs
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3890
        by (simp add: field_simps)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3891
    qed
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3892
  }
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3893
  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
  3894
    by auto
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3895
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3896
  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)"
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3897
    by simp
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3898
  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
  3899
      (\<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
  3900
      (\<exists>(b, a) \<in> set ?U. \<exists>(d, c) \<in> set ?U.
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3901
        ?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
  3902
    using fr_eq_msubst2[OF lq, of vs bs x] by simp
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3903
  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
  3904
      (\<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
  3905
      (\<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
  3906
    by (simp add: split_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3907
  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
  3908
      (\<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
  3909
      (\<exists>(x, y) \<in> set ?Up. ?I (?s (x, y)))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3910
    using alluopairs_bex[OF th0] by simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3911
  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
  3912
    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
  3913
  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
  3914
    unfolding ferrack2_def
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3915
    apply (cases "?mp = T")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3916
    apply (simp add: list_disj_def)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3917
    apply (cases "?pp = T")
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3918
    apply (simp add: list_disj_def)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3919
    apply (simp_all add: Let_def decr0[OF nb])
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3920
    done
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3921
  finally show ?thesis using decr0_qf[OF nb]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3922
    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
  3923
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3924
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3925
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
  3926
proof -
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3927
  from ferrack2
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3928
  have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3929
    by blast
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3930
  from qelim[OF th, of "prep p" bs]
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3931
  show ?thesis
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3932
    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
  3933
qed
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  3934
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  3935
oracle frpar_oracle = \<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3936
let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3937
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3938
fun binopT T = T --> T --> T;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3939
fun relT T = T --> T --> @{typ bool};
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3940
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58889
diff changeset
  3941
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
  3942
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
  3943
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
  3944
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3945
val dest_num = snd o HOLogic.dest_number;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3946
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3947
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
  3948
  handle TERM _ => NONE;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3949
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3950
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
  3951
  | dest_nat t = dest_num t;
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3952
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3953
fun the_index ts t =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3954
  let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3955
    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
  3956
  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
  3957
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3958
fun num_of_term ps (Const (@{const_name Groups.uminus}, _) $ t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3959
      @{code poly.Neg} (num_of_term ps t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3960
  | num_of_term ps (Const (@{const_name Groups.plus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3961
      @{code poly.Add} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3962
  | num_of_term ps (Const (@{const_name Groups.minus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3963
      @{code poly.Sub} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3964
  | num_of_term ps (Const (@{const_name Groups.times}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3965
      @{code poly.Mul} (num_of_term ps a, num_of_term ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3966
  | num_of_term ps (Const (@{const_name Power.power}, _) $ a $ n) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3967
      @{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
  3968
  | num_of_term ps (Const (@{const_name Rings.divide}, _) $ a $ b) =
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3969
      mk_C (dest_num a, dest_num b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3970
  | num_of_term ps t =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3971
      (case try_dest_num t of
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3972
        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
  3973
      | 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
  3974
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3975
fun tm_of_term fs ps (Const(@{const_name Groups.uminus}, _) $ t) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3976
      @{code Neg} (tm_of_term fs ps t)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3977
  | tm_of_term fs ps (Const(@{const_name Groups.plus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3978
      @{code Add} (tm_of_term fs ps a, tm_of_term fs ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3979
  | tm_of_term fs ps (Const(@{const_name Groups.minus}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3980
      @{code Sub} (tm_of_term fs ps a, tm_of_term fs ps b)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3981
  | tm_of_term fs ps (Const(@{const_name Groups.times}, _) $ a $ b) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3982
      @{code Mul} (num_of_term ps a, tm_of_term fs ps b)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  3983
  | 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
  3984
      handle TERM _ => mk_Bound (the_index fs t)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3985
        | 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
  3986
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  3987
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
  3988
  | fm_of_term fs ps @{term False} = @{code F}
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3989
  | fm_of_term fs ps (Const (@{const_name Not}, _) $ p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3990
      @{code NOT} (fm_of_term fs ps p)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3991
  | fm_of_term fs ps (Const (@{const_name HOL.conj}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3992
      @{code And} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3993
  | fm_of_term fs ps (Const (@{const_name HOL.disj}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3994
      @{code Or} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3995
  | fm_of_term fs ps (Const (@{const_name HOL.implies}, _) $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3996
      @{code Imp} (fm_of_term fs ps p, fm_of_term fs ps q)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3997
  | fm_of_term fs ps (@{term HOL.iff} $ p $ q) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  3998
      @{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
  3999
  | 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
  4000
      @{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
  4001
  | 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
  4002
      @{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
  4003
  | 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
  4004
      @{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
  4005
  | 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
  4006
      let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4007
        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
  4008
      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
  4009
  | 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
  4010
      let
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4011
        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
  4012
      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
  4013
  | 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
  4014
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4015
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
  4016
      let
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58889
diff changeset
  4017
        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
  4018
      in
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4019
        (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
  4020
        else if d = 0 then Const (@{const_name Groups.zero}, T)
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4021
        else
60352
d46de31a50c4 separate class for division operator, with particular syntax added in more specific classes
haftmann
parents: 59867
diff changeset
  4022
          Const (@{const_name Rings.divide}, binopT T) $
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4023
            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
  4024
      end
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50282
diff changeset
  4025
  | 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
  4026
  | term_of_num T ps (@{code poly.Add} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4027
      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
  4028
  | term_of_num T ps (@{code poly.Mul} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4029
      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
  4030
  | term_of_num T ps (@{code poly.Sub} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4031
      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
  4032
  | term_of_num T ps (@{code poly.Neg} a) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4033
      Const (@{const_name Groups.uminus}, T --> T) $ term_of_num T ps a
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4034
  | term_of_num T ps (@{code poly.Pw} (a, n)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4035
      Const (@{const_name Power.power}, T --> @{typ nat} --> T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4036
        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
  4037
  | 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
  4038
      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
  4039
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4040
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
  4041
  | 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
  4042
  | term_of_tm T fs ps (@{code Add} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4043
      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
  4044
  | term_of_tm T fs ps (@{code Mul} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4045
      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
  4046
  | term_of_tm T fs ps (@{code Sub} (a, b)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4047
      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
  4048
  | term_of_tm T fs ps (@{code Neg} a) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4049
      Const (@{const_name Groups.uminus}, T --> T) $ term_of_tm T fs ps a
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4050
  | term_of_tm T fs ps (@{code CNP} (n, c, p)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4051
      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
  4052
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4053
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
  4054
  | 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
  4055
  | 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
  4056
  | term_of_fm T fs ps (@{code And} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4057
      @{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
  4058
  | term_of_fm T fs ps (@{code Or} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4059
      @{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
  4060
  | term_of_fm T fs ps (@{code Imp} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4061
      @{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
  4062
  | term_of_fm T fs ps (@{code Iff} (p, q)) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4063
      @{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
  4064
  | term_of_fm T fs ps (@{code Lt} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4065
      Const (@{const_name Orderings.less}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4066
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4067
  | term_of_fm T fs ps (@{code Le} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4068
      Const (@{const_name Orderings.less_eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4069
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4070
  | term_of_fm T fs ps (@{code Eq} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4071
      Const (@{const_name HOL.eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4072
        term_of_tm T fs ps p $ Const (@{const_name Groups.zero}, T)
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4073
  | term_of_fm T fs ps (@{code NEq} p) =
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4074
      @{term Not} $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4075
        (Const (@{const_name HOL.eq}, relT T) $
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4076
          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
  4077
  | 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
  4078
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4079
fun frpar_procedure alternative T ps fm =
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4080
  let
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4081
    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
  4082
    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
  4083
    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
  4084
    val t = HOLogic.dest_Trueprop fm;
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4085
  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
  4086
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4087
in
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4088
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4089
  fn (ctxt, alternative, ty, ps, ct) =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
  4090
    Thm.cterm_of ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
  4091
      (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
  4092
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4093
end
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4094
\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4095
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4096
ML \<open>
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4097
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
  4098
struct
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4099
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4100
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
  4101
  Object_Logic.full_atomize_tac ctxt
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4102
  THEN' CSUBGOAL (fn (g, i) =>
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4103
    let
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4104
      val th = frpar_oracle (ctxt, alternative, T, ps, g);
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60567
diff changeset
  4105
    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
  4106
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4107
fun method alternative =
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4108
  let
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4109
    fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4110
    val parsN = "pars";
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4111
    val typN = "type";
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4112
    val any_keyword = keyword parsN || keyword typN;
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4113
    val terms = Scan.repeat (Scan.unless any_keyword Args.term);
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4114
    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
  4115
  in
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4116
    (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4117
      (fn (T, ps) => fn ctxt => SIMPLE_METHOD' (tactic ctxt alternative T ps))
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4118
  end;
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4119
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4120
end;
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4121
\<close>
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4122
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4123
method_setup frpar = \<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4124
  Parametric_Ferrante_Rackoff.method false
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4125
\<close> "parametric QE for linear Arithmetic over fields"
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4126
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4127
method_setup frpar2 = \<open>
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4128
  Parametric_Ferrante_Rackoff.method true
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 60352
diff changeset
  4129
\<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
  4130
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4131
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
  4132
  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
  4133
  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
  4134
  apply (rule spec[where x=y])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4135
  apply (frpar type: 'a pars: "z::'a")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4136
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4137
  done
33152
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4138
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4139
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
  4140
  apply (frpar2 type: 'a pars: y)
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4141
  apply (simp add: field_simps)
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4142
  apply (rule spec[where x=y])
55768
72c6ce5aea2a tuned specifications and proofs;
wenzelm
parents: 55754
diff changeset
  4143
  apply (frpar2 type: 'a pars: "z::'a")
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4144
  apply simp
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 55422
diff changeset
  4145
  done
50045
2214bc566f88 modernized, simplified and compacted oracle and proof method glue code;
haftmann
parents: 49962
diff changeset
  4146
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4147
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
  4148
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4149
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
  4150
proof -
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4151
  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
  4152
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
  4153
have "?rhs"
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
  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
  4156
  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
  4157
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4158
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4159
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4160
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
  4161
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
  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
60560
ce7030d9191d tuned proofs;
wenzelm
parents: 60533
diff changeset
  4165
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
  4166
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4167
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4168
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
  4169
proof -
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4170
  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
  4171
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
  4172
have "?rhs"
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4173
  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
  4174
  apply simp
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4175
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4176
*)
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4177
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4178
(*
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59621
diff changeset
  4179
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
  4180
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
  4181
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
  4182
apply ferrack
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4183
oops
241cfaed158f Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials)
chaieb
parents:
diff changeset
  4184
*)
45499
849d697adf1e Parametric_Ferrante_Rackoff.thy: restrict to class number_ring, replace '1+1' with '2' everywhere
huffman
parents: 44064
diff changeset
  4185
end