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