| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70092 | a19dd7006a3c | 
| child 74101 | d804e93ae9ff | 
| permissions | -rw-r--r-- | 
| 30439 | 1  | 
(* Title: HOL/Decision_Procs/Cooper.thy  | 
| 27456 | 2  | 
Author: Amine Chaieb  | 
3  | 
*)  | 
|
4  | 
||
| 70091 | 5  | 
section \<open>Presburger arithmetic based on Cooper's algorithm\<close>  | 
6  | 
||
| 29788 | 7  | 
theory Cooper  | 
| 55885 | 8  | 
imports  | 
9  | 
Complex_Main  | 
|
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
65024 
diff
changeset
 | 
10  | 
"HOL-Library.Code_Target_Numeral"  | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23315 
diff
changeset
 | 
11  | 
begin  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
12  | 
|
| 70091 | 13  | 
subsection \<open>Basic formulae\<close>  | 
| 23274 | 14  | 
|
| 66809 | 15  | 
datatype (plugins del: size) num = C int | Bound nat | CN nat int num  | 
16  | 
| Neg num | Add num num | Sub num num  | 
|
| 23274 | 17  | 
| Mul int num  | 
18  | 
||
| 66809 | 19  | 
instantiation num :: size  | 
20  | 
begin  | 
|
21  | 
||
22  | 
primrec size_num :: "num \<Rightarrow> nat"  | 
|
| 67123 | 23  | 
where  | 
24  | 
"size_num (C c) = 1"  | 
|
25  | 
| "size_num (Bound n) = 1"  | 
|
26  | 
| "size_num (Neg a) = 1 + size_num a"  | 
|
27  | 
| "size_num (Add a b) = 1 + size_num a + size_num b"  | 
|
28  | 
| "size_num (Sub a b) = 3 + size_num a + size_num b"  | 
|
29  | 
| "size_num (CN n c a) = 4 + size_num a"  | 
|
30  | 
| "size_num (Mul c a) = 1 + size_num a"  | 
|
| 66809 | 31  | 
|
32  | 
instance ..  | 
|
33  | 
||
34  | 
end  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
35  | 
|
| 55844 | 36  | 
primrec Inum :: "int list \<Rightarrow> num \<Rightarrow> int"  | 
| 67123 | 37  | 
where  | 
38  | 
"Inum bs (C c) = c"  | 
|
39  | 
| "Inum bs (Bound n) = bs ! n"  | 
|
40  | 
| "Inum bs (CN n c a) = c * (bs ! n) + Inum bs a"  | 
|
41  | 
| "Inum bs (Neg a) = - Inum bs a"  | 
|
42  | 
| "Inum bs (Add a b) = Inum bs a + Inum bs b"  | 
|
43  | 
| "Inum bs (Sub a b) = Inum bs a - Inum bs b"  | 
|
44  | 
| "Inum bs (Mul c a) = c * Inum bs a"  | 
|
| 23274 | 45  | 
|
| 66809 | 46  | 
datatype (plugins del: size) fm = T | F  | 
47  | 
| Lt num | Le num | Gt num | Ge num | Eq num | NEq num  | 
|
48  | 
| Dvd int num | NDvd int num  | 
|
49  | 
| NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm  | 
|
50  | 
| E fm | A fm | Closed nat | NClosed nat  | 
|
| 23274 | 51  | 
|
| 66809 | 52  | 
instantiation fm :: size  | 
53  | 
begin  | 
|
| 23274 | 54  | 
|
| 66809 | 55  | 
primrec size_fm :: "fm \<Rightarrow> nat"  | 
| 67123 | 56  | 
where  | 
57  | 
"size_fm (NOT p) = 1 + size_fm p"  | 
|
58  | 
| "size_fm (And p q) = 1 + size_fm p + size_fm q"  | 
|
59  | 
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"  | 
|
60  | 
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"  | 
|
61  | 
| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"  | 
|
62  | 
| "size_fm (E p) = 1 + size_fm p"  | 
|
63  | 
| "size_fm (A p) = 4 + size_fm p"  | 
|
64  | 
| "size_fm (Dvd i t) = 2"  | 
|
65  | 
| "size_fm (NDvd i t) = 2"  | 
|
66  | 
| "size_fm T = 1"  | 
|
67  | 
| "size_fm F = 1"  | 
|
68  | 
| "size_fm (Lt _) = 1"  | 
|
69  | 
| "size_fm (Le _) = 1"  | 
|
70  | 
| "size_fm (Gt _) = 1"  | 
|
71  | 
| "size_fm (Ge _) = 1"  | 
|
72  | 
| "size_fm (Eq _) = 1"  | 
|
73  | 
| "size_fm (NEq _) = 1"  | 
|
74  | 
| "size_fm (Closed _) = 1"  | 
|
75  | 
| "size_fm (NClosed _) = 1"  | 
|
| 50313 | 76  | 
|
| 66809 | 77  | 
instance ..  | 
78  | 
||
79  | 
end  | 
|
80  | 
||
| 67123 | 81  | 
lemma fmsize_pos [simp]: "size p > 0"  | 
82  | 
for p :: fm  | 
|
| 66809 | 83  | 
by (induct p) simp_all  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
84  | 
|
| 67123 | 85  | 
primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool" \<comment> \<open>Semantics of formulae (\<open>fm\<close>)\<close>  | 
86  | 
where  | 
|
87  | 
"Ifm bbs bs T \<longleftrightarrow> True"  | 
|
88  | 
| "Ifm bbs bs F \<longleftrightarrow> False"  | 
|
89  | 
| "Ifm bbs bs (Lt a) \<longleftrightarrow> Inum bs a < 0"  | 
|
90  | 
| "Ifm bbs bs (Gt a) \<longleftrightarrow> Inum bs a > 0"  | 
|
91  | 
| "Ifm bbs bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0"  | 
|
92  | 
| "Ifm bbs bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0"  | 
|
93  | 
| "Ifm bbs bs (Eq a) \<longleftrightarrow> Inum bs a = 0"  | 
|
94  | 
| "Ifm bbs bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"  | 
|
95  | 
| "Ifm bbs bs (Dvd i b) \<longleftrightarrow> i dvd Inum bs b"  | 
|
96  | 
| "Ifm bbs bs (NDvd i b) \<longleftrightarrow> \<not> i dvd Inum bs b"  | 
|
97  | 
| "Ifm bbs bs (NOT p) \<longleftrightarrow> \<not> Ifm bbs bs p"  | 
|
98  | 
| "Ifm bbs bs (And p q) \<longleftrightarrow> Ifm bbs bs p \<and> Ifm bbs bs q"  | 
|
99  | 
| "Ifm bbs bs (Or p q) \<longleftrightarrow> Ifm bbs bs p \<or> Ifm bbs bs q"  | 
|
100  | 
| "Ifm bbs bs (Imp p q) \<longleftrightarrow> (Ifm bbs bs p \<longrightarrow> Ifm bbs bs q)"  | 
|
101  | 
| "Ifm bbs bs (Iff p q) \<longleftrightarrow> Ifm bbs bs p = Ifm bbs bs q"  | 
|
102  | 
| "Ifm bbs bs (E p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) p)"  | 
|
103  | 
| "Ifm bbs bs (A p) \<longleftrightarrow> (\<forall>x. Ifm bbs (x # bs) p)"  | 
|
104  | 
| "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs ! n"  | 
|
105  | 
| "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs ! n"  | 
|
| 23274 | 106  | 
|
| 66809 | 107  | 
fun prep :: "fm \<Rightarrow> fm"  | 
| 67123 | 108  | 
where  | 
109  | 
"prep (E T) = T"  | 
|
110  | 
| "prep (E F) = F"  | 
|
111  | 
| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"  | 
|
112  | 
| "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"  | 
|
113  | 
| "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"  | 
|
114  | 
| "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"  | 
|
115  | 
| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"  | 
|
116  | 
| "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"  | 
|
117  | 
| "prep (E p) = E (prep p)"  | 
|
118  | 
| "prep (A (And p q)) = And (prep (A p)) (prep (A q))"  | 
|
119  | 
| "prep (A p) = prep (NOT (E (NOT p)))"  | 
|
120  | 
| "prep (NOT (NOT p)) = prep p"  | 
|
121  | 
| "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"  | 
|
122  | 
| "prep (NOT (A p)) = prep (E (NOT p))"  | 
|
123  | 
| "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"  | 
|
124  | 
| "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"  | 
|
125  | 
| "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"  | 
|
126  | 
| "prep (NOT p) = NOT (prep p)"  | 
|
127  | 
| "prep (Or p q) = Or (prep p) (prep q)"  | 
|
128  | 
| "prep (And p q) = And (prep p) (prep q)"  | 
|
129  | 
| "prep (Imp p q) = prep (Or (NOT p) q)"  | 
|
130  | 
| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"  | 
|
131  | 
| "prep p = p"  | 
|
| 50313 | 132  | 
|
| 23274 | 133  | 
lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p"  | 
| 50313 | 134  | 
by (induct p arbitrary: bs rule: prep.induct) auto  | 
| 23274 | 135  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
136  | 
|
| 61586 | 137  | 
fun qfree :: "fm \<Rightarrow> bool" \<comment> \<open>Quantifier freeness\<close>  | 
| 67123 | 138  | 
where  | 
139  | 
"qfree (E p) \<longleftrightarrow> False"  | 
|
140  | 
| "qfree (A p) \<longleftrightarrow> False"  | 
|
141  | 
| "qfree (NOT p) \<longleftrightarrow> qfree p"  | 
|
142  | 
| "qfree (And p q) \<longleftrightarrow> qfree p \<and> qfree q"  | 
|
143  | 
| "qfree (Or p q) \<longleftrightarrow> qfree p \<and> qfree q"  | 
|
144  | 
| "qfree (Imp p q) \<longleftrightarrow> qfree p \<and> qfree q"  | 
|
145  | 
| "qfree (Iff p q) \<longleftrightarrow> qfree p \<and> qfree q"  | 
|
146  | 
| "qfree p \<longleftrightarrow> True"  | 
|
| 23274 | 147  | 
|
| 50313 | 148  | 
|
| 70091 | 149  | 
subsection \<open>Boundedness and substitution\<close>  | 
| 50313 | 150  | 
|
| 67123 | 151  | 
primrec numbound0 :: "num \<Rightarrow> bool" \<comment> \<open>a \<open>num\<close> is \<^emph>\<open>independent\<close> of Bound 0\<close>  | 
152  | 
where  | 
|
153  | 
"numbound0 (C c) \<longleftrightarrow> True"  | 
|
154  | 
| "numbound0 (Bound n) \<longleftrightarrow> n > 0"  | 
|
155  | 
| "numbound0 (CN n i a) \<longleftrightarrow> n > 0 \<and> numbound0 a"  | 
|
156  | 
| "numbound0 (Neg a) \<longleftrightarrow> numbound0 a"  | 
|
157  | 
| "numbound0 (Add a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"  | 
|
158  | 
| "numbound0 (Sub a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b"  | 
|
159  | 
| "numbound0 (Mul i a) \<longleftrightarrow> numbound0 a"  | 
|
| 23274 | 160  | 
|
161  | 
lemma numbound0_I:  | 
|
| 67123 | 162  | 
assumes "numbound0 a"  | 
| 55981 | 163  | 
shows "Inum (b # bs) a = Inum (b' # bs) a"  | 
| 67123 | 164  | 
using assms by (induct a rule: num.induct) (auto simp add: gr0_conv_Suc)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
165  | 
|
| 67123 | 166  | 
primrec bound0 :: "fm \<Rightarrow> bool" \<comment> \<open>a formula is independent of Bound 0\<close>  | 
167  | 
where  | 
|
168  | 
"bound0 T \<longleftrightarrow> True"  | 
|
169  | 
| "bound0 F \<longleftrightarrow> True"  | 
|
170  | 
| "bound0 (Lt a) \<longleftrightarrow> numbound0 a"  | 
|
171  | 
| "bound0 (Le a) \<longleftrightarrow> numbound0 a"  | 
|
172  | 
| "bound0 (Gt a) \<longleftrightarrow> numbound0 a"  | 
|
173  | 
| "bound0 (Ge a) \<longleftrightarrow> numbound0 a"  | 
|
174  | 
| "bound0 (Eq a) \<longleftrightarrow> numbound0 a"  | 
|
175  | 
| "bound0 (NEq a) \<longleftrightarrow> numbound0 a"  | 
|
176  | 
| "bound0 (Dvd i a) \<longleftrightarrow> numbound0 a"  | 
|
177  | 
| "bound0 (NDvd i a) \<longleftrightarrow> numbound0 a"  | 
|
178  | 
| "bound0 (NOT p) \<longleftrightarrow> bound0 p"  | 
|
179  | 
| "bound0 (And p q) \<longleftrightarrow> bound0 p \<and> bound0 q"  | 
|
180  | 
| "bound0 (Or p q) \<longleftrightarrow> bound0 p \<and> bound0 q"  | 
|
181  | 
| "bound0 (Imp p q) \<longleftrightarrow> bound0 p \<and> bound0 q"  | 
|
182  | 
| "bound0 (Iff p q) \<longleftrightarrow> bound0 p \<and> bound0 q"  | 
|
183  | 
| "bound0 (E p) \<longleftrightarrow> False"  | 
|
184  | 
| "bound0 (A p) \<longleftrightarrow> False"  | 
|
185  | 
| "bound0 (Closed P) \<longleftrightarrow> True"  | 
|
186  | 
| "bound0 (NClosed P) \<longleftrightarrow> True"  | 
|
| 50313 | 187  | 
|
| 23274 | 188  | 
lemma bound0_I:  | 
| 67123 | 189  | 
assumes "bound0 p"  | 
| 55981 | 190  | 
shows "Ifm bbs (b # bs) p = Ifm bbs (b' # bs) p"  | 
| 67123 | 191  | 
using assms numbound0_I[where b="b" and bs="bs" and b'="b'"]  | 
| 50313 | 192  | 
by (induct p rule: fm.induct) (auto simp add: gr0_conv_Suc)  | 
| 23274 | 193  | 
|
| 50313 | 194  | 
fun numsubst0 :: "num \<Rightarrow> num \<Rightarrow> num"  | 
| 67123 | 195  | 
where  | 
196  | 
"numsubst0 t (C c) = (C c)"  | 
|
197  | 
| "numsubst0 t (Bound n) = (if n = 0 then t else Bound n)"  | 
|
198  | 
| "numsubst0 t (CN 0 i a) = Add (Mul i t) (numsubst0 t a)"  | 
|
199  | 
| "numsubst0 t (CN n i a) = CN n i (numsubst0 t a)"  | 
|
200  | 
| "numsubst0 t (Neg a) = Neg (numsubst0 t a)"  | 
|
201  | 
| "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"  | 
|
202  | 
| "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"  | 
|
203  | 
| "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"  | 
|
| 23274 | 204  | 
|
| 67123 | 205  | 
lemma numsubst0_I: "Inum (b # bs) (numsubst0 a t) = Inum ((Inum (b # bs) a) # bs) t"  | 
| 50313 | 206  | 
by (induct t rule: numsubst0.induct) (auto simp: nth_Cons')  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
207  | 
|
| 67123 | 208  | 
lemma numsubst0_I': "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"  | 
| 50313 | 209  | 
by (induct t rule: numsubst0.induct) (auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"])  | 
| 23274 | 210  | 
|
| 67123 | 211  | 
primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" \<comment> \<open>substitute a \<open>num\<close> into a formula for Bound 0\<close>  | 
212  | 
where  | 
|
213  | 
"subst0 t T = T"  | 
|
214  | 
| "subst0 t F = F"  | 
|
215  | 
| "subst0 t (Lt a) = Lt (numsubst0 t a)"  | 
|
216  | 
| "subst0 t (Le a) = Le (numsubst0 t a)"  | 
|
217  | 
| "subst0 t (Gt a) = Gt (numsubst0 t a)"  | 
|
218  | 
| "subst0 t (Ge a) = Ge (numsubst0 t a)"  | 
|
219  | 
| "subst0 t (Eq a) = Eq (numsubst0 t a)"  | 
|
220  | 
| "subst0 t (NEq a) = NEq (numsubst0 t a)"  | 
|
221  | 
| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"  | 
|
222  | 
| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"  | 
|
223  | 
| "subst0 t (NOT p) = NOT (subst0 t p)"  | 
|
224  | 
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"  | 
|
225  | 
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"  | 
|
226  | 
| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"  | 
|
227  | 
| "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"  | 
|
228  | 
| "subst0 t (Closed P) = (Closed P)"  | 
|
229  | 
| "subst0 t (NClosed P) = (NClosed P)"  | 
|
| 23274 | 230  | 
|
| 50313 | 231  | 
lemma subst0_I:  | 
| 55999 | 232  | 
assumes "qfree p"  | 
| 55885 | 233  | 
shows "Ifm bbs (b # bs) (subst0 a p) = Ifm bbs (Inum (b # bs) a # bs) p"  | 
| 55999 | 234  | 
using assms numsubst0_I[where b="b" and bs="bs" and a="a"]  | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
235  | 
by (induct p) (simp_all add: gr0_conv_Suc)  | 
| 23274 | 236  | 
|
| 50313 | 237  | 
fun decrnum:: "num \<Rightarrow> num"  | 
| 67123 | 238  | 
where  | 
239  | 
"decrnum (Bound n) = Bound (n - 1)"  | 
|
240  | 
| "decrnum (Neg a) = Neg (decrnum a)"  | 
|
241  | 
| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"  | 
|
242  | 
| "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"  | 
|
243  | 
| "decrnum (Mul c a) = Mul c (decrnum a)"  | 
|
244  | 
| "decrnum (CN n i a) = (CN (n - 1) i (decrnum a))"  | 
|
245  | 
| "decrnum a = a"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
246  | 
|
| 50313 | 247  | 
fun decr :: "fm \<Rightarrow> fm"  | 
| 67123 | 248  | 
where  | 
249  | 
"decr (Lt a) = Lt (decrnum a)"  | 
|
250  | 
| "decr (Le a) = Le (decrnum a)"  | 
|
251  | 
| "decr (Gt a) = Gt (decrnum a)"  | 
|
252  | 
| "decr (Ge a) = Ge (decrnum a)"  | 
|
253  | 
| "decr (Eq a) = Eq (decrnum a)"  | 
|
254  | 
| "decr (NEq a) = NEq (decrnum a)"  | 
|
255  | 
| "decr (Dvd i a) = Dvd i (decrnum a)"  | 
|
256  | 
| "decr (NDvd i a) = NDvd i (decrnum a)"  | 
|
257  | 
| "decr (NOT p) = NOT (decr p)"  | 
|
258  | 
| "decr (And p q) = And (decr p) (decr q)"  | 
|
259  | 
| "decr (Or p q) = Or (decr p) (decr q)"  | 
|
260  | 
| "decr (Imp p q) = Imp (decr p) (decr q)"  | 
|
261  | 
| "decr (Iff p q) = Iff (decr p) (decr q)"  | 
|
262  | 
| "decr p = p"  | 
|
| 23274 | 263  | 
|
| 50313 | 264  | 
lemma decrnum:  | 
| 67123 | 265  | 
assumes "numbound0 t"  | 
| 55885 | 266  | 
shows "Inum (x # bs) t = Inum bs (decrnum t)"  | 
| 67123 | 267  | 
using assms by (induct t rule: decrnum.induct) (auto simp add: gr0_conv_Suc)  | 
| 23274 | 268  | 
|
| 50313 | 269  | 
lemma decr:  | 
| 67123 | 270  | 
assumes assms: "bound0 p"  | 
| 55885 | 271  | 
shows "Ifm bbs (x # bs) p = Ifm bbs bs (decr p)"  | 
| 67123 | 272  | 
using assms by (induct p rule: decr.induct) (simp_all add: gr0_conv_Suc decrnum)  | 
| 23274 | 273  | 
|
274  | 
lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"  | 
|
| 50313 | 275  | 
by (induct p) simp_all  | 
| 23274 | 276  | 
|
| 61586 | 277  | 
fun isatom :: "fm \<Rightarrow> bool" \<comment> \<open>test for atomicity\<close>  | 
| 67123 | 278  | 
where  | 
279  | 
"isatom T \<longleftrightarrow> True"  | 
|
280  | 
| "isatom F \<longleftrightarrow> True"  | 
|
281  | 
| "isatom (Lt a) \<longleftrightarrow> True"  | 
|
282  | 
| "isatom (Le a) \<longleftrightarrow> True"  | 
|
283  | 
| "isatom (Gt a) \<longleftrightarrow> True"  | 
|
284  | 
| "isatom (Ge a) \<longleftrightarrow> True"  | 
|
285  | 
| "isatom (Eq a) \<longleftrightarrow> True"  | 
|
286  | 
| "isatom (NEq a) \<longleftrightarrow> True"  | 
|
287  | 
| "isatom (Dvd i b) \<longleftrightarrow> True"  | 
|
288  | 
| "isatom (NDvd i b) \<longleftrightarrow> True"  | 
|
289  | 
| "isatom (Closed P) \<longleftrightarrow> True"  | 
|
290  | 
| "isatom (NClosed P) \<longleftrightarrow> True"  | 
|
291  | 
| "isatom p \<longleftrightarrow> False"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
292  | 
|
| 50313 | 293  | 
lemma numsubst0_numbound0:  | 
| 55844 | 294  | 
assumes "numbound0 t"  | 
| 23274 | 295  | 
shows "numbound0 (numsubst0 t a)"  | 
| 55844 | 296  | 
using assms  | 
| 60708 | 297  | 
proof (induct a)  | 
| 67123 | 298  | 
case (CN n)  | 
| 60708 | 299  | 
then show ?case by (cases n) simp_all  | 
300  | 
qed simp_all  | 
|
| 23274 | 301  | 
|
| 50313 | 302  | 
lemma subst0_bound0:  | 
| 55844 | 303  | 
assumes qf: "qfree p"  | 
304  | 
and nb: "numbound0 t"  | 
|
| 23274 | 305  | 
shows "bound0 (subst0 t p)"  | 
| 50313 | 306  | 
using qf numsubst0_numbound0[OF nb] by (induct p) auto  | 
| 23274 | 307  | 
|
308  | 
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"  | 
|
| 50313 | 309  | 
by (induct p) simp_all  | 
| 23274 | 310  | 
|
311  | 
||
| 50313 | 312  | 
definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
 | 
313  | 
where  | 
|
314  | 
"djf f p q =  | 
|
| 55885 | 315  | 
(if q = T then T  | 
316  | 
else if q = F then f p  | 
|
317  | 
else  | 
|
318  | 
let fp = f p  | 
|
319  | 
in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q)"  | 
|
| 50313 | 320  | 
|
321  | 
definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
 | 
|
322  | 
where "evaldjf f ps = foldr (djf f) ps F"  | 
|
| 23274 | 323  | 
|
324  | 
lemma djf_Or: "Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)"  | 
|
| 55885 | 325  | 
by (cases "q=T", simp add: djf_def, cases "q = F", simp add: djf_def)  | 
| 50313 | 326  | 
(cases "f p", simp_all add: Let_def djf_def)  | 
| 23274 | 327  | 
|
| 55885 | 328  | 
lemma evaldjf_ex: "Ifm bbs bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm bbs bs (f p))"  | 
| 50313 | 329  | 
by (induct ps) (simp_all add: evaldjf_def djf_Or)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
330  | 
|
| 50313 | 331  | 
lemma evaldjf_bound0:  | 
332  | 
assumes nb: "\<forall>x\<in> set xs. bound0 (f x)"  | 
|
| 23274 | 333  | 
shows "bound0 (evaldjf f xs)"  | 
| 55422 | 334  | 
using nb by (induct xs) (auto simp add: evaldjf_def djf_def Let_def, case_tac "f a", auto)  | 
| 23274 | 335  | 
|
| 50313 | 336  | 
lemma evaldjf_qf:  | 
337  | 
assumes nb: "\<forall>x\<in> set xs. qfree (f x)"  | 
|
| 23274 | 338  | 
shows "qfree (evaldjf f xs)"  | 
| 55422 | 339  | 
using nb by (induct xs) (auto simp add: evaldjf_def djf_def Let_def, case_tac "f a", auto)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
340  | 
|
| 50313 | 341  | 
fun disjuncts :: "fm \<Rightarrow> fm list"  | 
| 67123 | 342  | 
where  | 
343  | 
"disjuncts (Or p q) = disjuncts p @ disjuncts q"  | 
|
344  | 
| "disjuncts F = []"  | 
|
345  | 
| "disjuncts p = [p]"  | 
|
| 23274 | 346  | 
|
| 55885 | 347  | 
lemma disjuncts: "(\<exists>q \<in> set (disjuncts p). Ifm bbs bs q) \<longleftrightarrow> Ifm bbs bs p"  | 
348  | 
by (induct p rule: disjuncts.induct) auto  | 
|
| 23274 | 349  | 
|
| 50313 | 350  | 
lemma disjuncts_nb:  | 
| 55999 | 351  | 
assumes "bound0 p"  | 
| 50313 | 352  | 
shows "\<forall>q \<in> set (disjuncts p). bound0 q"  | 
353  | 
proof -  | 
|
| 55999 | 354  | 
from assms have "list_all bound0 (disjuncts p)"  | 
| 50313 | 355  | 
by (induct p rule: disjuncts.induct) auto  | 
| 55999 | 356  | 
then show ?thesis  | 
357  | 
by (simp only: list_all_iff)  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
358  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
359  | 
|
| 50313 | 360  | 
lemma disjuncts_qf:  | 
| 55999 | 361  | 
assumes "qfree p"  | 
| 50313 | 362  | 
shows "\<forall>q \<in> set (disjuncts p). qfree q"  | 
363  | 
proof -  | 
|
| 55999 | 364  | 
from assms have "list_all qfree (disjuncts p)"  | 
| 50313 | 365  | 
by (induct p rule: disjuncts.induct) auto  | 
| 55885 | 366  | 
then show ?thesis by (simp only: list_all_iff)  | 
| 23274 | 367  | 
qed  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
368  | 
|
| 50313 | 369  | 
definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"  | 
370  | 
where "DJ f p = evaldjf f (disjuncts p)"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
371  | 
|
| 50313 | 372  | 
lemma DJ:  | 
| 55999 | 373  | 
assumes "\<forall>p q. f (Or p q) = Or (f p) (f q)"  | 
374  | 
and "f F = F"  | 
|
| 23274 | 375  | 
shows "Ifm bbs bs (DJ f p) = Ifm bbs bs (f p)"  | 
| 50313 | 376  | 
proof -  | 
| 55999 | 377  | 
have "Ifm bbs bs (DJ f p) \<longleftrightarrow> (\<exists>q \<in> set (disjuncts p). Ifm bbs bs (f q))"  | 
| 50313 | 378  | 
by (simp add: DJ_def evaldjf_ex)  | 
| 55999 | 379  | 
also from assms have "\<dots> = Ifm bbs bs (f p)"  | 
380  | 
by (induct p rule: disjuncts.induct) auto  | 
|
| 23274 | 381  | 
finally show ?thesis .  | 
382  | 
qed  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
383  | 
|
| 50313 | 384  | 
lemma DJ_qf:  | 
| 55999 | 385  | 
assumes "\<forall>p. qfree p \<longrightarrow> qfree (f p)"  | 
| 23274 | 386  | 
shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "  | 
| 50313 | 387  | 
proof clarify  | 
| 55844 | 388  | 
fix p  | 
389  | 
assume qf: "qfree p"  | 
|
390  | 
have th: "DJ f p = evaldjf f (disjuncts p)"  | 
|
391  | 
by (simp add: DJ_def)  | 
|
| 55925 | 392  | 
from disjuncts_qf[OF qf] have "\<forall>q \<in> set (disjuncts p). qfree q" .  | 
| 55999 | 393  | 
with assms have th': "\<forall>q \<in> set (disjuncts p). qfree (f q)"  | 
| 55844 | 394  | 
by blast  | 
395  | 
from evaldjf_qf[OF th'] th show "qfree (DJ f p)"  | 
|
396  | 
by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
397  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
398  | 
|
| 50313 | 399  | 
lemma DJ_qe:  | 
| 55885 | 400  | 
assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)"  | 
401  | 
shows "\<forall>bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> Ifm bbs bs ((DJ qe p)) = Ifm bbs bs (E p)"  | 
|
| 50313 | 402  | 
proof clarify  | 
| 55844 | 403  | 
fix p :: fm  | 
404  | 
fix bs  | 
|
| 23274 | 405  | 
assume qf: "qfree p"  | 
| 55844 | 406  | 
from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)"  | 
407  | 
by blast  | 
|
| 55925 | 408  | 
from DJ_qf[OF qth] qf have qfth: "qfree (DJ qe p)"  | 
| 55844 | 409  | 
by auto  | 
| 50313 | 410  | 
have "Ifm bbs bs (DJ qe p) = (\<exists>q\<in> set (disjuncts p). Ifm bbs bs (qe q))"  | 
| 23274 | 411  | 
by (simp add: DJ_def evaldjf_ex)  | 
| 55999 | 412  | 
also have "\<dots> \<longleftrightarrow> (\<exists>q \<in> set (disjuncts p). Ifm bbs bs (E q))"  | 
| 50313 | 413  | 
using qe disjuncts_qf[OF qf] by auto  | 
| 55925 | 414  | 
also have "\<dots> \<longleftrightarrow> Ifm bbs bs (E p)"  | 
| 50313 | 415  | 
by (induct p rule: disjuncts.induct) auto  | 
416  | 
finally show "qfree (DJ qe p) \<and> Ifm bbs bs (DJ qe p) = Ifm bbs bs (E p)"  | 
|
417  | 
using qfth by blast  | 
|
| 23274 | 418  | 
qed  | 
| 50313 | 419  | 
|
420  | 
||
| 70091 | 421  | 
subsection \<open>Simplification\<close>  | 
| 23274 | 422  | 
|
| 60533 | 423  | 
text \<open>Algebraic simplifications for nums\<close>  | 
| 41837 | 424  | 
|
| 50313 | 425  | 
fun bnds :: "num \<Rightarrow> nat list"  | 
| 67123 | 426  | 
where  | 
427  | 
"bnds (Bound n) = [n]"  | 
|
428  | 
| "bnds (CN n c a) = n # bnds a"  | 
|
429  | 
| "bnds (Neg a) = bnds a"  | 
|
430  | 
| "bnds (Add a b) = bnds a @ bnds b"  | 
|
431  | 
| "bnds (Sub a b) = bnds a @ bnds b"  | 
|
432  | 
| "bnds (Mul i a) = bnds a"  | 
|
433  | 
| "bnds a = []"  | 
|
| 41837 | 434  | 
|
| 50313 | 435  | 
fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"  | 
| 67123 | 436  | 
where  | 
437  | 
"lex_ns [] ms \<longleftrightarrow> True"  | 
|
438  | 
| "lex_ns ns [] \<longleftrightarrow> False"  | 
|
439  | 
| "lex_ns (n # ns) (m # ms) \<longleftrightarrow> n < m \<or> (n = m \<and> lex_ns ns ms)"  | 
|
| 23274 | 440  | 
|
| 50313 | 441  | 
definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"  | 
442  | 
where "lex_bnd t s = lex_ns (bnds t) (bnds s)"  | 
|
443  | 
||
| 66809 | 444  | 
fun numadd:: "num \<Rightarrow> num \<Rightarrow> num"  | 
| 67123 | 445  | 
where  | 
446  | 
"numadd (CN n1 c1 r1) (CN n2 c2 r2) =  | 
|
447  | 
(if n1 = n2 then  | 
|
448  | 
let c = c1 + c2  | 
|
449  | 
in if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)  | 
|
450  | 
else if n1 \<le> n2 then CN n1 c1 (numadd r1 (Add (Mul c2 (Bound n2)) r2))  | 
|
451  | 
else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1) r2))"  | 
|
452  | 
| "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"  | 
|
453  | 
| "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"  | 
|
454  | 
| "numadd (C b1) (C b2) = C (b1 + b2)"  | 
|
455  | 
| "numadd a b = Add a b"  | 
|
| 23274 | 456  | 
|
| 66809 | 457  | 
lemma numadd: "Inum bs (numadd t s) = Inum bs (Add t s)"  | 
458  | 
by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)  | 
|
| 23274 | 459  | 
|
| 66809 | 460  | 
lemma numadd_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd t s)"  | 
461  | 
by (induct t s rule: numadd.induct) (simp_all add: Let_def)  | 
|
| 23274 | 462  | 
|
| 50313 | 463  | 
fun nummul :: "int \<Rightarrow> num \<Rightarrow> num"  | 
| 67123 | 464  | 
where  | 
465  | 
"nummul i (C j) = C (i * j)"  | 
|
466  | 
| "nummul i (CN n c t) = CN n (c * i) (nummul i t)"  | 
|
467  | 
| "nummul i t = Mul i t"  | 
|
| 23274 | 468  | 
|
| 50313 | 469  | 
lemma nummul: "Inum bs (nummul i t) = Inum bs (Mul i t)"  | 
| 66809 | 470  | 
by (induct t arbitrary: i rule: nummul.induct) (simp_all add: algebra_simps)  | 
| 23274 | 471  | 
|
| 50313 | 472  | 
lemma nummul_nb: "numbound0 t \<Longrightarrow> numbound0 (nummul i t)"  | 
| 66809 | 473  | 
by (induct t arbitrary: i rule: nummul.induct) (simp_all add: numadd_nb)  | 
| 23274 | 474  | 
|
| 50313 | 475  | 
definition numneg :: "num \<Rightarrow> num"  | 
476  | 
where "numneg t = nummul (- 1) t"  | 
|
| 23274 | 477  | 
|
| 50313 | 478  | 
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"  | 
| 66809 | 479  | 
where "numsub s t = (if s = t then C 0 else numadd s (numneg t))"  | 
| 23274 | 480  | 
|
481  | 
lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)"  | 
|
| 50313 | 482  | 
using numneg_def nummul by simp  | 
| 23274 | 483  | 
|
484  | 
lemma numneg_nb: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"  | 
|
| 50313 | 485  | 
using numneg_def nummul_nb by simp  | 
| 23274 | 486  | 
|
487  | 
lemma numsub: "Inum bs (numsub a b) = Inum bs (Sub a b)"  | 
|
| 50313 | 488  | 
using numneg numadd numsub_def by simp  | 
| 23274 | 489  | 
|
| 50313 | 490  | 
lemma numsub_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numsub t s)"  | 
491  | 
using numsub_def numadd_nb numneg_nb by simp  | 
|
| 23274 | 492  | 
|
| 50313 | 493  | 
fun simpnum :: "num \<Rightarrow> num"  | 
| 67123 | 494  | 
where  | 
495  | 
"simpnum (C j) = C j"  | 
|
496  | 
| "simpnum (Bound n) = CN n 1 (C 0)"  | 
|
497  | 
| "simpnum (Neg t) = numneg (simpnum t)"  | 
|
498  | 
| "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"  | 
|
499  | 
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"  | 
|
500  | 
| "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))"  | 
|
501  | 
| "simpnum t = t"  | 
|
| 23274 | 502  | 
|
503  | 
lemma simpnum_ci: "Inum bs (simpnum t) = Inum bs t"  | 
|
| 50313 | 504  | 
by (induct t rule: simpnum.induct) (auto simp add: numneg numadd numsub nummul)  | 
| 23274 | 505  | 
|
| 50313 | 506  | 
lemma simpnum_numbound0: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"  | 
507  | 
by (induct t rule: simpnum.induct) (auto simp add: numadd_nb numsub_nb nummul_nb numneg_nb)  | 
|
| 23274 | 508  | 
|
| 50313 | 509  | 
fun not :: "fm \<Rightarrow> fm"  | 
| 67123 | 510  | 
where  | 
511  | 
"not (NOT p) = p"  | 
|
512  | 
| "not T = F"  | 
|
513  | 
| "not F = T"  | 
|
514  | 
| "not p = NOT p"  | 
|
| 50313 | 515  | 
|
| 23274 | 516  | 
lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)"  | 
| 41807 | 517  | 
by (cases p) auto  | 
| 50313 | 518  | 
|
| 23274 | 519  | 
lemma not_qf: "qfree p \<Longrightarrow> qfree (not p)"  | 
| 41807 | 520  | 
by (cases p) auto  | 
| 50313 | 521  | 
|
| 23274 | 522  | 
lemma not_bn: "bound0 p \<Longrightarrow> bound0 (not p)"  | 
| 41807 | 523  | 
by (cases p) auto  | 
| 23274 | 524  | 
|
| 50313 | 525  | 
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"  | 
| 67123 | 526  | 
where "conj p q =  | 
| 60708 | 527  | 
(if p = F \<or> q = F then F  | 
528  | 
else if p = T then q  | 
|
529  | 
else if q = T then p  | 
|
530  | 
else And p q)"  | 
|
| 50313 | 531  | 
|
| 23274 | 532  | 
lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)"  | 
| 55844 | 533  | 
by (cases "p = F \<or> q = F", simp_all add: conj_def) (cases p, simp_all)  | 
| 23274 | 534  | 
|
| 50313 | 535  | 
lemma conj_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (conj p q)"  | 
536  | 
using conj_def by auto  | 
|
| 23274 | 537  | 
|
| 50313 | 538  | 
lemma conj_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (conj p q)"  | 
539  | 
using conj_def by auto  | 
|
540  | 
||
541  | 
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"  | 
|
| 67123 | 542  | 
where "disj p q =  | 
| 60708 | 543  | 
(if p = T \<or> q = T then T  | 
544  | 
else if p = F then q  | 
|
545  | 
else if q = F then p  | 
|
546  | 
else Or p q)"  | 
|
| 23274 | 547  | 
|
548  | 
lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)"  | 
|
| 55885 | 549  | 
by (cases "p = T \<or> q = T", simp_all add: disj_def) (cases p, simp_all)  | 
| 50313 | 550  | 
|
| 55844 | 551  | 
lemma disj_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (disj p q)"  | 
| 50313 | 552  | 
using disj_def by auto  | 
553  | 
||
| 55844 | 554  | 
lemma disj_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (disj p q)"  | 
| 50313 | 555  | 
using disj_def by auto  | 
| 23274 | 556  | 
|
| 50313 | 557  | 
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"  | 
| 67123 | 558  | 
where "imp p q =  | 
| 60708 | 559  | 
(if p = F \<or> q = T then T  | 
560  | 
else if p = T then q  | 
|
561  | 
else if q = F then not p  | 
|
562  | 
else Imp p q)"  | 
|
| 50313 | 563  | 
|
| 23274 | 564  | 
lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)"  | 
| 55844 | 565  | 
by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) (simp_all add: not)  | 
| 50313 | 566  | 
|
567  | 
lemma imp_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (imp p q)"  | 
|
| 55844 | 568  | 
using imp_def by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) (simp_all add: not_qf)  | 
| 50313 | 569  | 
|
570  | 
lemma imp_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (imp p q)"  | 
|
| 55844 | 571  | 
using imp_def by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) simp_all  | 
| 23274 | 572  | 
|
| 50313 | 573  | 
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"  | 
| 67123 | 574  | 
where "iff p q =  | 
| 55885 | 575  | 
(if p = q then T  | 
576  | 
else if p = not q \<or> not p = q then F  | 
|
577  | 
else if p = F then not q  | 
|
578  | 
else if q = F then not p  | 
|
579  | 
else if p = T then q  | 
|
580  | 
else if q = T then p  | 
|
581  | 
else Iff p q)"  | 
|
| 50313 | 582  | 
|
| 23274 | 583  | 
lemma iff: "Ifm bbs bs (iff p q) = Ifm bbs bs (Iff p q)"  | 
| 55885 | 584  | 
by (unfold iff_def, cases "p = q", simp, cases "p = not q", simp add: not)  | 
585  | 
(cases "not p = q", auto simp add: not)  | 
|
| 50313 | 586  | 
|
| 55885 | 587  | 
lemma iff_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (iff p q)"  | 
588  | 
by (unfold iff_def, cases "p = q", auto simp add: not_qf)  | 
|
| 23274 | 589  | 
|
| 55885 | 590  | 
lemma iff_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)"  | 
591  | 
using iff_def by (unfold iff_def, cases "p = q", auto simp add: not_bn)  | 
|
| 50313 | 592  | 
|
| 66809 | 593  | 
fun simpfm :: "fm \<Rightarrow> fm"  | 
| 67123 | 594  | 
where  | 
595  | 
"simpfm (And p q) = conj (simpfm p) (simpfm q)"  | 
|
596  | 
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"  | 
|
597  | 
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"  | 
|
598  | 
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"  | 
|
599  | 
| "simpfm (NOT p) = not (simpfm p)"  | 
|
600  | 
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v < 0 then T else F | _ \<Rightarrow> Lt a')"  | 
|
601  | 
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<le> 0 then T else F | _ \<Rightarrow> Le a')"  | 
|
602  | 
| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v > 0 then T else F | _ \<Rightarrow> Gt a')"  | 
|
603  | 
| "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<ge> 0 then T else F | _ \<Rightarrow> Ge a')"  | 
|
604  | 
| "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v = 0 then T else F | _ \<Rightarrow> Eq a')"  | 
|
605  | 
| "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<noteq> 0 then T else F | _ \<Rightarrow> NEq a')"  | 
|
606  | 
| "simpfm (Dvd i a) =  | 
|
607  | 
(if i = 0 then simpfm (Eq a)  | 
|
608  | 
else if \<bar>i\<bar> = 1 then T  | 
|
609  | 
else let a' = simpnum a in case a' of C v \<Rightarrow> if i dvd v then T else F | _ \<Rightarrow> Dvd i a')"  | 
|
610  | 
| "simpfm (NDvd i a) =  | 
|
611  | 
(if i = 0 then simpfm (NEq a)  | 
|
612  | 
else if \<bar>i\<bar> = 1 then F  | 
|
613  | 
else let a' = simpnum a in case a' of C v \<Rightarrow> if \<not>( i dvd v) then T else F | _ \<Rightarrow> NDvd i a')"  | 
|
614  | 
| "simpfm p = p"  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
615  | 
|
| 23274 | 616  | 
lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs bs p"  | 
| 55844 | 617  | 
proof (induct p rule: simpfm.induct)  | 
| 50313 | 618  | 
case (6 a)  | 
619  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 620  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
621  | 
by simp  | 
|
| 60708 | 622  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
623  | 
then show ?case  | 
|
624  | 
proof cases  | 
|
625  | 
case 1  | 
|
626  | 
with sa show ?thesis by simp  | 
|
627  | 
next  | 
|
628  | 
case 2  | 
|
629  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
630  | 
qed  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
631  | 
next  | 
| 50313 | 632  | 
case (7 a)  | 
633  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 634  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
635  | 
by simp  | 
|
| 60708 | 636  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
637  | 
then show ?case  | 
|
638  | 
proof cases  | 
|
639  | 
case 1  | 
|
640  | 
with sa show ?thesis by simp  | 
|
641  | 
next  | 
|
642  | 
case 2  | 
|
643  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
644  | 
qed  | 
|
| 23274 | 645  | 
next  | 
| 50313 | 646  | 
case (8 a)  | 
647  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 648  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
649  | 
by simp  | 
|
| 60708 | 650  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
651  | 
then show ?case  | 
|
652  | 
proof cases  | 
|
653  | 
case 1  | 
|
654  | 
with sa show ?thesis by simp  | 
|
655  | 
next  | 
|
656  | 
case 2  | 
|
657  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
658  | 
qed  | 
|
| 23274 | 659  | 
next  | 
| 50313 | 660  | 
case (9 a)  | 
661  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 662  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
663  | 
by simp  | 
|
| 60708 | 664  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
665  | 
then show ?case  | 
|
666  | 
proof cases  | 
|
667  | 
case 1  | 
|
668  | 
with sa show ?thesis by simp  | 
|
669  | 
next  | 
|
670  | 
case 2  | 
|
671  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
672  | 
qed  | 
|
| 23274 | 673  | 
next  | 
| 50313 | 674  | 
case (10 a)  | 
675  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 676  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
677  | 
by simp  | 
|
| 60708 | 678  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
679  | 
then show ?case  | 
|
680  | 
proof cases  | 
|
681  | 
case 1  | 
|
682  | 
with sa show ?thesis by simp  | 
|
683  | 
next  | 
|
684  | 
case 2  | 
|
685  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
686  | 
qed  | 
|
| 23274 | 687  | 
next  | 
| 50313 | 688  | 
case (11 a)  | 
689  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 690  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
691  | 
by simp  | 
|
| 60708 | 692  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
693  | 
then show ?case  | 
|
694  | 
proof cases  | 
|
695  | 
case 1  | 
|
696  | 
with sa show ?thesis by simp  | 
|
697  | 
next  | 
|
698  | 
case 2  | 
|
699  | 
with sa show ?thesis by (cases ?sa) (simp_all add: Let_def)  | 
|
700  | 
qed  | 
|
| 23274 | 701  | 
next  | 
| 50313 | 702  | 
case (12 i a)  | 
703  | 
let ?sa = "simpnum a"  | 
|
| 55925 | 704  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
705  | 
by simp  | 
|
| 61945 | 706  | 
consider "i = 0" | "\<bar>i\<bar> = 1" | "i \<noteq> 0" "\<bar>i\<bar> \<noteq> 1" by blast  | 
| 60708 | 707  | 
then show ?case  | 
708  | 
proof cases  | 
|
709  | 
case 1  | 
|
710  | 
then show ?thesis  | 
|
711  | 
using "12.hyps" by (simp add: dvd_def Let_def)  | 
|
712  | 
next  | 
|
713  | 
case 2  | 
|
714  | 
with one_dvd[of "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"]  | 
|
715  | 
show ?thesis  | 
|
| 55925 | 716  | 
apply (cases "i = 0")  | 
717  | 
apply (simp_all add: Let_def)  | 
|
718  | 
apply (cases "i > 0")  | 
|
719  | 
apply simp_all  | 
|
| 50313 | 720  | 
done  | 
| 60708 | 721  | 
next  | 
722  | 
case i: 3  | 
|
723  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
|
724  | 
then show ?thesis  | 
|
725  | 
proof cases  | 
|
726  | 
case 1  | 
|
727  | 
with sa[symmetric] i show ?thesis  | 
|
| 61945 | 728  | 
by (cases "\<bar>i\<bar> = 1") auto  | 
| 60708 | 729  | 
next  | 
730  | 
case 2  | 
|
| 55925 | 731  | 
then have "simpfm (Dvd i a) = Dvd i ?sa"  | 
| 60708 | 732  | 
using i by (cases ?sa) (auto simp add: Let_def)  | 
733  | 
with sa show ?thesis by simp  | 
|
734  | 
qed  | 
|
735  | 
qed  | 
|
| 50313 | 736  | 
next  | 
737  | 
case (13 i a)  | 
|
| 55925 | 738  | 
let ?sa = "simpnum a"  | 
739  | 
from simpnum_ci have sa: "Inum bs ?sa = Inum bs a"  | 
|
740  | 
by simp  | 
|
| 61945 | 741  | 
consider "i = 0" | "\<bar>i\<bar> = 1" | "i \<noteq> 0" "\<bar>i\<bar> \<noteq> 1" by blast  | 
| 60708 | 742  | 
then show ?case  | 
743  | 
proof cases  | 
|
744  | 
case 1  | 
|
745  | 
then show ?thesis  | 
|
746  | 
using "13.hyps" by (simp add: dvd_def Let_def)  | 
|
747  | 
next  | 
|
748  | 
case 2  | 
|
749  | 
with one_dvd[of "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"]  | 
|
750  | 
show ?thesis  | 
|
| 55925 | 751  | 
apply (cases "i = 0")  | 
752  | 
apply (simp_all add: Let_def)  | 
|
753  | 
apply (cases "i > 0")  | 
|
754  | 
apply simp_all  | 
|
| 50313 | 755  | 
done  | 
| 60708 | 756  | 
next  | 
757  | 
case i: 3  | 
|
758  | 
consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast  | 
|
759  | 
then show ?thesis  | 
|
760  | 
proof cases  | 
|
761  | 
case 1  | 
|
762  | 
with sa[symmetric] i show ?thesis  | 
|
| 61945 | 763  | 
by (cases "\<bar>i\<bar> = 1") auto  | 
| 60708 | 764  | 
next  | 
765  | 
case 2  | 
|
| 55925 | 766  | 
then have "simpfm (NDvd i a) = NDvd i ?sa"  | 
| 60708 | 767  | 
using i by (cases ?sa) (auto simp add: Let_def)  | 
768  | 
with sa show ?thesis by simp  | 
|
769  | 
qed  | 
|
770  | 
qed  | 
|
| 50313 | 771  | 
qed (simp_all add: conj disj imp iff not)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
772  | 
|
| 23274 | 773  | 
lemma simpfm_bound0: "bound0 p \<Longrightarrow> bound0 (simpfm p)"  | 
| 50313 | 774  | 
proof (induct p rule: simpfm.induct)  | 
| 55925 | 775  | 
case (6 a)  | 
776  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 777  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
778  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 779  | 
next  | 
| 55925 | 780  | 
case (7 a)  | 
781  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 782  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
783  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 784  | 
next  | 
| 55925 | 785  | 
case (8 a)  | 
786  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 787  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
788  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 789  | 
next  | 
| 55925 | 790  | 
case (9 a)  | 
791  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 792  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
793  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 794  | 
next  | 
| 55925 | 795  | 
case (10 a)  | 
796  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 797  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
798  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 799  | 
next  | 
| 55925 | 800  | 
case (11 a)  | 
801  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 802  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
803  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 804  | 
next  | 
| 55925 | 805  | 
case (12 i a)  | 
806  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 807  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
808  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 23274 | 809  | 
next  | 
| 55925 | 810  | 
case (13 i a)  | 
811  | 
then have nb: "numbound0 a" by simp  | 
|
| 55885 | 812  | 
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])  | 
813  | 
then show ?case by (cases "simpnum a") (auto simp add: Let_def)  | 
|
| 50313 | 814  | 
qed (auto simp add: disj_def imp_def iff_def conj_def not_bn)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
815  | 
|
| 23274 | 816  | 
lemma simpfm_qf: "qfree p \<Longrightarrow> qfree (simpfm p)"  | 
| 60708 | 817  | 
apply (induct p rule: simpfm.induct)  | 
818  | 
apply (auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def)  | 
|
819  | 
apply (case_tac "simpnum a", auto)+  | 
|
820  | 
done  | 
|
| 23274 | 821  | 
|
| 70091 | 822  | 
|
823  | 
subsection \<open>Generic quantifier elimination\<close>  | 
|
824  | 
||
| 66809 | 825  | 
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"  | 
| 67123 | 826  | 
where  | 
827  | 
"qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"  | 
|
828  | 
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"  | 
|
829  | 
| "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"  | 
|
830  | 
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"  | 
|
831  | 
| "qelim (Or p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"  | 
|
832  | 
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"  | 
|
833  | 
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"  | 
|
834  | 
| "qelim p = (\<lambda>y. simpfm p)"  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
835  | 
|
| 23274 | 836  | 
lemma qelim_ci:  | 
| 55885 | 837  | 
assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)"  | 
838  | 
shows "\<And>bs. qfree (qelim p qe) \<and> Ifm bbs bs (qelim p qe) = Ifm bbs bs p"  | 
|
| 50313 | 839  | 
using qe_inv DJ_qe[OF qe_inv]  | 
| 55964 | 840  | 
by (induct p rule: qelim.induct)  | 
841  | 
(auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf  | 
|
842  | 
simpfm simpfm_qf simp del: simpfm.simps)  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
843  | 
|
| 61586 | 844  | 
text \<open>Linearity for fm where Bound 0 ranges over \<open>\<int>\<close>\<close>  | 
| 50313 | 845  | 
|
| 61586 | 846  | 
fun zsplit0 :: "num \<Rightarrow> int \<times> num" \<comment> \<open>splits the bounded from the unbounded part\<close>  | 
| 67123 | 847  | 
where  | 
848  | 
"zsplit0 (C c) = (0, C c)"  | 
|
849  | 
| "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))"  | 
|
850  | 
| "zsplit0 (CN n i a) =  | 
|
851  | 
(let (i', a') = zsplit0 a  | 
|
852  | 
in if n = 0 then (i + i', a') else (i', CN n i a'))"  | 
|
853  | 
| "zsplit0 (Neg a) = (let (i', a') = zsplit0 a in (-i', Neg a'))"  | 
|
854  | 
| "zsplit0 (Add a b) =  | 
|
855  | 
(let  | 
|
856  | 
(ia, a') = zsplit0 a;  | 
|
857  | 
(ib, b') = zsplit0 b  | 
|
858  | 
in (ia + ib, Add a' b'))"  | 
|
859  | 
| "zsplit0 (Sub a b) =  | 
|
860  | 
(let  | 
|
861  | 
(ia, a') = zsplit0 a;  | 
|
862  | 
(ib, b') = zsplit0 b  | 
|
863  | 
in (ia - ib, Sub a' b'))"  | 
|
864  | 
| "zsplit0 (Mul i a) = (let (i', a') = zsplit0 a in (i*i', Mul i a'))"  | 
|
| 23274 | 865  | 
|
866  | 
lemma zsplit0_I:  | 
|
| 55964 | 867  | 
"\<And>n a. zsplit0 t = (n, a) \<Longrightarrow>  | 
| 55921 | 868  | 
(Inum ((x::int) # bs) (CN 0 n a) = Inum (x # bs) t) \<and> numbound0 a"  | 
| 50313 | 869  | 
(is "\<And>n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a")  | 
870  | 
proof (induct t rule: zsplit0.induct)  | 
|
| 55844 | 871  | 
case (1 c n a)  | 
872  | 
then show ?case by auto  | 
|
| 23274 | 873  | 
next  | 
| 55844 | 874  | 
case (2 m n a)  | 
875  | 
then show ?case by (cases "m = 0") auto  | 
|
| 23274 | 876  | 
next  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
877  | 
case (3 m i a n a')  | 
| 23274 | 878  | 
let ?j = "fst (zsplit0 a)"  | 
879  | 
let ?b = "snd (zsplit0 a)"  | 
|
| 55844 | 880  | 
have abj: "zsplit0 a = (?j, ?b)" by simp  | 
| 60708 | 881  | 
show ?case  | 
882  | 
proof (cases "m = 0")  | 
|
883  | 
case False  | 
|
884  | 
with 3(1)[OF abj] 3(2) show ?thesis  | 
|
| 55844 | 885  | 
by (auto simp add: Let_def split_def)  | 
| 60708 | 886  | 
next  | 
887  | 
case m: True  | 
|
| 55964 | 888  | 
with abj have th: "a' = ?b \<and> n = i + ?j"  | 
889  | 
using 3 by (simp add: Let_def split_def)  | 
|
| 60708 | 890  | 
from abj 3 m have th2: "(?I x (CN 0 ?j ?b) = ?I x a) \<and> ?N ?b"  | 
| 55844 | 891  | 
by blast  | 
| 55964 | 892  | 
from th have "?I x (CN 0 n a') = ?I x (CN 0 (i + ?j) ?b)"  | 
| 55844 | 893  | 
by simp  | 
894  | 
also from th2 have "\<dots> = ?I x (CN 0 i (CN 0 ?j ?b))"  | 
|
895  | 
by (simp add: distrib_right)  | 
|
896  | 
finally have "?I x (CN 0 n a') = ?I x (CN 0 i a)"  | 
|
897  | 
using th2 by simp  | 
|
| 60708 | 898  | 
with th2 th m show ?thesis  | 
| 55844 | 899  | 
by blast  | 
| 60708 | 900  | 
qed  | 
| 23274 | 901  | 
next  | 
902  | 
case (4 t n a)  | 
|
903  | 
let ?nt = "fst (zsplit0 t)"  | 
|
904  | 
let ?at = "snd (zsplit0 t)"  | 
|
| 55964 | 905  | 
have abj: "zsplit0 t = (?nt, ?at)"  | 
906  | 
by simp  | 
|
907  | 
then have th: "a = Neg ?at \<and> n = - ?nt"  | 
|
| 55844 | 908  | 
using 4 by (simp add: Let_def split_def)  | 
909  | 
from abj 4 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at"  | 
|
910  | 
by blast  | 
|
911  | 
from th2[simplified] th[simplified] show ?case  | 
|
912  | 
by simp  | 
|
| 23274 | 913  | 
next  | 
914  | 
case (5 s t n a)  | 
|
915  | 
let ?ns = "fst (zsplit0 s)"  | 
|
916  | 
let ?as = "snd (zsplit0 s)"  | 
|
917  | 
let ?nt = "fst (zsplit0 t)"  | 
|
918  | 
let ?at = "snd (zsplit0 t)"  | 
|
| 55844 | 919  | 
have abjs: "zsplit0 s = (?ns, ?as)"  | 
920  | 
by simp  | 
|
921  | 
moreover have abjt: "zsplit0 t = (?nt, ?at)"  | 
|
922  | 
by simp  | 
|
| 55964 | 923  | 
ultimately have th: "a = Add ?as ?at \<and> n = ?ns + ?nt"  | 
| 55844 | 924  | 
using 5 by (simp add: Let_def split_def)  | 
| 55964 | 925  | 
from abjs[symmetric] have bluddy: "\<exists>x y. (x, y) = zsplit0 s"  | 
| 55844 | 926  | 
by blast  | 
927  | 
from 5 have "(\<exists>x y. (x, y) = zsplit0 s) \<longrightarrow>  | 
|
928  | 
(\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)"  | 
|
929  | 
by auto  | 
|
930  | 
with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at"  | 
|
931  | 
by blast  | 
|
932  | 
from abjs 5 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as"  | 
|
933  | 
by blast  | 
|
| 50313 | 934  | 
from th3[simplified] th2[simplified] th[simplified] show ?case  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
48891 
diff
changeset
 | 
935  | 
by (simp add: distrib_right)  | 
| 23274 | 936  | 
next  | 
937  | 
case (6 s t n a)  | 
|
938  | 
let ?ns = "fst (zsplit0 s)"  | 
|
939  | 
let ?as = "snd (zsplit0 s)"  | 
|
940  | 
let ?nt = "fst (zsplit0 t)"  | 
|
941  | 
let ?at = "snd (zsplit0 t)"  | 
|
| 55844 | 942  | 
have abjs: "zsplit0 s = (?ns, ?as)"  | 
943  | 
by simp  | 
|
944  | 
moreover have abjt: "zsplit0 t = (?nt, ?at)"  | 
|
945  | 
by simp  | 
|
| 55964 | 946  | 
ultimately have th: "a = Sub ?as ?at \<and> n = ?ns - ?nt"  | 
| 55844 | 947  | 
using 6 by (simp add: Let_def split_def)  | 
| 55964 | 948  | 
from abjs[symmetric] have bluddy: "\<exists>x y. (x, y) = zsplit0 s"  | 
| 55844 | 949  | 
by blast  | 
| 50313 | 950  | 
from 6 have "(\<exists>x y. (x,y) = zsplit0 s) \<longrightarrow>  | 
951  | 
(\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)"  | 
|
952  | 
by auto  | 
|
| 55844 | 953  | 
with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at"  | 
954  | 
by blast  | 
|
955  | 
from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as"  | 
|
956  | 
by blast  | 
|
| 50313 | 957  | 
from th3[simplified] th2[simplified] th[simplified] show ?case  | 
| 23274 | 958  | 
by (simp add: left_diff_distrib)  | 
959  | 
next  | 
|
960  | 
case (7 i t n a)  | 
|
961  | 
let ?nt = "fst (zsplit0 t)"  | 
|
962  | 
let ?at = "snd (zsplit0 t)"  | 
|
| 55844 | 963  | 
have abj: "zsplit0 t = (?nt,?at)"  | 
964  | 
by simp  | 
|
| 55964 | 965  | 
then have th: "a = Mul i ?at \<and> n = i * ?nt"  | 
| 55844 | 966  | 
using 7 by (simp add: Let_def split_def)  | 
967  | 
from abj 7 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at"  | 
|
968  | 
by blast  | 
|
969  | 
then have "?I x (Mul i t) = i * ?I x (CN 0 ?nt ?at)"  | 
|
970  | 
by simp  | 
|
971  | 
also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))"  | 
|
972  | 
by (simp add: distrib_left)  | 
|
973  | 
finally show ?case using th th2  | 
|
974  | 
by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
975  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
976  | 
|
| 67123 | 977  | 
fun iszlfm :: "fm \<Rightarrow> bool" \<comment> \<open>linearity test for fm\<close>  | 
978  | 
where  | 
|
979  | 
"iszlfm (And p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"  | 
|
980  | 
| "iszlfm (Or p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q"  | 
|
981  | 
| "iszlfm (Eq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
982  | 
| "iszlfm (NEq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
983  | 
| "iszlfm (Lt (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
984  | 
| "iszlfm (Le (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
985  | 
| "iszlfm (Gt (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
986  | 
| "iszlfm (Ge (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e"  | 
|
987  | 
| "iszlfm (Dvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"  | 
|
988  | 
| "iszlfm (NDvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e"  | 
|
989  | 
| "iszlfm p \<longleftrightarrow> isatom p \<and> bound0 p"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
990  | 
|
| 23274 | 991  | 
lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p"  | 
992  | 
by (induct p rule: iszlfm.induct) auto  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
993  | 
|
| 67123 | 994  | 
fun zlfm :: "fm \<Rightarrow> fm" \<comment> \<open>linearity transformation for fm\<close>  | 
995  | 
where  | 
|
996  | 
"zlfm (And p q) = And (zlfm p) (zlfm q)"  | 
|
997  | 
| "zlfm (Or p q) = Or (zlfm p) (zlfm q)"  | 
|
998  | 
| "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)"  | 
|
999  | 
| "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))"  | 
|
1000  | 
| "zlfm (Lt a) =  | 
|
1001  | 
(let (c, r) = zsplit0 a in  | 
|
1002  | 
if c = 0 then Lt r else  | 
|
1003  | 
if c > 0 then (Lt (CN 0 c r))  | 
|
1004  | 
else Gt (CN 0 (- c) (Neg r)))"  | 
|
1005  | 
| "zlfm (Le a) =  | 
|
1006  | 
(let (c, r) = zsplit0 a in  | 
|
1007  | 
if c = 0 then Le r  | 
|
1008  | 
else if c > 0 then Le (CN 0 c r)  | 
|
1009  | 
else Ge (CN 0 (- c) (Neg r)))"  | 
|
1010  | 
| "zlfm (Gt a) =  | 
|
1011  | 
(let (c, r) = zsplit0 a in  | 
|
1012  | 
if c = 0 then Gt r else  | 
|
1013  | 
if c > 0 then Gt (CN 0 c r)  | 
|
1014  | 
else Lt (CN 0 (- c) (Neg r)))"  | 
|
1015  | 
| "zlfm (Ge a) =  | 
|
1016  | 
(let (c, r) = zsplit0 a in  | 
|
1017  | 
if c = 0 then Ge r  | 
|
1018  | 
else if c > 0 then Ge (CN 0 c r)  | 
|
1019  | 
else Le (CN 0 (- c) (Neg r)))"  | 
|
1020  | 
| "zlfm (Eq a) =  | 
|
1021  | 
(let (c, r) = zsplit0 a in  | 
|
1022  | 
if c = 0 then Eq r  | 
|
1023  | 
else if c > 0 then Eq (CN 0 c r)  | 
|
1024  | 
else Eq (CN 0 (- c) (Neg r)))"  | 
|
1025  | 
| "zlfm (NEq a) =  | 
|
1026  | 
(let (c, r) = zsplit0 a in  | 
|
1027  | 
if c = 0 then NEq r  | 
|
1028  | 
else if c > 0 then NEq (CN 0 c r)  | 
|
1029  | 
else NEq (CN 0 (- c) (Neg r)))"  | 
|
1030  | 
| "zlfm (Dvd i a) =  | 
|
1031  | 
(if i = 0 then zlfm (Eq a)  | 
|
1032  | 
else  | 
|
1033  | 
let (c, r) = zsplit0 a in  | 
|
1034  | 
if c = 0 then Dvd \<bar>i\<bar> r  | 
|
1035  | 
else if c > 0 then Dvd \<bar>i\<bar> (CN 0 c r)  | 
|
1036  | 
else Dvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"  | 
|
1037  | 
| "zlfm (NDvd i a) =  | 
|
1038  | 
(if i = 0 then zlfm (NEq a)  | 
|
1039  | 
else  | 
|
1040  | 
let (c, r) = zsplit0 a in  | 
|
1041  | 
if c = 0 then NDvd \<bar>i\<bar> r  | 
|
1042  | 
else if c > 0 then NDvd \<bar>i\<bar> (CN 0 c r)  | 
|
1043  | 
else NDvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"  | 
|
1044  | 
| "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))"  | 
|
1045  | 
| "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))"  | 
|
1046  | 
| "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))"  | 
|
1047  | 
| "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))"  | 
|
1048  | 
| "zlfm (NOT (NOT p)) = zlfm p"  | 
|
1049  | 
| "zlfm (NOT T) = F"  | 
|
1050  | 
| "zlfm (NOT F) = T"  | 
|
1051  | 
| "zlfm (NOT (Lt a)) = zlfm (Ge a)"  | 
|
1052  | 
| "zlfm (NOT (Le a)) = zlfm (Gt a)"  | 
|
1053  | 
| "zlfm (NOT (Gt a)) = zlfm (Le a)"  | 
|
1054  | 
| "zlfm (NOT (Ge a)) = zlfm (Lt a)"  | 
|
1055  | 
| "zlfm (NOT (Eq a)) = zlfm (NEq a)"  | 
|
1056  | 
| "zlfm (NOT (NEq a)) = zlfm (Eq a)"  | 
|
1057  | 
| "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"  | 
|
1058  | 
| "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"  | 
|
1059  | 
| "zlfm (NOT (Closed P)) = NClosed P"  | 
|
1060  | 
| "zlfm (NOT (NClosed P)) = Closed P"  | 
|
1061  | 
| "zlfm p = p"  | 
|
| 23274 | 1062  | 
|
1063  | 
lemma zlfm_I:  | 
|
1064  | 
assumes qfp: "qfree p"  | 
|
| 55981 | 1065  | 
shows "Ifm bbs (i # bs) (zlfm p) = Ifm bbs (i # bs) p \<and> iszlfm (zlfm p)"  | 
| 60708 | 1066  | 
(is "?I (?l p) = ?I p \<and> ?L (?l p)")  | 
| 50313 | 1067  | 
using qfp  | 
1068  | 
proof (induct p rule: zlfm.induct)  | 
|
1069  | 
case (5 a)  | 
|
| 23274 | 1070  | 
let ?c = "fst (zsplit0 a)"  | 
1071  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1072  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1073  | 
by simp  | 
|
| 50313 | 1074  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55964 | 1075  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
| 55844 | 1076  | 
by auto  | 
| 55964 | 1077  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
1078  | 
from 5 Ia nb show ?case  | 
|
| 50313 | 1079  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
| 55844 | 1080  | 
apply (cases "?r")  | 
1081  | 
apply auto  | 
|
| 60708 | 1082  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1083  | 
done  | 
| 23274 | 1084  | 
next  | 
| 50313 | 1085  | 
case (6 a)  | 
| 23274 | 1086  | 
let ?c = "fst (zsplit0 a)"  | 
1087  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1088  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1089  | 
by simp  | 
|
| 50313 | 1090  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55964 | 1091  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
| 55844 | 1092  | 
by auto  | 
| 55964 | 1093  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 50313 | 1094  | 
from 6 Ia nb show ?case  | 
1095  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
|
| 55844 | 1096  | 
apply (cases "?r")  | 
1097  | 
apply auto  | 
|
| 60708 | 1098  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1099  | 
done  | 
| 23274 | 1100  | 
next  | 
| 50313 | 1101  | 
case (7 a)  | 
| 23274 | 1102  | 
let ?c = "fst (zsplit0 a)"  | 
1103  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1104  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1105  | 
by simp  | 
|
| 50313 | 1106  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55844 | 1107  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
1108  | 
by auto  | 
|
| 55964 | 1109  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 50313 | 1110  | 
from 7 Ia nb show ?case  | 
1111  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
|
| 55844 | 1112  | 
apply (cases "?r")  | 
1113  | 
apply auto  | 
|
| 60708 | 1114  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1115  | 
done  | 
| 23274 | 1116  | 
next  | 
| 50313 | 1117  | 
case (8 a)  | 
| 23274 | 1118  | 
let ?c = "fst (zsplit0 a)"  | 
1119  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1120  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1121  | 
by simp  | 
|
| 50313 | 1122  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55964 | 1123  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
| 55844 | 1124  | 
by auto  | 
| 55964 | 1125  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 55844 | 1126  | 
from 8 Ia nb show ?case  | 
| 50313 | 1127  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
| 55844 | 1128  | 
apply (cases "?r")  | 
1129  | 
apply auto  | 
|
| 60708 | 1130  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1131  | 
done  | 
| 23274 | 1132  | 
next  | 
| 50313 | 1133  | 
case (9 a)  | 
| 23274 | 1134  | 
let ?c = "fst (zsplit0 a)"  | 
1135  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1136  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1137  | 
by simp  | 
|
| 50313 | 1138  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55844 | 1139  | 
have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
1140  | 
by auto  | 
|
| 55964 | 1141  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 55844 | 1142  | 
from 9 Ia nb show ?case  | 
| 50313 | 1143  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
| 55844 | 1144  | 
apply (cases "?r")  | 
1145  | 
apply auto  | 
|
| 60708 | 1146  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1147  | 
done  | 
| 23274 | 1148  | 
next  | 
| 50313 | 1149  | 
case (10 a)  | 
| 23274 | 1150  | 
let ?c = "fst (zsplit0 a)"  | 
1151  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1152  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1153  | 
by simp  | 
|
| 50313 | 1154  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55844 | 1155  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
1156  | 
by auto  | 
|
| 55964 | 1157  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 55844 | 1158  | 
from 10 Ia nb show ?case  | 
| 50313 | 1159  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
| 55844 | 1160  | 
apply (cases "?r")  | 
1161  | 
apply auto  | 
|
| 60708 | 1162  | 
subgoal for nat a b by (cases nat) auto  | 
| 
23995
 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 
chaieb 
parents: 
23984 
diff
changeset
 | 
1163  | 
done  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1164  | 
next  | 
| 50313 | 1165  | 
case (11 j a)  | 
| 23274 | 1166  | 
let ?c = "fst (zsplit0 a)"  | 
1167  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1168  | 
have spl: "zsplit0 a = (?c,?r)"  | 
1169  | 
by simp  | 
|
| 50313 | 1170  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 55844 | 1171  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
1172  | 
by auto  | 
|
| 50313 | 1173  | 
let ?N = "\<lambda>t. Inum (i#bs) t"  | 
| 60708 | 1174  | 
consider "j = 0" | "j \<noteq> 0" "?c = 0" | "j \<noteq> 0" "?c > 0" | "j \<noteq> 0" "?c < 0"  | 
| 55844 | 1175  | 
by arith  | 
| 60708 | 1176  | 
then show ?case  | 
1177  | 
proof cases  | 
|
1178  | 
case 1  | 
|
| 55844 | 1179  | 
then have z: "zlfm (Dvd j a) = (zlfm (Eq a))"  | 
1180  | 
by (simp add: Let_def)  | 
|
| 60708 | 1181  | 
with 11 \<open>j = 0\<close> show ?thesis  | 
| 55844 | 1182  | 
by (simp del: zlfm.simps)  | 
| 60708 | 1183  | 
next  | 
1184  | 
case 2  | 
|
1185  | 
with zsplit0_I[OF spl, where x="i" and bs="bs"] show ?thesis  | 
|
1186  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
|
1187  | 
apply (cases "?r")  | 
|
1188  | 
apply auto  | 
|
1189  | 
subgoal for nat a b by (cases nat) auto  | 
|
1190  | 
done  | 
|
1191  | 
next  | 
|
1192  | 
case 3  | 
|
| 55844 | 1193  | 
then have l: "?L (?l (Dvd j a))"  | 
| 23274 | 1194  | 
by (simp add: nb Let_def split_def)  | 
| 60708 | 1195  | 
with Ia 3 show ?thesis  | 
1196  | 
by (simp add: Let_def split_def)  | 
|
1197  | 
next  | 
|
1198  | 
case 4  | 
|
| 55844 | 1199  | 
then have l: "?L (?l (Dvd j a))"  | 
| 23274 | 1200  | 
by (simp add: nb Let_def split_def)  | 
| 61945 | 1201  | 
with Ia 4 dvd_minus_iff[of "\<bar>j\<bar>" "?c*i + ?N ?r"] show ?thesis  | 
| 55844 | 1202  | 
by (simp add: Let_def split_def)  | 
| 60708 | 1203  | 
qed  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1204  | 
next  | 
| 50313 | 1205  | 
case (12 j a)  | 
| 23274 | 1206  | 
let ?c = "fst (zsplit0 a)"  | 
1207  | 
let ?r = "snd (zsplit0 a)"  | 
|
| 55844 | 1208  | 
have spl: "zsplit0 a = (?c, ?r)"  | 
1209  | 
by simp  | 
|
| 50313 | 1210  | 
from zsplit0_I[OF spl, where x="i" and bs="bs"]  | 
| 67123 | 1211  | 
have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r"  | 
| 55844 | 1212  | 
by auto  | 
| 55964 | 1213  | 
let ?N = "\<lambda>t. Inum (i # bs) t"  | 
| 60708 | 1214  | 
consider "j = 0" | "j \<noteq> 0" "?c = 0" | "j \<noteq> 0" "?c > 0" | "j \<noteq> 0" "?c < 0"  | 
| 55844 | 1215  | 
by arith  | 
| 60708 | 1216  | 
then show ?case  | 
1217  | 
proof cases  | 
|
1218  | 
case 1  | 
|
| 55964 | 1219  | 
then have z: "zlfm (NDvd j a) = zlfm (NEq a)"  | 
| 55844 | 1220  | 
by (simp add: Let_def)  | 
| 60708 | 1221  | 
with assms 12 \<open>j = 0\<close> show ?thesis  | 
1222  | 
by (simp del: zlfm.simps)  | 
|
1223  | 
next  | 
|
1224  | 
case 2  | 
|
1225  | 
with zsplit0_I[OF spl, where x="i" and bs="bs"] show ?thesis  | 
|
1226  | 
apply (auto simp add: Let_def split_def algebra_simps)  | 
|
1227  | 
apply (cases "?r")  | 
|
1228  | 
apply auto  | 
|
1229  | 
subgoal for nat a b by (cases nat) auto  | 
|
1230  | 
done  | 
|
1231  | 
next  | 
|
1232  | 
case 3  | 
|
| 55844 | 1233  | 
then have l: "?L (?l (Dvd j a))"  | 
| 23274 | 1234  | 
by (simp add: nb Let_def split_def)  | 
| 60708 | 1235  | 
with Ia 3 show ?thesis  | 
| 55844 | 1236  | 
by (simp add: Let_def split_def)  | 
| 60708 | 1237  | 
next  | 
1238  | 
case 4  | 
|
| 55844 | 1239  | 
then have l: "?L (?l (Dvd j a))"  | 
| 23274 | 1240  | 
by (simp add: nb Let_def split_def)  | 
| 61945 | 1241  | 
with Ia 4 dvd_minus_iff[of "\<bar>j\<bar>" "?c*i + ?N ?r"] show ?thesis  | 
| 55844 | 1242  | 
by (simp add: Let_def split_def)  | 
| 60708 | 1243  | 
qed  | 
| 23274 | 1244  | 
qed auto  | 
1245  | 
||
| 67123 | 1246  | 
fun minusinf :: "fm \<Rightarrow> fm" \<comment> \<open>virtual substitution of \<open>-\<infinity>\<close>\<close>  | 
1247  | 
where  | 
|
1248  | 
"minusinf (And p q) = And (minusinf p) (minusinf q)"  | 
|
1249  | 
| "minusinf (Or p q) = Or (minusinf p) (minusinf q)"  | 
|
1250  | 
| "minusinf (Eq (CN 0 c e)) = F"  | 
|
1251  | 
| "minusinf (NEq (CN 0 c e)) = T"  | 
|
1252  | 
| "minusinf (Lt (CN 0 c e)) = T"  | 
|
1253  | 
| "minusinf (Le (CN 0 c e)) = T"  | 
|
1254  | 
| "minusinf (Gt (CN 0 c e)) = F"  | 
|
1255  | 
| "minusinf (Ge (CN 0 c e)) = F"  | 
|
1256  | 
| "minusinf p = p"  | 
|
| 23274 | 1257  | 
|
1258  | 
lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"  | 
|
| 50313 | 1259  | 
by (induct p rule: minusinf.induct) auto  | 
| 23274 | 1260  | 
|
| 67123 | 1261  | 
fun plusinf :: "fm \<Rightarrow> fm" \<comment> \<open>virtual substitution of \<open>+\<infinity>\<close>\<close>  | 
1262  | 
where  | 
|
1263  | 
"plusinf (And p q) = And (plusinf p) (plusinf q)"  | 
|
1264  | 
| "plusinf (Or p q) = Or (plusinf p) (plusinf q)"  | 
|
1265  | 
| "plusinf (Eq (CN 0 c e)) = F"  | 
|
1266  | 
| "plusinf (NEq (CN 0 c e)) = T"  | 
|
1267  | 
| "plusinf (Lt (CN 0 c e)) = F"  | 
|
1268  | 
| "plusinf (Le (CN 0 c e)) = F"  | 
|
1269  | 
| "plusinf (Gt (CN 0 c e)) = T"  | 
|
1270  | 
| "plusinf (Ge (CN 0 c e)) = T"  | 
|
1271  | 
| "plusinf p = p"  | 
|
| 23274 | 1272  | 
|
| 67123 | 1273  | 
fun \<delta> :: "fm \<Rightarrow> int"  \<comment> \<open>compute \<open>lcm {d| N\<^sup>? Dvd c*x+t \<in> p}\<close>\<close>
 | 
1274  | 
where  | 
|
1275  | 
"\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"  | 
|
1276  | 
| "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"  | 
|
1277  | 
| "\<delta> (Dvd i (CN 0 c e)) = i"  | 
|
1278  | 
| "\<delta> (NDvd i (CN 0 c e)) = i"  | 
|
1279  | 
| "\<delta> p = 1"  | 
|
| 23274 | 1280  | 
|
| 67123 | 1281  | 
fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" \<comment> \<open>check if a given \<open>l\<close> divides all the \<open>ds\<close> above\<close>  | 
1282  | 
where  | 
|
1283  | 
"d_\<delta> (And p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"  | 
|
1284  | 
| "d_\<delta> (Or p q) d \<longleftrightarrow> d_\<delta> p d \<and> d_\<delta> q d"  | 
|
1285  | 
| "d_\<delta> (Dvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"  | 
|
1286  | 
| "d_\<delta> (NDvd i (CN 0 c e)) d \<longleftrightarrow> i dvd d"  | 
|
1287  | 
| "d_\<delta> p d \<longleftrightarrow> True"  | 
|
| 23274 | 1288  | 
|
| 50313 | 1289  | 
lemma delta_mono:  | 
| 23274 | 1290  | 
assumes lin: "iszlfm p"  | 
| 50313 | 1291  | 
and d: "d dvd d'"  | 
1292  | 
and ad: "d_\<delta> p d"  | 
|
| 50252 | 1293  | 
shows "d_\<delta> p d'"  | 
| 55999 | 1294  | 
using lin ad  | 
| 50313 | 1295  | 
proof (induct p rule: iszlfm.induct)  | 
| 55844 | 1296  | 
case (9 i c e)  | 
1297  | 
then show ?case using d  | 
|
| 30042 | 1298  | 
by (simp add: dvd_trans[of "i" "d" "d'"])  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1299  | 
next  | 
| 55844 | 1300  | 
case (10 i c e)  | 
1301  | 
then show ?case using d  | 
|
| 30042 | 1302  | 
by (simp add: dvd_trans[of "i" "d" "d'"])  | 
| 23274 | 1303  | 
qed simp_all  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1304  | 
|
| 50313 | 1305  | 
lemma \<delta>:  | 
| 55885 | 1306  | 
assumes lin: "iszlfm p"  | 
| 50252 | 1307  | 
shows "d_\<delta> p (\<delta> p) \<and> \<delta> p >0"  | 
| 50313 | 1308  | 
using lin  | 
| 67123 | 1309  | 
by (induct p rule: iszlfm.induct) (auto intro: delta_mono simp add: lcm_pos_int)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1310  | 
|
| 65024 | 1311  | 
fun a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" \<comment> \<open>adjust the coefficients of a formula\<close>  | 
| 67123 | 1312  | 
where  | 
1313  | 
"a_\<beta> (And p q) k = And (a_\<beta> p k) (a_\<beta> q k)"  | 
|
1314  | 
| "a_\<beta> (Or p q) k = Or (a_\<beta> p k) (a_\<beta> q k)"  | 
|
1315  | 
| "a_\<beta> (Eq (CN 0 c e)) k = Eq (CN 0 1 (Mul (k div c) e))"  | 
|
1316  | 
| "a_\<beta> (NEq (CN 0 c e)) k = NEq (CN 0 1 (Mul (k div c) e))"  | 
|
1317  | 
| "a_\<beta> (Lt (CN 0 c e)) k = Lt (CN 0 1 (Mul (k div c) e))"  | 
|
1318  | 
| "a_\<beta> (Le (CN 0 c e)) k = Le (CN 0 1 (Mul (k div c) e))"  | 
|
1319  | 
| "a_\<beta> (Gt (CN 0 c e)) k = Gt (CN 0 1 (Mul (k div c) e))"  | 
|
1320  | 
| "a_\<beta> (Ge (CN 0 c e)) k = Ge (CN 0 1 (Mul (k div c) e))"  | 
|
1321  | 
| "a_\<beta> (Dvd i (CN 0 c e)) k = Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"  | 
|
1322  | 
| "a_\<beta> (NDvd i (CN 0 c e)) k = NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e))"  | 
|
1323  | 
| "a_\<beta> p k = p"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1324  | 
|
| 67123 | 1325  | 
fun d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" \<comment> \<open>test if all coeffs of \<open>c\<close> divide a given \<open>l\<close>\<close>  | 
1326  | 
where  | 
|
1327  | 
"d_\<beta> (And p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"  | 
|
1328  | 
| "d_\<beta> (Or p q) k \<longleftrightarrow> d_\<beta> p k \<and> d_\<beta> q k"  | 
|
1329  | 
| "d_\<beta> (Eq (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1330  | 
| "d_\<beta> (NEq (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1331  | 
| "d_\<beta> (Lt (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1332  | 
| "d_\<beta> (Le (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1333  | 
| "d_\<beta> (Gt (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1334  | 
| "d_\<beta> (Ge (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1335  | 
| "d_\<beta> (Dvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1336  | 
| "d_\<beta> (NDvd i (CN 0 c e)) k \<longleftrightarrow> c dvd k"  | 
|
1337  | 
| "d_\<beta> p k \<longleftrightarrow> True"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1338  | 
|
| 67123 | 1339  | 
fun \<zeta> :: "fm \<Rightarrow> int" \<comment> \<open>computes the lcm of all coefficients of \<open>x\<close>\<close>  | 
1340  | 
where  | 
|
1341  | 
"\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"  | 
|
1342  | 
| "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"  | 
|
1343  | 
| "\<zeta> (Eq (CN 0 c e)) = c"  | 
|
1344  | 
| "\<zeta> (NEq (CN 0 c e)) = c"  | 
|
1345  | 
| "\<zeta> (Lt (CN 0 c e)) = c"  | 
|
1346  | 
| "\<zeta> (Le (CN 0 c e)) = c"  | 
|
1347  | 
| "\<zeta> (Gt (CN 0 c e)) = c"  | 
|
1348  | 
| "\<zeta> (Ge (CN 0 c e)) = c"  | 
|
1349  | 
| "\<zeta> (Dvd i (CN 0 c e)) = c"  | 
|
1350  | 
| "\<zeta> (NDvd i (CN 0 c e))= c"  | 
|
1351  | 
| "\<zeta> p = 1"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1352  | 
|
| 65024 | 1353  | 
fun \<beta> :: "fm \<Rightarrow> num list"  | 
| 67123 | 1354  | 
where  | 
1355  | 
"\<beta> (And p q) = (\<beta> p @ \<beta> q)"  | 
|
1356  | 
| "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"  | 
|
1357  | 
| "\<beta> (Eq (CN 0 c e)) = [Sub (C (- 1)) e]"  | 
|
1358  | 
| "\<beta> (NEq (CN 0 c e)) = [Neg e]"  | 
|
1359  | 
| "\<beta> (Lt (CN 0 c e)) = []"  | 
|
1360  | 
| "\<beta> (Le (CN 0 c e)) = []"  | 
|
1361  | 
| "\<beta> (Gt (CN 0 c e)) = [Neg e]"  | 
|
1362  | 
| "\<beta> (Ge (CN 0 c e)) = [Sub (C (- 1)) e]"  | 
|
1363  | 
| "\<beta> p = []"  | 
|
| 19736 | 1364  | 
|
| 65024 | 1365  | 
fun \<alpha> :: "fm \<Rightarrow> num list"  | 
| 67123 | 1366  | 
where  | 
1367  | 
"\<alpha> (And p q) = \<alpha> p @ \<alpha> q"  | 
|
1368  | 
| "\<alpha> (Or p q) = \<alpha> p @ \<alpha> q"  | 
|
1369  | 
| "\<alpha> (Eq (CN 0 c e)) = [Add (C (- 1)) e]"  | 
|
1370  | 
| "\<alpha> (NEq (CN 0 c e)) = [e]"  | 
|
1371  | 
| "\<alpha> (Lt (CN 0 c e)) = [e]"  | 
|
1372  | 
| "\<alpha> (Le (CN 0 c e)) = [Add (C (- 1)) e]"  | 
|
1373  | 
| "\<alpha> (Gt (CN 0 c e)) = []"  | 
|
1374  | 
| "\<alpha> (Ge (CN 0 c e)) = []"  | 
|
1375  | 
| "\<alpha> p = []"  | 
|
| 50313 | 1376  | 
|
| 65024 | 1377  | 
fun mirror :: "fm \<Rightarrow> fm"  | 
| 67123 | 1378  | 
where  | 
1379  | 
"mirror (And p q) = And (mirror p) (mirror q)"  | 
|
1380  | 
| "mirror (Or p q) = Or (mirror p) (mirror q)"  | 
|
1381  | 
| "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))"  | 
|
1382  | 
| "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"  | 
|
1383  | 
| "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))"  | 
|
1384  | 
| "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))"  | 
|
1385  | 
| "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))"  | 
|
1386  | 
| "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))"  | 
|
1387  | 
| "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"  | 
|
1388  | 
| "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"  | 
|
1389  | 
| "mirror p = p"  | 
|
| 50313 | 1390  | 
|
| 61586 | 1391  | 
text \<open>Lemmas for the correctness of \<open>\<sigma>_\<rho>\<close>\<close>  | 
| 50313 | 1392  | 
|
| 67123 | 1393  | 
lemma dvd1_eq1: "x > 0 \<Longrightarrow> x dvd 1 \<longleftrightarrow> x = 1"  | 
1394  | 
for x :: int  | 
|
| 41807 | 1395  | 
by simp  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1396  | 
|
| 23274 | 1397  | 
lemma minusinf_inf:  | 
1398  | 
assumes linp: "iszlfm p"  | 
|
| 50313 | 1399  | 
and u: "d_\<beta> p 1"  | 
| 55964 | 1400  | 
shows "\<exists>z::int. \<forall>x < z. Ifm bbs (x # bs) (minusinf p) = Ifm bbs (x # bs) p"  | 
| 50313 | 1401  | 
(is "?P p" is "\<exists>(z::int). \<forall>x < z. ?I x (?M p) = ?I x p")  | 
1402  | 
using linp u  | 
|
| 23274 | 1403  | 
proof (induct p rule: minusinf.induct)  | 
| 55844 | 1404  | 
case (1 p q)  | 
1405  | 
then show ?case  | 
|
| 60708 | 1406  | 
apply auto  | 
1407  | 
subgoal for z z' by (rule exI [where x = "min z z'"]) simp  | 
|
1408  | 
done  | 
|
| 23274 | 1409  | 
next  | 
| 55844 | 1410  | 
case (2 p q)  | 
1411  | 
then show ?case  | 
|
| 60708 | 1412  | 
apply auto  | 
1413  | 
subgoal for z z' by (rule exI [where x = "min z z'"]) simp  | 
|
1414  | 
done  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1415  | 
next  | 
| 55844 | 1416  | 
case (3 c e)  | 
1417  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1418  | 
by simp_all  | 
|
| 26934 | 1419  | 
fix a  | 
| 55999 | 1420  | 
from 3 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e \<noteq> 0"  | 
| 55844 | 1421  | 
proof clarsimp  | 
1422  | 
fix x  | 
|
| 55999 | 1423  | 
assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e = 0"  | 
| 23274 | 1424  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
| 55844 | 1425  | 
show False by simp  | 
| 23274 | 1426  | 
qed  | 
| 55844 | 1427  | 
then show ?case by auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1428  | 
next  | 
| 55844 | 1429  | 
case (4 c e)  | 
1430  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1431  | 
by simp_all  | 
|
| 26934 | 1432  | 
fix a  | 
| 55964 | 1433  | 
from 4 have "\<forall>x < (- Inum (a # bs) e). c * x + Inum (x # bs) e \<noteq> 0"  | 
| 55921 | 1434  | 
proof clarsimp  | 
1435  | 
fix x  | 
|
| 55964 | 1436  | 
assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e = 0"  | 
| 23274 | 1437  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
1438  | 
show "False" by simp  | 
|
1439  | 
qed  | 
|
| 55885 | 1440  | 
then show ?case by auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1441  | 
next  | 
| 55921 | 1442  | 
case (5 c e)  | 
1443  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1444  | 
by simp_all  | 
|
| 26934 | 1445  | 
fix a  | 
| 55999 | 1446  | 
from 5 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e < 0"  | 
| 55921 | 1447  | 
proof clarsimp  | 
1448  | 
fix x  | 
|
| 55964 | 1449  | 
assume "x < (- Inum (a # bs) e)"  | 
| 23274 | 1450  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
| 55921 | 1451  | 
show "x + Inum (x # bs) e < 0"  | 
1452  | 
by simp  | 
|
| 23274 | 1453  | 
qed  | 
| 55885 | 1454  | 
then show ?case by auto  | 
| 23274 | 1455  | 
next  | 
| 55921 | 1456  | 
case (6 c e)  | 
1457  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1458  | 
by simp_all  | 
|
| 26934 | 1459  | 
fix a  | 
| 55964 | 1460  | 
from 6 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e \<le> 0"  | 
| 55921 | 1461  | 
proof clarsimp  | 
1462  | 
fix x  | 
|
| 55964 | 1463  | 
assume "x < (- Inum (a # bs) e)"  | 
| 23274 | 1464  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
| 55964 | 1465  | 
show "x + Inum (x # bs) e \<le> 0" by simp  | 
| 23274 | 1466  | 
qed  | 
| 55885 | 1467  | 
then show ?case by auto  | 
| 23274 | 1468  | 
next  | 
| 55921 | 1469  | 
case (7 c e)  | 
1470  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1471  | 
by simp_all  | 
|
| 26934 | 1472  | 
fix a  | 
| 55964 | 1473  | 
from 7 have "\<forall>x<(- Inum (a # bs) e). \<not> (c * x + Inum (x # bs) e > 0)"  | 
| 55921 | 1474  | 
proof clarsimp  | 
1475  | 
fix x  | 
|
| 55964 | 1476  | 
assume "x < - Inum (a # bs) e" and "x + Inum (x # bs) e > 0"  | 
| 23274 | 1477  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
| 55921 | 1478  | 
show False by simp  | 
| 23274 | 1479  | 
qed  | 
| 55885 | 1480  | 
then show ?case by auto  | 
| 23274 | 1481  | 
next  | 
| 55921 | 1482  | 
case (8 c e)  | 
1483  | 
then have c1: "c = 1" and nb: "numbound0 e"  | 
|
1484  | 
by simp_all  | 
|
| 26934 | 1485  | 
fix a  | 
| 55999 | 1486  | 
from 8 have "\<forall>x<(- Inum (a # bs) e). \<not> c * x + Inum (x # bs) e \<ge> 0"  | 
| 55921 | 1487  | 
proof clarsimp  | 
1488  | 
fix x  | 
|
| 55999 | 1489  | 
assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e \<ge> 0"  | 
| 23274 | 1490  | 
with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"]  | 
| 55921 | 1491  | 
show False by simp  | 
| 23274 | 1492  | 
qed  | 
| 55885 | 1493  | 
then show ?case by auto  | 
| 23274 | 1494  | 
qed auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1495  | 
|
| 23274 | 1496  | 
lemma minusinf_repeats:  | 
| 55921 | 1497  | 
assumes d: "d_\<delta> p d"  | 
1498  | 
and linp: "iszlfm p"  | 
|
1499  | 
shows "Ifm bbs ((x - k * d) # bs) (minusinf p) = Ifm bbs (x # bs) (minusinf p)"  | 
|
| 50313 | 1500  | 
using linp d  | 
1501  | 
proof (induct p rule: iszlfm.induct)  | 
|
1502  | 
case (9 i c e)  | 
|
| 55921 | 1503  | 
then have nbe: "numbound0 e" and id: "i dvd d"  | 
1504  | 
by simp_all  | 
|
1505  | 
then have "\<exists>k. d = i * k"  | 
|
1506  | 
by (simp add: dvd_def)  | 
|
1507  | 
then obtain "di" where di_def: "d = i * di"  | 
|
1508  | 
by blast  | 
|
| 50313 | 1509  | 
show ?case  | 
1510  | 
proof (simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib,  | 
|
1511  | 
rule iffI)  | 
|
| 55921 | 1512  | 
assume "i dvd c * x - c * (k * d) + Inum (x # bs) e"  | 
| 55999 | 1513  | 
(is "?ri dvd ?rc * ?rx - ?rc * (?rk * ?rd) + ?I x e" is "?ri dvd ?rt")  | 
| 55921 | 1514  | 
then have "\<exists>l::int. ?rt = i * l"  | 
1515  | 
by (simp add: dvd_def)  | 
|
| 55964 | 1516  | 
then have "\<exists>l::int. c * x + ?I x e = i * l + c * (k * i * di)"  | 
| 50313 | 1517  | 
by (simp add: algebra_simps di_def)  | 
| 55964 | 1518  | 
then have "\<exists>l::int. c * x + ?I x e = i* (l + c * k * di)"  | 
| 50313 | 1519  | 
by (simp add: algebra_simps)  | 
| 55921 | 1520  | 
then have "\<exists>l::int. c * x + ?I x e = i * l"  | 
1521  | 
by blast  | 
|
1522  | 
then show "i dvd c * x + Inum (x # bs) e"  | 
|
1523  | 
by (simp add: dvd_def)  | 
|
| 50313 | 1524  | 
next  | 
| 55964 | 1525  | 
assume "i dvd c * x + Inum (x # bs) e" (is "?ri dvd ?rc * ?rx + ?e")  | 
| 55921 | 1526  | 
then have "\<exists>l::int. c * x + ?e = i * l"  | 
1527  | 
by (simp add: dvd_def)  | 
|
1528  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * d)"  | 
|
1529  | 
by simp  | 
|
1530  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * i * di)"  | 
|
1531  | 
by (simp add: di_def)  | 
|
| 55964 | 1532  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * (l - c * k * di)"  | 
| 55921 | 1533  | 
by (simp add: algebra_simps)  | 
1534  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l"  | 
|
1535  | 
by blast  | 
|
1536  | 
then show "i dvd c * x - c * (k * d) + Inum (x # bs) e"  | 
|
1537  | 
by (simp add: dvd_def)  | 
|
| 50313 | 1538  | 
qed  | 
| 23274 | 1539  | 
next  | 
| 50313 | 1540  | 
case (10 i c e)  | 
| 55921 | 1541  | 
then have nbe: "numbound0 e" and id: "i dvd d"  | 
1542  | 
by simp_all  | 
|
1543  | 
then have "\<exists>k. d = i * k"  | 
|
1544  | 
by (simp add: dvd_def)  | 
|
1545  | 
then obtain di where di_def: "d = i * di"  | 
|
1546  | 
by blast  | 
|
| 50313 | 1547  | 
show ?case  | 
| 55999 | 1548  | 
proof (simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib,  | 
1549  | 
rule iffI)  | 
|
| 55921 | 1550  | 
assume "i dvd c * x - c * (k * d) + Inum (x # bs) e"  | 
| 55999 | 1551  | 
(is "?ri dvd ?rc * ?rx - ?rc * (?rk * ?rd) + ?I x e" is "?ri dvd ?rt")  | 
| 55921 | 1552  | 
then have "\<exists>l::int. ?rt = i * l"  | 
1553  | 
by (simp add: dvd_def)  | 
|
1554  | 
then have "\<exists>l::int. c * x + ?I x e = i * l + c * (k * i * di)"  | 
|
| 50313 | 1555  | 
by (simp add: algebra_simps di_def)  | 
| 55921 | 1556  | 
then have "\<exists>l::int. c * x+ ?I x e = i * (l + c * k * di)"  | 
| 50313 | 1557  | 
by (simp add: algebra_simps)  | 
| 55921 | 1558  | 
then have "\<exists>l::int. c * x + ?I x e = i * l"  | 
1559  | 
by blast  | 
|
1560  | 
then show "i dvd c * x + Inum (x # bs) e"  | 
|
1561  | 
by (simp add: dvd_def)  | 
|
| 50313 | 1562  | 
next  | 
| 55921 | 1563  | 
assume "i dvd c * x + Inum (x # bs) e" (is "?ri dvd ?rc * ?rx + ?e")  | 
1564  | 
then have "\<exists>l::int. c * x + ?e = i * l"  | 
|
1565  | 
by (simp add: dvd_def)  | 
|
1566  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * d)"  | 
|
1567  | 
by simp  | 
|
1568  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * i * di)"  | 
|
1569  | 
by (simp add: di_def)  | 
|
| 55999 | 1570  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * (l - c * k * di)"  | 
| 55921 | 1571  | 
by (simp add: algebra_simps)  | 
1572  | 
then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l"  | 
|
| 50313 | 1573  | 
by blast  | 
| 55921 | 1574  | 
then show "i dvd c * x - c * (k * d) + Inum (x # bs) e"  | 
1575  | 
by (simp add: dvd_def)  | 
|
| 50313 | 1576  | 
qed  | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
1577  | 
qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="x - k*d" and b'="x"])  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1578  | 
|
| 50252 | 1579  | 
lemma mirror_\<alpha>_\<beta>:  | 
| 23274 | 1580  | 
assumes lp: "iszlfm p"  | 
| 55964 | 1581  | 
shows "Inum (i # bs) ` set (\<alpha> p) = Inum (i # bs) ` set (\<beta> (mirror p))"  | 
| 50313 | 1582  | 
using lp by (induct p rule: mirror.induct) auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1583  | 
|
| 50313 | 1584  | 
lemma mirror:  | 
| 23274 | 1585  | 
assumes lp: "iszlfm p"  | 
| 55921 | 1586  | 
shows "Ifm bbs (x # bs) (mirror p) = Ifm bbs ((- x) # bs) p"  | 
| 50313 | 1587  | 
using lp  | 
1588  | 
proof (induct p rule: iszlfm.induct)  | 
|
1589  | 
case (9 j c e)  | 
|
| 55964 | 1590  | 
then have nb: "numbound0 e"  | 
1591  | 
by simp  | 
|
1592  | 
have "Ifm bbs (x # bs) (mirror (Dvd j (CN 0 c e))) \<longleftrightarrow> j dvd c * x - Inum (x # bs) e"  | 
|
| 50313 | 1593  | 
(is "_ = (j dvd c*x - ?e)") by simp  | 
| 55964 | 1594  | 
also have "\<dots> \<longleftrightarrow> j dvd (- (c * x - ?e))"  | 
| 30042 | 1595  | 
by (simp only: dvd_minus_iff)  | 
| 55964 | 1596  | 
also have "\<dots> \<longleftrightarrow> j dvd (c * (- x)) + ?e"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
55999 
diff
changeset
 | 
1597  | 
by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] ac_simps minus_add_distrib)  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
53168 
diff
changeset
 | 
1598  | 
(simp add: algebra_simps)  | 
| 55964 | 1599  | 
also have "\<dots> = Ifm bbs ((- x) # bs) (Dvd j (CN 0 c e))"  | 
| 50313 | 1600  | 
using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp  | 
| 23274 | 1601  | 
finally show ?case .  | 
1602  | 
next  | 
|
| 55964 | 1603  | 
case (10 j c e)  | 
1604  | 
then have nb: "numbound0 e"  | 
|
1605  | 
by simp  | 
|
1606  | 
have "Ifm bbs (x # bs) (mirror (Dvd j (CN 0 c e))) \<longleftrightarrow> j dvd c * x - Inum (x # bs) e"  | 
|
1607  | 
(is "_ = (j dvd c * x - ?e)") by simp  | 
|
1608  | 
also have "\<dots> \<longleftrightarrow> j dvd (- (c * x - ?e))"  | 
|
| 30042 | 1609  | 
by (simp only: dvd_minus_iff)  | 
| 55964 | 1610  | 
also have "\<dots> \<longleftrightarrow> j dvd (c * (- x)) + ?e"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
55999 
diff
changeset
 | 
1611  | 
by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] ac_simps minus_add_distrib)  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
53168 
diff
changeset
 | 
1612  | 
(simp add: algebra_simps)  | 
| 55964 | 1613  | 
also have "\<dots> \<longleftrightarrow> Ifm bbs ((- x) # bs) (Dvd j (CN 0 c e))"  | 
| 50313 | 1614  | 
using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp  | 
| 23274 | 1615  | 
finally show ?case by simp  | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
1616  | 
qed (auto simp add: numbound0_I[where bs="bs" and b="x" and b'="- x"] gr0_conv_Suc)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1617  | 
|
| 50313 | 1618  | 
lemma mirror_l: "iszlfm p \<and> d_\<beta> p 1 \<Longrightarrow> iszlfm (mirror p) \<and> d_\<beta> (mirror p) 1"  | 
| 41807 | 1619  | 
by (induct p rule: mirror.induct) auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1620  | 
|
| 23274 | 1621  | 
lemma mirror_\<delta>: "iszlfm p \<Longrightarrow> \<delta> (mirror p) = \<delta> p"  | 
| 41807 | 1622  | 
by (induct p rule: mirror.induct) auto  | 
| 23274 | 1623  | 
|
| 50313 | 1624  | 
lemma \<beta>_numbound0:  | 
1625  | 
assumes lp: "iszlfm p"  | 
|
| 55964 | 1626  | 
shows "\<forall>b \<in> set (\<beta> p). numbound0 b"  | 
| 41807 | 1627  | 
using lp by (induct p rule: \<beta>.induct) auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1628  | 
|
| 50313 | 1629  | 
lemma d_\<beta>_mono:  | 
| 23274 | 1630  | 
assumes linp: "iszlfm p"  | 
| 50313 | 1631  | 
and dr: "d_\<beta> p l"  | 
1632  | 
and d: "l dvd l'"  | 
|
| 50252 | 1633  | 
shows "d_\<beta> p l'"  | 
| 50313 | 1634  | 
using dr linp dvd_trans[of _ "l" "l'", simplified d]  | 
| 41807 | 1635  | 
by (induct p rule: iszlfm.induct) simp_all  | 
| 23274 | 1636  | 
|
| 50313 | 1637  | 
lemma \<alpha>_l:  | 
| 55999 | 1638  | 
assumes "iszlfm p"  | 
| 50313 | 1639  | 
shows "\<forall>b \<in> set (\<alpha> p). numbound0 b"  | 
| 55999 | 1640  | 
using assms by (induct p rule: \<alpha>.induct) auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1641  | 
|
| 50313 | 1642  | 
lemma \<zeta>:  | 
| 55999 | 1643  | 
assumes "iszlfm p"  | 
| 50252 | 1644  | 
shows "\<zeta> p > 0 \<and> d_\<beta> p (\<zeta> p)"  | 
| 55999 | 1645  | 
using assms  | 
| 50313 | 1646  | 
proof (induct p rule: iszlfm.induct)  | 
| 23274 | 1647  | 
case (1 p q)  | 
| 55964 | 1648  | 
from 1 have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)"  | 
1649  | 
by simp  | 
|
1650  | 
from 1 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)"  | 
|
1651  | 
by simp  | 
|
| 50313 | 1652  | 
from 1 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]  | 
| 55964 | 1653  | 
d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]  | 
1654  | 
dl1 dl2  | 
|
1655  | 
show ?case  | 
|
1656  | 
by (auto simp add: lcm_pos_int)  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1657  | 
next  | 
| 23274 | 1658  | 
case (2 p q)  | 
| 55964 | 1659  | 
from 2 have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)"  | 
1660  | 
by simp  | 
|
1661  | 
from 2 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)"  | 
|
1662  | 
by simp  | 
|
| 50313 | 1663  | 
from 2 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]  | 
| 55964 | 1664  | 
d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]  | 
1665  | 
dl1 dl2  | 
|
1666  | 
show ?case  | 
|
1667  | 
by (auto simp add: lcm_pos_int)  | 
|
| 
31952
 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 
nipkow 
parents: 
31730 
diff
changeset
 | 
1668  | 
qed (auto simp add: lcm_pos_int)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1669  | 
|
| 50313 | 1670  | 
lemma a_\<beta>:  | 
| 55921 | 1671  | 
assumes linp: "iszlfm p"  | 
1672  | 
and d: "d_\<beta> p l"  | 
|
1673  | 
and lp: "l > 0"  | 
|
| 55964 | 1674  | 
shows "iszlfm (a_\<beta> p l) \<and> d_\<beta> (a_\<beta> p l) 1 \<and> Ifm bbs (l * x # bs) (a_\<beta> p l) = Ifm bbs (x # bs) p"  | 
| 50313 | 1675  | 
using linp d  | 
| 23274 | 1676  | 
proof (induct p rule: iszlfm.induct)  | 
| 50313 | 1677  | 
case (5 c e)  | 
| 55964 | 1678  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
| 55921 | 1679  | 
by simp_all  | 
1680  | 
from lp cp have clel: "c \<le> l"  | 
|
1681  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1682  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1683  | 
by simp  | 
|
1684  | 
have "c div c \<le> l div c"  | 
|
| 50313 | 1685  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
| 55999 | 1686  | 
then have ldcp: "0 < l div c"  | 
| 50313 | 1687  | 
by (simp add: div_self[OF cnz])  | 
| 55921 | 1688  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
1689  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
|
1690  | 
then have cl: "c * (l div c) =l"  | 
|
| 64246 | 1691  | 
using mult_div_mod_eq [where a="l" and b="c"] by simp  | 
| 55964 | 1692  | 
then have "(l * x + (l div c) * Inum (x # bs) e < 0) \<longleftrightarrow>  | 
| 50313 | 1693  | 
((c * (l div c)) * x + (l div c) * Inum (x # bs) e < 0)"  | 
1694  | 
by simp  | 
|
| 55999 | 1695  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) < (l div c) * 0"  | 
| 50313 | 1696  | 
by (simp add: algebra_simps)  | 
| 55964 | 1697  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e < 0"  | 
1698  | 
using mult_less_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp  | 
|
1699  | 
by simp  | 
|
| 50313 | 1700  | 
finally show ?case  | 
| 55964 | 1701  | 
using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be  | 
1702  | 
by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1703  | 
next  | 
| 50313 | 1704  | 
case (6 c e)  | 
| 55921 | 1705  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
1706  | 
by simp_all  | 
|
1707  | 
from lp cp have clel: "c \<le> l"  | 
|
1708  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1709  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1710  | 
by simp  | 
|
| 55964 | 1711  | 
have "c div c \<le> l div c"  | 
| 50313 | 1712  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
1713  | 
then have ldcp:"0 < l div c"  | 
|
1714  | 
by (simp add: div_self[OF cnz])  | 
|
| 55964 | 1715  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 55921 | 1716  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
1717  | 
then have cl: "c * (l div c) = l"  | 
|
| 64246 | 1718  | 
using mult_div_mod_eq [where a="l" and b="c"] by simp  | 
| 55964 | 1719  | 
then have "l * x + (l div c) * Inum (x # bs) e \<le> 0 \<longleftrightarrow>  | 
1720  | 
(c * (l div c)) * x + (l div c) * Inum (x # bs) e \<le> 0"  | 
|
| 50313 | 1721  | 
by simp  | 
| 55964 | 1722  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<le> (l div c) * 0"  | 
| 50313 | 1723  | 
by (simp add: algebra_simps)  | 
| 55964 | 1724  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<le> 0"  | 
| 23274 | 1725  | 
using mult_le_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp  | 
| 50313 | 1726  | 
finally show ?case  | 
1727  | 
using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1728  | 
next  | 
| 50313 | 1729  | 
case (7 c e)  | 
| 55921 | 1730  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
1731  | 
by simp_all  | 
|
1732  | 
from lp cp have clel: "c \<le> l"  | 
|
1733  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1734  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1735  | 
by simp  | 
|
1736  | 
have "c div c \<le> l div c"  | 
|
1737  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
|
| 55964 | 1738  | 
then have ldcp: "0 < l div c"  | 
| 55921 | 1739  | 
by (simp add: div_self[OF cnz])  | 
| 55964 | 1740  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 55921 | 1741  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
| 55964 | 1742  | 
then have cl: "c * (l div c) = l"  | 
| 64246 | 1743  | 
using mult_div_mod_eq [where a="l" and b="c"] by simp  | 
| 55964 | 1744  | 
then have "l * x + (l div c) * Inum (x # bs) e > 0 \<longleftrightarrow>  | 
1745  | 
(c * (l div c)) * x + (l div c) * Inum (x # bs) e > 0"  | 
|
| 55921 | 1746  | 
by simp  | 
| 55964 | 1747  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) > (l div c) * 0"  | 
| 55921 | 1748  | 
by (simp add: algebra_simps)  | 
| 55964 | 1749  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e > 0"  | 
| 55921 | 1750  | 
using zero_less_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp  | 
1751  | 
by simp  | 
|
1752  | 
finally show ?case  | 
|
1753  | 
using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be  | 
|
1754  | 
by simp  | 
|
1755  | 
next  | 
|
1756  | 
case (8 c e)  | 
|
1757  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
|
1758  | 
by simp_all  | 
|
1759  | 
from lp cp have clel: "c \<le> l"  | 
|
1760  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1761  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1762  | 
by simp  | 
|
1763  | 
have "c div c \<le> l div c"  | 
|
1764  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
|
1765  | 
then have ldcp: "0 < l div c"  | 
|
1766  | 
by (simp add: div_self[OF cnz])  | 
|
| 55964 | 1767  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 55921 | 1768  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
1769  | 
then have cl: "c * (l div c) =l"  | 
|
| 64246 | 1770  | 
using mult_div_mod_eq [where a="l" and b="c"]  | 
| 55921 | 1771  | 
by simp  | 
| 55964 | 1772  | 
then have "l * x + (l div c) * Inum (x # bs) e \<ge> 0 \<longleftrightarrow>  | 
1773  | 
(c * (l div c)) * x + (l div c) * Inum (x # bs) e \<ge> 0"  | 
|
| 55921 | 1774  | 
by simp  | 
| 55964 | 1775  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<ge> (l div c) * 0"  | 
| 55921 | 1776  | 
by (simp add: algebra_simps)  | 
| 55964 | 1777  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<ge> 0"  | 
| 55921 | 1778  | 
using ldcp zero_le_mult_iff [where a="l div c" and b="c*x + Inum (x # bs) e"]  | 
1779  | 
by simp  | 
|
1780  | 
finally show ?case  | 
|
1781  | 
using be numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"]  | 
|
1782  | 
by simp  | 
|
1783  | 
next  | 
|
1784  | 
case (3 c e)  | 
|
1785  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
|
1786  | 
by simp_all  | 
|
1787  | 
from lp cp have clel: "c \<le> l"  | 
|
1788  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1789  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1790  | 
by simp  | 
|
1791  | 
have "c div c \<le> l div c"  | 
|
| 50313 | 1792  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
1793  | 
then have ldcp:"0 < l div c"  | 
|
1794  | 
by (simp add: div_self[OF cnz])  | 
|
| 55964 | 1795  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 50313 | 1796  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
| 64246 | 1797  | 
then have cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"]  | 
| 50313 | 1798  | 
by simp  | 
| 55964 | 1799  | 
then have "l * x + (l div c) * Inum (x # bs) e = 0 \<longleftrightarrow>  | 
1800  | 
(c * (l div c)) * x + (l div c) * Inum (x # bs) e = 0"  | 
|
| 23274 | 1801  | 
by simp  | 
| 55964 | 1802  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0"  | 
| 50313 | 1803  | 
by (simp add: algebra_simps)  | 
| 55964 | 1804  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e = 0"  | 
| 55921 | 1805  | 
using mult_eq_0_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp  | 
1806  | 
by simp  | 
|
| 50313 | 1807  | 
finally show ?case  | 
| 55921 | 1808  | 
using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be  | 
1809  | 
by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1810  | 
next  | 
| 50313 | 1811  | 
case (4 c e)  | 
| 55921 | 1812  | 
then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l"  | 
1813  | 
by simp_all  | 
|
1814  | 
from lp cp have clel: "c \<le> l"  | 
|
1815  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1816  | 
from cp have cnz: "c \<noteq> 0"  | 
|
1817  | 
by simp  | 
|
1818  | 
have "c div c \<le> l div c"  | 
|
| 50313 | 1819  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
1820  | 
then have ldcp:"0 < l div c"  | 
|
1821  | 
by (simp add: div_self[OF cnz])  | 
|
| 55964 | 1822  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 55921 | 1823  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
1824  | 
then have cl: "c * (l div c) = l"  | 
|
| 64246 | 1825  | 
using mult_div_mod_eq [where a="l" and b="c"] by simp  | 
| 55964 | 1826  | 
then have "l * x + (l div c) * Inum (x # bs) e \<noteq> 0 \<longleftrightarrow>  | 
| 55921 | 1827  | 
(c * (l div c)) * x + (l div c) * Inum (x # bs) e \<noteq> 0"  | 
| 50313 | 1828  | 
by simp  | 
| 55921 | 1829  | 
also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<noteq> (l div c) * 0"  | 
| 50313 | 1830  | 
by (simp add: algebra_simps)  | 
| 55921 | 1831  | 
also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<noteq> 0"  | 
1832  | 
using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp  | 
|
1833  | 
by simp  | 
|
| 50313 | 1834  | 
finally show ?case  | 
| 55921 | 1835  | 
using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be  | 
1836  | 
by simp  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1837  | 
next  | 
| 50313 | 1838  | 
case (9 j c e)  | 
| 55921 | 1839  | 
then have cp: "c > 0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l"  | 
1840  | 
by simp_all  | 
|
1841  | 
from lp cp have clel: "c \<le> l"  | 
|
1842  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
| 50313 | 1843  | 
from cp have cnz: "c \<noteq> 0" by simp  | 
1844  | 
have "c div c\<le> l div c"  | 
|
1845  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
|
1846  | 
then have ldcp:"0 < l div c"  | 
|
1847  | 
by (simp add: div_self[OF cnz])  | 
|
| 55964 | 1848  | 
have "c * (l div c) = c * (l div c) + l mod c"  | 
| 55885 | 1849  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
| 55921 | 1850  | 
then have cl: "c * (l div c) = l"  | 
| 64246 | 1851  | 
using mult_div_mod_eq [where a="l" and b="c"] by simp  | 
| 55921 | 1852  | 
then have "(\<exists>k::int. l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) \<longleftrightarrow>  | 
| 55964 | 1853  | 
(\<exists>k::int. (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)"  | 
| 55921 | 1854  | 
by simp  | 
| 55964 | 1855  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c) * 0)"  | 
| 55921 | 1856  | 
by (simp add: algebra_simps)  | 
| 60708 | 1857  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e - j * k = 0)"  | 
1858  | 
using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k" for k] ldcp  | 
|
| 55921 | 1859  | 
by simp  | 
1860  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e = j * k)"  | 
|
1861  | 
by simp  | 
|
1862  | 
finally show ?case  | 
|
1863  | 
using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"]  | 
|
1864  | 
be mult_strict_mono[OF ldcp jp ldcp ]  | 
|
1865  | 
by (simp add: dvd_def)  | 
|
1866  | 
next  | 
|
1867  | 
case (10 j c e)  | 
|
1868  | 
then have cp: "c > 0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l"  | 
|
1869  | 
by simp_all  | 
|
1870  | 
from lp cp have clel: "c \<le> l"  | 
|
1871  | 
by (simp add: zdvd_imp_le [OF d' lp])  | 
|
1872  | 
from cp have cnz: "c \<noteq> 0"  | 
|
| 50313 | 1873  | 
by simp  | 
| 55921 | 1874  | 
have "c div c \<le> l div c"  | 
1875  | 
by (simp add: zdiv_mono1[OF clel cp])  | 
|
1876  | 
then have ldcp: "0 < l div c"  | 
|
1877  | 
by (simp add: div_self[OF cnz])  | 
|
1878  | 
have "c * (l div c) = c* (l div c) + l mod c"  | 
|
1879  | 
using d' dvd_eq_mod_eq_0[of "c" "l"] by simp  | 
|
1880  | 
then have cl:"c * (l div c) =l"  | 
|
| 64246 | 1881  | 
using mult_div_mod_eq [where a="l" and b="c"]  | 
| 55921 | 1882  | 
by simp  | 
1883  | 
then have "(\<exists>k::int. l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) \<longleftrightarrow>  | 
|
1884  | 
(\<exists>k::int. (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)"  | 
|
1885  | 
by simp  | 
|
1886  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c) * 0)"  | 
|
1887  | 
by (simp add: algebra_simps)  | 
|
| 60708 | 1888  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e - j * k = 0)"  | 
1889  | 
using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k" for k] ldcp  | 
|
| 55921 | 1890  | 
by simp  | 
| 55964 | 1891  | 
also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e = j * k)"  | 
| 55921 | 1892  | 
by simp  | 
1893  | 
finally show ?case  | 
|
1894  | 
using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be  | 
|
1895  | 
mult_strict_mono[OF ldcp jp ldcp ]  | 
|
1896  | 
by (simp add: dvd_def)  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
1897  | 
qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="(l * x)" and b'="x"])  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1898  | 
|
| 55921 | 1899  | 
lemma a_\<beta>_ex:  | 
1900  | 
assumes linp: "iszlfm p"  | 
|
1901  | 
and d: "d_\<beta> p l"  | 
|
1902  | 
and lp: "l > 0"  | 
|
1903  | 
shows "(\<exists>x. l dvd x \<and> Ifm bbs (x #bs) (a_\<beta> p l)) \<longleftrightarrow> (\<exists>x::int. Ifm bbs (x#bs) p)"  | 
|
1904  | 
(is "(\<exists>x. l dvd x \<and> ?P x) \<longleftrightarrow> (\<exists>x. ?P' x)")  | 
|
| 23274 | 1905  | 
proof-  | 
| 55999 | 1906  | 
have "(\<exists>x. l dvd x \<and> ?P x) \<longleftrightarrow> (\<exists>x::int. ?P (l * x))"  | 
| 23274 | 1907  | 
using unity_coeff_ex[where l="l" and P="?P", simplified] by simp  | 
| 55921 | 1908  | 
also have "\<dots> = (\<exists>x::int. ?P' x)"  | 
1909  | 
using a_\<beta>[OF linp d lp] by simp  | 
|
| 50313 | 1910  | 
finally show ?thesis .  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1911  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
1912  | 
|
| 23274 | 1913  | 
lemma \<beta>:  | 
| 55999 | 1914  | 
assumes "iszlfm p"  | 
1915  | 
and "d_\<beta> p 1"  | 
|
1916  | 
and "d_\<delta> p d"  | 
|
| 55885 | 1917  | 
and dp: "d > 0"  | 
| 55999 | 1918  | 
    and "\<not> (\<exists>j::int \<in> {1 .. d}. \<exists>b \<in> Inum (a # bs) ` set (\<beta> p). x = b + j)"
 | 
| 55964 | 1919  | 
and p: "Ifm bbs (x # bs) p" (is "?P x")  | 
| 23274 | 1920  | 
shows "?P (x - d)"  | 
| 55999 | 1921  | 
using assms  | 
| 55885 | 1922  | 
proof (induct p rule: iszlfm.induct)  | 
1923  | 
case (5 c e)  | 
|
1924  | 
then have c1: "c = 1" and bn: "numbound0 e"  | 
|
1925  | 
by simp_all  | 
|
| 55964 | 1926  | 
with dp p c1 numbound0_I[OF bn,where b = "(x - d)" and b' = "x" and bs = "bs"] 5  | 
| 41807 | 1927  | 
show ?case by simp  | 
| 23274 | 1928  | 
next  | 
| 55885 | 1929  | 
case (6 c e)  | 
1930  | 
then have c1: "c = 1" and bn: "numbound0 e"  | 
|
1931  | 
by simp_all  | 
|
| 41807 | 1932  | 
with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] 6  | 
1933  | 
show ?case by simp  | 
|
| 23274 | 1934  | 
next  | 
| 55885 | 1935  | 
case (7 c e)  | 
| 55964 | 1936  | 
then have p: "Ifm bbs (x # bs) (Gt (CN 0 c e))" and c1: "c=1" and bn: "numbound0 e"  | 
| 55885 | 1937  | 
by simp_all  | 
| 41807 | 1938  | 
let ?e = "Inum (x # bs) e"  | 
| 60708 | 1939  | 
show ?case  | 
1940  | 
proof (cases "(x - d) + ?e > 0")  | 
|
1941  | 
case True  | 
|
1942  | 
then show ?thesis  | 
|
| 55885 | 1943  | 
using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] by simp  | 
| 60708 | 1944  | 
next  | 
1945  | 
case False  | 
|
| 55964 | 1946  | 
let ?v = "Neg e"  | 
1947  | 
have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))"  | 
|
1948  | 
by simp  | 
|
| 
57816
 
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
 
blanchet 
parents: 
57514 
diff
changeset
 | 
1949  | 
from 7(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]  | 
| 55964 | 1950  | 
    have nob: "\<not> (\<exists>j\<in> {1 ..d}. x = - ?e + j)"
 | 
| 55885 | 1951  | 
by auto  | 
| 60708 | 1952  | 
from False p have "x + ?e > 0 \<and> x + ?e \<le> d"  | 
| 55885 | 1953  | 
by (simp add: c1)  | 
1954  | 
then have "x + ?e \<ge> 1 \<and> x + ?e \<le> d"  | 
|
1955  | 
by simp  | 
|
| 55964 | 1956  | 
    then have "\<exists>j::int \<in> {1 .. d}. j = x + ?e"
 | 
| 55885 | 1957  | 
by simp  | 
| 55964 | 1958  | 
    then have "\<exists>j::int \<in> {1 .. d}. x = (- ?e + j)"
 | 
| 41807 | 1959  | 
by (simp add: algebra_simps)  | 
| 60708 | 1960  | 
with nob show ?thesis  | 
| 55885 | 1961  | 
by auto  | 
| 60708 | 1962  | 
qed  | 
| 23274 | 1963  | 
next  | 
| 55885 | 1964  | 
case (8 c e)  | 
1965  | 
then have p: "Ifm bbs (x # bs) (Ge (CN 0 c e))" and c1: "c = 1" and bn: "numbound0 e"  | 
|
| 50313 | 1966  | 
by simp_all  | 
| 55885 | 1967  | 
let ?e = "Inum (x # bs) e"  | 
| 60708 | 1968  | 
show ?case  | 
1969  | 
proof (cases "(x - d) + ?e \<ge> 0")  | 
|
1970  | 
case True  | 
|
1971  | 
then show ?thesis  | 
|
| 55885 | 1972  | 
using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"]  | 
1973  | 
by simp  | 
|
| 60708 | 1974  | 
next  | 
1975  | 
case False  | 
|
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1976  | 
let ?v = "Sub (C (- 1)) e"  | 
| 55885 | 1977  | 
have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))"  | 
1978  | 
by simp  | 
|
| 
57816
 
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
 
blanchet 
parents: 
57514 
diff
changeset
 | 
1979  | 
from 8(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]  | 
| 55885 | 1980  | 
    have nob: "\<not> (\<exists>j\<in> {1 ..d}. x =  - ?e - 1 + j)"
 | 
1981  | 
by auto  | 
|
| 60708 | 1982  | 
from False p have "x + ?e \<ge> 0 \<and> x + ?e < d"  | 
| 55885 | 1983  | 
by (simp add: c1)  | 
1984  | 
then have "x + ?e +1 \<ge> 1 \<and> x + ?e + 1 \<le> d"  | 
|
1985  | 
by simp  | 
|
| 55964 | 1986  | 
    then have "\<exists>j::int \<in> {1 .. d}. j = x + ?e + 1"
 | 
| 55885 | 1987  | 
by simp  | 
| 55964 | 1988  | 
    then have "\<exists>j::int \<in> {1 .. d}. x= - ?e - 1 + j"
 | 
| 55885 | 1989  | 
by (simp add: algebra_simps)  | 
| 60708 | 1990  | 
with nob show ?thesis  | 
| 55885 | 1991  | 
by simp  | 
| 60708 | 1992  | 
qed  | 
| 23274 | 1993  | 
next  | 
| 55885 | 1994  | 
case (3 c e)  | 
1995  | 
then  | 
|
1996  | 
have p: "Ifm bbs (x #bs) (Eq (CN 0 c e))" (is "?p x")  | 
|
| 55964 | 1997  | 
and c1: "c = 1"  | 
| 55885 | 1998  | 
and bn: "numbound0 e"  | 
1999  | 
by simp_all  | 
|
2000  | 
let ?e = "Inum (x # bs) e"  | 
|
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
2001  | 
let ?v="(Sub (C (- 1)) e)"  | 
| 55885 | 2002  | 
have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))"  | 
2003  | 
by simp  | 
|
| 55964 | 2004  | 
from p have "x= - ?e"  | 
2005  | 
by (simp add: c1) with 3(5)  | 
|
2006  | 
show ?case  | 
|
| 60708 | 2007  | 
using dp apply simp  | 
2008  | 
apply (erule ballE[where x="1"])  | 
|
2009  | 
apply (simp_all add:algebra_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"])  | 
|
2010  | 
done  | 
|
| 23274 | 2011  | 
next  | 
| 55885 | 2012  | 
case (4 c e)  | 
2013  | 
then  | 
|
| 55964 | 2014  | 
have p: "Ifm bbs (x # bs) (NEq (CN 0 c e))" (is "?p x")  | 
| 55885 | 2015  | 
and c1: "c = 1"  | 
2016  | 
and bn: "numbound0 e"  | 
|
2017  | 
by simp_all  | 
|
2018  | 
let ?e = "Inum (x # bs) e"  | 
|
2019  | 
let ?v="Neg e"  | 
|
| 55964 | 2020  | 
have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))"  | 
2021  | 
by simp  | 
|
| 60708 | 2022  | 
show ?case  | 
2023  | 
proof (cases "x - d + Inum ((x - d) # bs) e = 0")  | 
|
2024  | 
case False  | 
|
2025  | 
then show ?thesis by (simp add: c1)  | 
|
2026  | 
next  | 
|
2027  | 
case True  | 
|
| 55964 | 2028  | 
then have "x = - Inum ((x - d) # bs) e + d"  | 
| 55885 | 2029  | 
by simp  | 
2030  | 
then have "x = - Inum (a # bs) e + d"  | 
|
2031  | 
by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"])  | 
|
| 60708 | 2032  | 
with 4(5) show ?thesis  | 
| 55885 | 2033  | 
using dp by simp  | 
| 60708 | 2034  | 
qed  | 
| 50313 | 2035  | 
next  | 
| 55885 | 2036  | 
case (9 j c e)  | 
2037  | 
then  | 
|
2038  | 
have p: "Ifm bbs (x # bs) (Dvd j (CN 0 c e))" (is "?p x")  | 
|
2039  | 
and c1: "c = 1"  | 
|
2040  | 
and bn: "numbound0 e"  | 
|
2041  | 
by simp_all  | 
|
2042  | 
let ?e = "Inum (x # bs) e"  | 
|
2043  | 
from 9 have id: "j dvd d"  | 
|
2044  | 
by simp  | 
|
| 55964 | 2045  | 
from c1 have "?p x \<longleftrightarrow> j dvd (x + ?e)"  | 
| 55885 | 2046  | 
by simp  | 
| 55964 | 2047  | 
also have "\<dots> \<longleftrightarrow> j dvd x - d + ?e"  | 
| 55885 | 2048  | 
using zdvd_period[OF id, where x="x" and c="-1" and t="?e"]  | 
2049  | 
by simp  | 
|
2050  | 
finally show ?case  | 
|
2051  | 
using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p  | 
|
2052  | 
by simp  | 
|
| 23274 | 2053  | 
next  | 
| 55885 | 2054  | 
case (10 j c e)  | 
2055  | 
then  | 
|
| 55964 | 2056  | 
have p: "Ifm bbs (x # bs) (NDvd j (CN 0 c e))" (is "?p x")  | 
| 55885 | 2057  | 
and c1: "c = 1"  | 
2058  | 
and bn: "numbound0 e"  | 
|
2059  | 
by simp_all  | 
|
2060  | 
let ?e = "Inum (x # bs) e"  | 
|
2061  | 
from 10 have id: "j dvd d"  | 
|
2062  | 
by simp  | 
|
| 55964 | 2063  | 
from c1 have "?p x \<longleftrightarrow> \<not> j dvd (x + ?e)"  | 
| 55885 | 2064  | 
by simp  | 
| 55964 | 2065  | 
also have "\<dots> \<longleftrightarrow> \<not> j dvd x - d + ?e"  | 
| 55885 | 2066  | 
using zdvd_period[OF id, where x="x" and c="-1" and t="?e"]  | 
2067  | 
by simp  | 
|
2068  | 
finally show ?case  | 
|
2069  | 
using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p  | 
|
2070  | 
by simp  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
2071  | 
qed (auto simp add: numbound0_I[where bs="bs" and b="(x - d)" and b'="x"] gr0_conv_Suc)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2072  | 
|
| 50313 | 2073  | 
lemma \<beta>':  | 
| 23274 | 2074  | 
assumes lp: "iszlfm p"  | 
| 60708 | 2075  | 
and u: "d_\<beta> p 1"  | 
2076  | 
and d: "d_\<delta> p d"  | 
|
2077  | 
and dp: "d > 0"  | 
|
| 55964 | 2078  | 
  shows "\<forall>x. \<not> (\<exists>j::int \<in> {1 .. d}. \<exists>b \<in> set(\<beta> p). Ifm bbs ((Inum (a#bs) b + j) #bs) p) \<longrightarrow>
 | 
2079  | 
Ifm bbs (x # bs) p \<longrightarrow> Ifm bbs ((x - d) # bs) p" (is "\<forall>x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")  | 
|
| 55885 | 2080  | 
proof clarify  | 
| 50313 | 2081  | 
fix x  | 
| 60708 | 2082  | 
assume nb: "?b" and px: "?P x"  | 
| 55964 | 2083  | 
  then have nb2: "\<not> (\<exists>j::int \<in> {1 .. d}. \<exists>b \<in> Inum (a # bs) ` set (\<beta> p). x = b + j)"
 | 
| 23274 | 2084  | 
by auto  | 
| 60708 | 2085  | 
show "?P (x - d)" by (rule \<beta>[OF lp u d dp nb2 px])  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2086  | 
qed  | 
| 55885 | 2087  | 
|
2088  | 
lemma cpmi_eq:  | 
|
| 55999 | 2089  | 
fixes P P1 :: "int \<Rightarrow> bool"  | 
2090  | 
assumes "0 < D"  | 
|
2091  | 
and "\<exists>z. \<forall>x. x < z \<longrightarrow> P x = P1 x"  | 
|
2092  | 
    and "\<forall>x. \<not> (\<exists>j \<in> {1..D}. \<exists>b \<in> B. P (b + j)) \<longrightarrow> P x \<longrightarrow> P (x - D)"
 | 
|
2093  | 
and "\<forall>x k. P1 x = P1 (x - k * D)"  | 
|
2094  | 
  shows "(\<exists>x. P x) \<longleftrightarrow> (\<exists>j \<in> {1..D}. P1 j) \<or> (\<exists>j \<in> {1..D}. \<exists>b \<in> B. P (b + j))"
 | 
|
2095  | 
apply (insert assms)  | 
|
2096  | 
apply (rule iffI)  | 
|
| 55885 | 2097  | 
prefer 2  | 
| 55981 | 2098  | 
apply (drule minusinfinity)  | 
| 55885 | 2099  | 
apply assumption+  | 
| 55981 | 2100  | 
apply fastforce  | 
| 55885 | 2101  | 
apply clarsimp  | 
| 55981 | 2102  | 
apply (subgoal_tac "\<And>k. 0 \<le> k \<Longrightarrow> \<forall>x. P x \<longrightarrow> P (x - k * D)")  | 
2103  | 
apply (frule_tac x = x and z=z in decr_lemma)  | 
|
2104  | 
apply (subgoal_tac "P1 (x - (\<bar>x - z\<bar> + 1) * D)")  | 
|
| 55885 | 2105  | 
prefer 2  | 
| 55981 | 2106  | 
apply (subgoal_tac "0 \<le> \<bar>x - z\<bar> + 1")  | 
| 55885 | 2107  | 
prefer 2 apply arith  | 
2108  | 
apply fastforce  | 
|
| 55981 | 2109  | 
apply (drule (1) periodic_finite_ex)  | 
| 55885 | 2110  | 
apply blast  | 
| 55981 | 2111  | 
apply (blast dest: decr_mult_lemma)  | 
| 55885 | 2112  | 
done  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2113  | 
|
| 23274 | 2114  | 
theorem cp_thm:  | 
2115  | 
assumes lp: "iszlfm p"  | 
|
| 55885 | 2116  | 
and u: "d_\<beta> p 1"  | 
2117  | 
and d: "d_\<delta> p d"  | 
|
2118  | 
and dp: "d > 0"  | 
|
| 55999 | 2119  | 
shows "(\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow>  | 
| 55964 | 2120  | 
    (\<exists>j \<in> {1.. d}. Ifm bbs (j # bs) (minusinf p) \<or>
 | 
2121  | 
(\<exists>b \<in> set (\<beta> p). Ifm bbs ((Inum (i # bs) b + j) # bs) p))"  | 
|
| 55999 | 2122  | 
(is "(\<exists>x. ?P x) \<longleftrightarrow> (\<exists>j \<in> ?D. ?M j \<or> (\<exists>b \<in> ?B. ?P (?I b + j)))")  | 
| 55885 | 2123  | 
proof -  | 
| 50313 | 2124  | 
from minusinf_inf[OF lp u]  | 
| 55999 | 2125  | 
have th: "\<exists>z. \<forall>x<z. ?P x = ?M x"  | 
| 55885 | 2126  | 
by blast  | 
| 55964 | 2127  | 
  let ?B' = "{?I b | b. b \<in> ?B}"
 | 
2128  | 
have BB': "(\<exists>j\<in>?D. \<exists>b \<in> ?B. ?P (?I b + j)) \<longleftrightarrow> (\<exists>j \<in> ?D. \<exists>b \<in> ?B'. ?P (b + j))"  | 
|
| 55885 | 2129  | 
by auto  | 
| 55964 | 2130  | 
then have th2: "\<forall>x. \<not> (\<exists>j \<in> ?D. \<exists>b \<in> ?B'. ?P (b + j)) \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)"  | 
| 23274 | 2131  | 
using \<beta>'[OF lp u d dp, where a="i" and bbs = "bbs"] by blast  | 
2132  | 
from minusinf_repeats[OF d lp]  | 
|
| 55885 | 2133  | 
have th3: "\<forall>x k. ?M x = ?M (x-k*d)"  | 
2134  | 
by simp  | 
|
2135  | 
from cpmi_eq[OF dp th th2 th3] BB' show ?thesis  | 
|
2136  | 
by blast  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2137  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2138  | 
|
| 67123 | 2139  | 
text \<open>Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff.\<close>  | 
2140  | 
||
| 50313 | 2141  | 
lemma mirror_ex:  | 
| 55999 | 2142  | 
assumes "iszlfm p"  | 
2143  | 
shows "(\<exists>x. Ifm bbs (x#bs) (mirror p)) \<longleftrightarrow> (\<exists>x. Ifm bbs (x#bs) p)"  | 
|
| 50313 | 2144  | 
(is "(\<exists>x. ?I x ?mp) = (\<exists>x. ?I x p)")  | 
| 55964 | 2145  | 
proof auto  | 
2146  | 
fix x  | 
|
2147  | 
assume "?I x ?mp"  | 
|
2148  | 
then have "?I (- x) p"  | 
|
| 55999 | 2149  | 
using mirror[OF assms] by blast  | 
| 55964 | 2150  | 
then show "\<exists>x. ?I x p"  | 
2151  | 
by blast  | 
|
| 23274 | 2152  | 
next  | 
| 55964 | 2153  | 
fix x  | 
2154  | 
assume "?I x p"  | 
|
2155  | 
then have "?I (- x) ?mp"  | 
|
| 55999 | 2156  | 
using mirror[OF assms, where x="- x", symmetric] by auto  | 
| 55964 | 2157  | 
then show "\<exists>x. ?I x ?mp"  | 
2158  | 
by blast  | 
|
| 23274 | 2159  | 
qed  | 
| 24349 | 2160  | 
|
| 50313 | 2161  | 
lemma cp_thm':  | 
| 55999 | 2162  | 
assumes "iszlfm p"  | 
2163  | 
and "d_\<beta> p 1"  | 
|
2164  | 
and "d_\<delta> p d"  | 
|
2165  | 
and "d > 0"  | 
|
| 55964 | 2166  | 
shows "(\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow>  | 
2167  | 
    ((\<exists>j\<in> {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \<or>
 | 
|
2168  | 
      (\<exists>j\<in> {1.. d}. \<exists>b\<in> (Inum (i#bs)) ` set (\<beta> p). Ifm bbs ((b + j) # bs) p))"
 | 
|
| 55999 | 2169  | 
using cp_thm[OF assms,where i="i"] by auto  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2170  | 
|
| 50313 | 2171  | 
definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int"  | 
2172  | 
where  | 
|
| 55964 | 2173  | 
"unit p =  | 
2174  | 
(let  | 
|
2175  | 
p' = zlfm p;  | 
|
2176  | 
l = \<zeta> p';  | 
|
2177  | 
q = And (Dvd l (CN 0 1 (C 0))) (a_\<beta> p' l);  | 
|
2178  | 
d = \<delta> q;  | 
|
2179  | 
B = remdups (map simpnum (\<beta> q));  | 
|
2180  | 
a = remdups (map simpnum (\<alpha> q))  | 
|
2181  | 
in if length B \<le> length a then (q, B, d) else (mirror q, a, d))"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2182  | 
|
| 50313 | 2183  | 
lemma unit:  | 
2184  | 
assumes qf: "qfree p"  | 
|
| 60708 | 2185  | 
fixes q B d  | 
2186  | 
assumes qBd: "unit p = (q, B, d)"  | 
|
2187  | 
shows "((\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) q)) \<and>  | 
|
| 55964 | 2188  | 
(Inum (i # bs)) ` set B = (Inum (i # bs)) ` set (\<beta> q) \<and> d_\<beta> q 1 \<and> d_\<delta> q d \<and> d > 0 \<and>  | 
2189  | 
iszlfm q \<and> (\<forall>b\<in> set B. numbound0 b)"  | 
|
| 50313 | 2190  | 
proof -  | 
2191  | 
let ?I = "\<lambda>x p. Ifm bbs (x#bs) p"  | 
|
| 23274 | 2192  | 
let ?p' = "zlfm p"  | 
2193  | 
let ?l = "\<zeta> ?p'"  | 
|
| 50252 | 2194  | 
let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\<beta> ?p' ?l)"  | 
| 23274 | 2195  | 
let ?d = "\<delta> ?q"  | 
2196  | 
let ?B = "set (\<beta> ?q)"  | 
|
2197  | 
let ?B'= "remdups (map simpnum (\<beta> ?q))"  | 
|
2198  | 
let ?A = "set (\<alpha> ?q)"  | 
|
2199  | 
let ?A'= "remdups (map simpnum (\<alpha> ?q))"  | 
|
| 50313 | 2200  | 
from conjunct1[OF zlfm_I[OF qf, where bs="bs"]]  | 
2201  | 
have pp': "\<forall>i. ?I i ?p' = ?I i p" by auto  | 
|
| 23274 | 2202  | 
from conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]  | 
| 50313 | 2203  | 
have lp': "iszlfm ?p'" .  | 
| 50252 | 2204  | 
from lp' \<zeta>[where p="?p'"] have lp: "?l >0" and dl: "d_\<beta> ?p' ?l" by auto  | 
2205  | 
from a_\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp' dl lp] pp'  | 
|
| 50313 | 2206  | 
have pq_ex:"(\<exists>(x::int). ?I x p) = (\<exists>x. ?I x ?q)" by simp  | 
| 50252 | 2207  | 
from lp' lp a_\<beta>[OF lp' dl lp] have lq:"iszlfm ?q" and uq: "d_\<beta> ?q 1" by auto  | 
2208  | 
from \<delta>[OF lq] have dp:"?d >0" and dd: "d_\<delta> ?q ?d" by blast+  | 
|
| 50313 | 2209  | 
let ?N = "\<lambda>t. Inum (i#bs) t"  | 
| 55981 | 2210  | 
have "?N ` set ?B' = ((?N \<circ> simpnum) ` ?B)"  | 
2211  | 
by auto  | 
|
2212  | 
also have "\<dots> = ?N ` ?B"  | 
|
2213  | 
using simpnum_ci[where bs="i#bs"] by auto  | 
|
| 23274 | 2214  | 
finally have BB': "?N ` set ?B' = ?N ` ?B" .  | 
| 55981 | 2215  | 
have "?N ` set ?A' = ((?N \<circ> simpnum) ` ?A)"  | 
2216  | 
by auto  | 
|
2217  | 
also have "\<dots> = ?N ` ?A"  | 
|
2218  | 
using simpnum_ci[where bs="i#bs"] by auto  | 
|
| 23274 | 2219  | 
finally have AA': "?N ` set ?A' = ?N ` ?A" .  | 
| 50313 | 2220  | 
from \<beta>_numbound0[OF lq] have B_nb:"\<forall>b\<in> set ?B'. numbound0 b"  | 
| 23274 | 2221  | 
by (simp add: simpnum_numbound0)  | 
| 50313 | 2222  | 
from \<alpha>_l[OF lq] have A_nb: "\<forall>b\<in> set ?A'. numbound0 b"  | 
| 23274 | 2223  | 
by (simp add: simpnum_numbound0)  | 
| 60708 | 2224  | 
show ?thesis  | 
2225  | 
proof (cases "length ?B' \<le> length ?A'")  | 
|
2226  | 
case True  | 
|
| 55981 | 2227  | 
then have q: "q = ?q" and "B = ?B'" and d: "d = ?d"  | 
| 23274 | 2228  | 
using qBd by (auto simp add: Let_def unit_def)  | 
| 55981 | 2229  | 
with BB' B_nb  | 
2230  | 
have b: "?N ` (set B) = ?N ` set (\<beta> q)" and bn: "\<forall>b\<in> set B. numbound0 b"  | 
|
2231  | 
by simp_all  | 
|
| 60708 | 2232  | 
with pq_ex dp uq dd lq q d show ?thesis  | 
| 55981 | 2233  | 
by simp  | 
| 60708 | 2234  | 
next  | 
2235  | 
case False  | 
|
| 55885 | 2236  | 
then have q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"  | 
| 23274 | 2237  | 
using qBd by (auto simp add: Let_def unit_def)  | 
| 50313 | 2238  | 
with AA' mirror_\<alpha>_\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)"  | 
2239  | 
and bn: "\<forall>b\<in> set B. numbound0 b" by simp_all  | 
|
2240  | 
from mirror_ex[OF lq] pq_ex q  | 
|
| 55981 | 2241  | 
have pqm_eq:"(\<exists>(x::int). ?I x p) = (\<exists>(x::int). ?I x q)"  | 
2242  | 
by simp  | 
|
| 23274 | 2243  | 
from lq uq q mirror_l[where p="?q"]  | 
| 55981 | 2244  | 
have lq': "iszlfm q" and uq: "d_\<beta> q 1"  | 
2245  | 
by auto  | 
|
2246  | 
from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq: "d_\<delta> q d"  | 
|
2247  | 
by auto  | 
|
| 60708 | 2248  | 
from pqm_eq b bn uq lq' dp dq q dp d show ?thesis  | 
| 55981 | 2249  | 
by simp  | 
| 60708 | 2250  | 
qed  | 
| 23274 | 2251  | 
qed  | 
| 50313 | 2252  | 
|
2253  | 
||
| 70091 | 2254  | 
subsection \<open>Cooper's Algorithm\<close>  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2255  | 
|
| 55981 | 2256  | 
definition cooper :: "fm \<Rightarrow> fm"  | 
2257  | 
where  | 
|
| 50313 | 2258  | 
"cooper p =  | 
2259  | 
(let  | 
|
2260  | 
(q, B, d) = unit p;  | 
|
2261  | 
js = [1..d];  | 
|
2262  | 
mq = simpfm (minusinf q);  | 
|
2263  | 
md = evaldjf (\<lambda>j. simpfm (subst0 (C j) mq)) js  | 
|
2264  | 
in  | 
|
2265  | 
if md = T then T  | 
|
2266  | 
else  | 
|
2267  | 
(let  | 
|
2268  | 
qd = evaldjf (\<lambda>(b, j). simpfm (subst0 (Add b (C j)) q)) [(b, j). b \<leftarrow> B, j \<leftarrow> js]  | 
|
2269  | 
in decr (disj md qd)))"  | 
|
2270  | 
||
2271  | 
lemma cooper:  | 
|
2272  | 
assumes qf: "qfree p"  | 
|
| 60708 | 2273  | 
shows "(\<exists>x. Ifm bbs (x#bs) p) = Ifm bbs bs (cooper p) \<and> qfree (cooper p)"  | 
2274  | 
(is "?lhs = ?rhs \<and> _")  | 
|
| 50313 | 2275  | 
proof -  | 
2276  | 
let ?I = "\<lambda>x p. Ifm bbs (x#bs) p"  | 
|
| 23274 | 2277  | 
let ?q = "fst (unit p)"  | 
2278  | 
let ?B = "fst (snd(unit p))"  | 
|
2279  | 
let ?d = "snd (snd (unit p))"  | 
|
| 41836 | 2280  | 
let ?js = "[1..?d]"  | 
| 23274 | 2281  | 
let ?mq = "minusinf ?q"  | 
2282  | 
let ?smq = "simpfm ?mq"  | 
|
| 50313 | 2283  | 
let ?md = "evaldjf (\<lambda>j. simpfm (subst0 (C j) ?smq)) ?js"  | 
| 26934 | 2284  | 
fix i  | 
| 50313 | 2285  | 
let ?N = "\<lambda>t. Inum (i#bs) t"  | 
| 24336 | 2286  | 
let ?Bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]"  | 
| 50313 | 2287  | 
let ?qd = "evaldjf (\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs"  | 
| 23274 | 2288  | 
have qbf:"unit p = (?q,?B,?d)" by simp  | 
| 55981 | 2289  | 
from unit[OF qf qbf]  | 
2290  | 
have pq_ex: "(\<exists>(x::int). ?I x p) \<longleftrightarrow> (\<exists>(x::int). ?I x ?q)"  | 
|
2291  | 
and B: "?N ` set ?B = ?N ` set (\<beta> ?q)"  | 
|
2292  | 
and uq: "d_\<beta> ?q 1"  | 
|
2293  | 
and dd: "d_\<delta> ?q ?d"  | 
|
2294  | 
and dp: "?d > 0"  | 
|
2295  | 
and lq: "iszlfm ?q"  | 
|
2296  | 
and Bn: "\<forall>b\<in> set ?B. numbound0 b"  | 
|
2297  | 
by auto  | 
|
| 23274 | 2298  | 
from zlin_qfree[OF lq] have qfq: "qfree ?q" .  | 
| 55921 | 2299  | 
from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq" .  | 
| 55981 | 2300  | 
have jsnb: "\<forall>j \<in> set ?js. numbound0 (C j)"  | 
2301  | 
by simp  | 
|
| 55885 | 2302  | 
then have "\<forall>j\<in> set ?js. bound0 (subst0 (C j) ?smq)"  | 
| 23274 | 2303  | 
by (auto simp only: subst0_bound0[OF qfmq])  | 
| 55885 | 2304  | 
then have th: "\<forall>j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"  | 
| 23274 | 2305  | 
by (auto simp add: simpfm_bound0)  | 
| 55981 | 2306  | 
from evaldjf_bound0[OF th] have mdb: "bound0 ?md"  | 
2307  | 
by simp  | 
|
| 50313 | 2308  | 
from Bn jsnb have "\<forall>(b,j) \<in> set ?Bjs. numbound0 (Add b (C j))"  | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23515 
diff
changeset
 | 
2309  | 
by simp  | 
| 55885 | 2310  | 
then have "\<forall>(b,j) \<in> set ?Bjs. bound0 (subst0 (Add b (C j)) ?q)"  | 
| 23274 | 2311  | 
using subst0_bound0[OF qfq] by blast  | 
| 55885 | 2312  | 
then have "\<forall>(b,j) \<in> set ?Bjs. bound0 (simpfm (subst0 (Add b (C j)) ?q))"  | 
| 55981 | 2313  | 
using simpfm_bound0 by blast  | 
| 55885 | 2314  | 
then have th': "\<forall>x \<in> set ?Bjs. bound0 ((\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) x)"  | 
| 50313 | 2315  | 
by auto  | 
| 55981 | 2316  | 
from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd"  | 
2317  | 
by simp  | 
|
2318  | 
from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)"  | 
|
2319  | 
unfolding disj_def by (cases "?md = T \<or> ?qd = T") simp_all  | 
|
| 23274 | 2320  | 
from trans [OF pq_ex cp_thm'[OF lq uq dd dp,where i="i"]] B  | 
| 55981 | 2321  | 
  have "?lhs \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq \<or> (\<exists>b \<in> ?N ` set ?B. Ifm bbs ((b + j) # bs) ?q))"
 | 
| 55921 | 2322  | 
by auto  | 
| 55981 | 2323  | 
  also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq \<or> (\<exists>b \<in> set ?B. Ifm bbs ((?N b + j) # bs) ?q))"
 | 
| 55921 | 2324  | 
by simp  | 
| 55981 | 2325  | 
  also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq ) \<or>
 | 
2326  | 
      (\<exists>j\<in> {1.. ?d}. \<exists>b \<in> set ?B. Ifm bbs ((?N (Add b (C j))) # bs) ?q)"
 | 
|
| 50313 | 2327  | 
by (simp only: Inum.simps) blast  | 
| 55981 | 2328  | 
  also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?smq) \<or>
 | 
2329  | 
      (\<exists>j \<in> {1.. ?d}. \<exists>b \<in> set ?B. Ifm bbs ((?N (Add b (C j))) # bs) ?q)"
 | 
|
| 50313 | 2330  | 
by (simp add: simpfm)  | 
| 55981 | 2331  | 
also have "\<dots> \<longleftrightarrow> (\<exists>j\<in> set ?js. (\<lambda>j. ?I i (simpfm (subst0 (C j) ?smq))) j) \<or>  | 
2332  | 
(\<exists>j\<in> set ?js. \<exists>b\<in> set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q)"  | 
|
| 41836 | 2333  | 
by (simp only: simpfm subst0_I[OF qfmq] set_upto) auto  | 
| 55981 | 2334  | 
also have "\<dots> \<longleftrightarrow> ?I i (evaldjf (\<lambda>j. simpfm (subst0 (C j) ?smq)) ?js) \<or>  | 
2335  | 
(\<exists>j\<in> set ?js. \<exists>b\<in> set ?B. ?I i (subst0 (Add b (C j)) ?q))"  | 
|
| 50313 | 2336  | 
by (simp only: evaldjf_ex subst0_I[OF qfq])  | 
| 55981 | 2337  | 
also have "\<dots> \<longleftrightarrow> ?I i ?md \<or>  | 
2338  | 
(\<exists>(b,j) \<in> set ?Bjs. (\<lambda>(b,j). ?I i (simpfm (subst0 (Add b (C j)) ?q))) (b,j))"  | 
|
| 50313 | 2339  | 
by (simp only: simpfm set_concat set_map concat_map_singleton UN_simps) blast  | 
| 55981 | 2340  | 
also have "\<dots> \<longleftrightarrow> ?I i ?md \<or> ?I i (evaldjf (\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs)"  | 
| 50313 | 2341  | 
by (simp only: evaldjf_ex[where bs="i#bs" and f="\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)" and ps="?Bjs"])  | 
2342  | 
(auto simp add: split_def)  | 
|
| 55981 | 2343  | 
finally have mdqd: "?lhs \<longleftrightarrow> ?I i ?md \<or> ?I i ?qd"  | 
| 55921 | 2344  | 
by simp  | 
| 55981 | 2345  | 
also have "\<dots> \<longleftrightarrow> ?I i (disj ?md ?qd)"  | 
| 55921 | 2346  | 
by (simp add: disj)  | 
| 55981 | 2347  | 
also have "\<dots> \<longleftrightarrow> Ifm bbs bs (decr (disj ?md ?qd))"  | 
| 55921 | 2348  | 
by (simp only: decr [OF mdqdb])  | 
| 55981 | 2349  | 
finally have mdqd2: "?lhs \<longleftrightarrow> Ifm bbs bs (decr (disj ?md ?qd))" .  | 
| 60708 | 2350  | 
show ?thesis  | 
2351  | 
proof (cases "?md = T")  | 
|
2352  | 
case True  | 
|
| 55921 | 2353  | 
then have cT: "cooper p = T"  | 
| 23274 | 2354  | 
by (simp only: cooper_def unit_def split_def Let_def if_True) simp  | 
| 60708 | 2355  | 
from True have lhs: "?lhs"  | 
| 55921 | 2356  | 
using mdqd by simp  | 
| 60708 | 2357  | 
from True have "?rhs"  | 
| 55921 | 2358  | 
by (simp add: cooper_def unit_def split_def)  | 
| 60708 | 2359  | 
with lhs cT show ?thesis  | 
| 55981 | 2360  | 
by simp  | 
| 60708 | 2361  | 
next  | 
2362  | 
case False  | 
|
| 55921 | 2363  | 
then have "cooper p = decr (disj ?md ?qd)"  | 
| 50313 | 2364  | 
by (simp only: cooper_def unit_def split_def Let_def if_False)  | 
| 60708 | 2365  | 
with mdqd2 decr_qf[OF mdqdb] show ?thesis  | 
| 55921 | 2366  | 
by simp  | 
| 60708 | 2367  | 
qed  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2368  | 
qed  | 
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2369  | 
|
| 55921 | 2370  | 
definition pa :: "fm \<Rightarrow> fm"  | 
2371  | 
where "pa p = qelim (prep p) cooper"  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2372  | 
|
| 55921 | 2373  | 
theorem mirqe: "Ifm bbs bs (pa p) = Ifm bbs bs p \<and> qfree (pa p)"  | 
| 23274 | 2374  | 
using qelim_ci cooper prep by (auto simp add: pa_def)  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2375  | 
|
| 
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2376  | 
|
| 70091 | 2377  | 
subsection \<open>Setup\<close>  | 
| 27456 | 2378  | 
|
| 60533 | 2379  | 
oracle linzqe_oracle = \<open>  | 
| 27456 | 2380  | 
let  | 
2381  | 
||
| 55814 | 2382  | 
fun num_of_term vs (t as Free (xn, xT)) =  | 
| 67399 | 2383  | 
(case AList.lookup (=) vs t of  | 
| 55814 | 2384  | 
NONE => error "Variable not found in the list!"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2385  | 
      | SOME n => @{code Bound} (@{code nat_of_integer} n))
 | 
| 69597 | 2386  | 
  | num_of_term vs \<^term>\<open>0::int\<close> = @{code C} (@{code int_of_integer} 0)
 | 
2387  | 
  | num_of_term vs \<^term>\<open>1::int\<close> = @{code C} (@{code int_of_integer} 1)
 | 
|
2388  | 
  | num_of_term vs \<^term>\<open>- 1::int\<close> = @{code C} (@{code int_of_integer} (~ 1))
 | 
|
2389  | 
| num_of_term vs (\<^term>\<open>numeral :: _ \<Rightarrow> int\<close> $ t) =  | 
|
| 62342 | 2390  | 
      @{code C} (@{code int_of_integer} (HOLogic.dest_numeral t))
 | 
| 69597 | 2391  | 
| num_of_term vs (\<^term>\<open>- numeral :: _ \<Rightarrow> int\<close> $ t) =  | 
| 62342 | 2392  | 
      @{code C} (@{code int_of_integer} (~(HOLogic.dest_numeral t)))
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2393  | 
  | num_of_term vs (Bound i) = @{code Bound} (@{code nat_of_integer} i)
 | 
| 69597 | 2394  | 
  | num_of_term vs (\<^term>\<open>uminus :: int \<Rightarrow> int\<close> $ t') = @{code Neg} (num_of_term vs t')
 | 
2395  | 
| num_of_term vs (\<^term>\<open>(+) :: int \<Rightarrow> int \<Rightarrow> int\<close> $ t1 $ t2) =  | 
|
| 27456 | 2396  | 
      @{code Add} (num_of_term vs t1, num_of_term vs t2)
 | 
| 69597 | 2397  | 
| num_of_term vs (\<^term>\<open>(-) :: int \<Rightarrow> int \<Rightarrow> int\<close> $ t1 $ t2) =  | 
| 27456 | 2398  | 
      @{code Sub} (num_of_term vs t1, num_of_term vs t2)
 | 
| 69597 | 2399  | 
| num_of_term vs (\<^term>\<open>(*) :: int \<Rightarrow> int \<Rightarrow> int\<close> $ t1 $ t2) =  | 
| 55814 | 2400  | 
(case try HOLogic.dest_number t1 of  | 
2401  | 
        SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t2)
 | 
|
2402  | 
| NONE =>  | 
|
2403  | 
(case try HOLogic.dest_number t2 of  | 
|
2404  | 
            SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t1)
 | 
|
2405  | 
| NONE => error "num_of_term: unsupported multiplication"))  | 
|
| 69597 | 2406  | 
  | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term \<^context> t);
 | 
| 27456 | 2407  | 
|
| 69597 | 2408  | 
fun fm_of_term ps vs \<^term>\<open>True\<close> = @{code T}
 | 
2409  | 
  | fm_of_term ps vs \<^term>\<open>False\<close> = @{code F}
 | 
|
2410  | 
| fm_of_term ps vs (\<^term>\<open>(<) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ t1 $ t2) =  | 
|
| 27456 | 2411  | 
      @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | 
| 69597 | 2412  | 
| fm_of_term ps vs (\<^term>\<open>(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ t1 $ t2) =  | 
| 27456 | 2413  | 
      @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | 
| 69597 | 2414  | 
| fm_of_term ps vs (\<^term>\<open>(=) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ t1 $ t2) =  | 
| 50313 | 2415  | 
      @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | 
| 69597 | 2416  | 
| fm_of_term ps vs (\<^term>\<open>(dvd) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ t1 $ t2) =  | 
| 55814 | 2417  | 
(case try HOLogic.dest_number t1 of  | 
2418  | 
        SOME (_, i) => @{code Dvd} (@{code int_of_integer} i, num_of_term vs t2)
 | 
|
2419  | 
| NONE => error "num_of_term: unsupported dvd")  | 
|
| 69597 | 2420  | 
| fm_of_term ps vs (\<^term>\<open>(=) :: bool \<Rightarrow> bool \<Rightarrow> bool\<close> $ t1 $ t2) =  | 
| 27456 | 2421  | 
      @{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 69597 | 2422  | 
| fm_of_term ps vs (\<^term>\<open>HOL.conj\<close> $ t1 $ t2) =  | 
| 27456 | 2423  | 
      @{code And} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 69597 | 2424  | 
| fm_of_term ps vs (\<^term>\<open>HOL.disj\<close> $ t1 $ t2) =  | 
| 27456 | 2425  | 
      @{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 69597 | 2426  | 
| fm_of_term ps vs (\<^term>\<open>HOL.implies\<close> $ t1 $ t2) =  | 
| 27456 | 2427  | 
      @{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 69597 | 2428  | 
| fm_of_term ps vs (\<^term>\<open>Not\<close> $ t') =  | 
| 27456 | 2429  | 
      @{code NOT} (fm_of_term ps vs t')
 | 
| 69597 | 2430  | 
| fm_of_term ps vs (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (xn, xT, p)) =  | 
| 27456 | 2431  | 
let  | 
| 42284 | 2432  | 
val (xn', p') = Syntax_Trans.variant_abs (xn, xT, p); (* FIXME !? *)  | 
| 27456 | 2433  | 
val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs;  | 
2434  | 
      in @{code E} (fm_of_term ps vs' p) end
 | 
|
| 69597 | 2435  | 
| fm_of_term ps vs (Const (\<^const_name>\<open>All\<close>, _) $ Abs (xn, xT, p)) =  | 
| 27456 | 2436  | 
let  | 
| 42284 | 2437  | 
val (xn', p') = Syntax_Trans.variant_abs (xn, xT, p); (* FIXME !? *)  | 
| 27456 | 2438  | 
val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs;  | 
2439  | 
      in @{code A} (fm_of_term ps vs' p) end
 | 
|
| 69597 | 2440  | 
  | fm_of_term ps vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term \<^context> t);
 | 
| 23515 | 2441  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2442  | 
fun term_of_num vs (@{code C} i) = HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2443  | 
  | term_of_num vs (@{code Bound} n) =
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2444  | 
let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2445  | 
        val q = @{code integer_of_nat} n
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2446  | 
in fst (the (find_first (fn (_, m) => q = m) vs)) end  | 
| 69597 | 2447  | 
  | term_of_num vs (@{code Neg} t') = \<^term>\<open>uminus :: int \<Rightarrow> int\<close> $ term_of_num vs t'
 | 
2448  | 
  | term_of_num vs (@{code Add} (t1, t2)) = \<^term>\<open>(+) :: int \<Rightarrow> int \<Rightarrow> int\<close> $
 | 
|
| 27456 | 2449  | 
term_of_num vs t1 $ term_of_num vs t2  | 
| 69597 | 2450  | 
  | term_of_num vs (@{code Sub} (t1, t2)) = \<^term>\<open>(-) :: int \<Rightarrow> int \<Rightarrow> int\<close> $
 | 
| 27456 | 2451  | 
term_of_num vs t1 $ term_of_num vs t2  | 
| 69597 | 2452  | 
  | term_of_num vs (@{code Mul} (i, t2)) = \<^term>\<open>(*) :: int \<Rightarrow> int \<Rightarrow> int\<close> $
 | 
| 27456 | 2453  | 
      term_of_num vs (@{code C} i) $ term_of_num vs t2
 | 
| 55814 | 2454  | 
  | term_of_num vs (@{code CN} (n, i, t)) =
 | 
2455  | 
      term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t));
 | 
|
| 27456 | 2456  | 
|
| 69597 | 2457  | 
fun term_of_fm ps vs @{code T} = \<^term>\<open>True\<close>
 | 
2458  | 
  | term_of_fm ps vs @{code F} = \<^term>\<open>False\<close>
 | 
|
| 27456 | 2459  | 
  | term_of_fm ps vs (@{code Lt} t) =
 | 
| 69597 | 2460  | 
\<^term>\<open>(<) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs t $ \<^term>\<open>0::int\<close>  | 
| 27456 | 2461  | 
  | term_of_fm ps vs (@{code Le} t) =
 | 
| 69597 | 2462  | 
\<^term>\<open>(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs t $ \<^term>\<open>0::int\<close>  | 
| 27456 | 2463  | 
  | term_of_fm ps vs (@{code Gt} t) =
 | 
| 69597 | 2464  | 
\<^term>\<open>(<) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ \<^term>\<open>0::int\<close> $ term_of_num vs t  | 
| 27456 | 2465  | 
  | term_of_fm ps vs (@{code Ge} t) =
 | 
| 69597 | 2466  | 
\<^term>\<open>(\<le>) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ \<^term>\<open>0::int\<close> $ term_of_num vs t  | 
| 27456 | 2467  | 
  | term_of_fm ps vs (@{code Eq} t) =
 | 
| 69597 | 2468  | 
\<^term>\<open>(=) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs t $ \<^term>\<open>0::int\<close>  | 
| 27456 | 2469  | 
  | term_of_fm ps vs (@{code NEq} t) =
 | 
2470  | 
      term_of_fm ps vs (@{code NOT} (@{code Eq} t))
 | 
|
2471  | 
  | term_of_fm ps vs (@{code Dvd} (i, t)) =
 | 
|
| 69597 | 2472  | 
      \<^term>\<open>(dvd) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs (@{code C} i) $ term_of_num vs t
 | 
| 27456 | 2473  | 
  | term_of_fm ps vs (@{code NDvd} (i, t)) =
 | 
2474  | 
      term_of_fm ps vs (@{code NOT} (@{code Dvd} (i, t)))
 | 
|
2475  | 
  | term_of_fm ps vs (@{code NOT} t') =
 | 
|
2476  | 
HOLogic.Not $ term_of_fm ps vs t'  | 
|
2477  | 
  | term_of_fm ps vs (@{code And} (t1, t2)) =
 | 
|
2478  | 
HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2  | 
|
2479  | 
  | term_of_fm ps vs (@{code Or} (t1, t2)) =
 | 
|
2480  | 
HOLogic.disj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2  | 
|
2481  | 
  | term_of_fm ps vs (@{code Imp} (t1, t2)) =
 | 
|
2482  | 
HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2  | 
|
2483  | 
  | term_of_fm ps vs (@{code Iff} (t1, t2)) =
 | 
|
| 69597 | 2484  | 
\<^term>\<open>(=) :: bool \<Rightarrow> bool \<Rightarrow> bool\<close> $ term_of_fm ps vs t1 $ term_of_fm ps vs t2  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2485  | 
  | term_of_fm ps vs (@{code Closed} n) =
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2486  | 
let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2487  | 
        val q = @{code integer_of_nat} n
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
50313 
diff
changeset
 | 
2488  | 
in (fst o the) (find_first (fn (_, m) => m = q) ps) end  | 
| 29788 | 2489  | 
  | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code NOT} (@{code Closed} n));
 | 
| 27456 | 2490  | 
|
2491  | 
fun term_bools acc t =  | 
|
2492  | 
let  | 
|
| 55814 | 2493  | 
val is_op =  | 
| 69597 | 2494  | 
member (=) [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>, \<^term>\<open>HOL.implies\<close>,  | 
2495  | 
\<^term>\<open>(=) :: bool \<Rightarrow> _\<close>,  | 
|
2496  | 
\<^term>\<open>(=) :: int \<Rightarrow> _\<close>, \<^term>\<open>(<) :: int \<Rightarrow> _\<close>,  | 
|
2497  | 
\<^term>\<open>(\<le>) :: int \<Rightarrow> _\<close>, \<^term>\<open>Not\<close>, \<^term>\<open>All :: (int \<Rightarrow> _) \<Rightarrow> _\<close>,  | 
|
2498  | 
\<^term>\<open>Ex :: (int \<Rightarrow> _) \<Rightarrow> _\<close>, \<^term>\<open>True\<close>, \<^term>\<open>False\<close>]  | 
|
| 50313 | 2499  | 
fun is_ty t = not (fastype_of t = HOLogic.boolT)  | 
| 55814 | 2500  | 
in  | 
2501  | 
(case t of  | 
|
2502  | 
(l as f $ a) $ b =>  | 
|
2503  | 
if is_ty t orelse is_op t then term_bools (term_bools acc l) b  | 
|
| 69214 | 2504  | 
else insert (op aconv) t acc  | 
| 55814 | 2505  | 
| f $ a =>  | 
2506  | 
if is_ty t orelse is_op t then term_bools (term_bools acc f) a  | 
|
| 69214 | 2507  | 
else insert (op aconv) t acc  | 
| 42284 | 2508  | 
| Abs p => term_bools acc (snd (Syntax_Trans.variant_abs p)) (* FIXME !? *)  | 
| 69214 | 2509  | 
| _ => if is_ty t orelse is_op t then acc else insert (op aconv) t acc)  | 
| 27456 | 2510  | 
end;  | 
2511  | 
||
| 55814 | 2512  | 
in  | 
| 60325 | 2513  | 
fn (ctxt, t) =>  | 
| 55814 | 2514  | 
let  | 
2515  | 
val fs = Misc_Legacy.term_frees t;  | 
|
2516  | 
val bs = term_bools [] t;  | 
|
2517  | 
val vs = map_index swap fs;  | 
|
2518  | 
val ps = map_index swap bs;  | 
|
| 60325 | 2519  | 
      val t' = term_of_fm ps vs (@{code pa} (fm_of_term ps vs t));
 | 
2520  | 
in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end  | 
|
| 
69266
 
7cc2d66a92a6
proper ML expressions, without trailing semicolons;
 
wenzelm 
parents: 
69214 
diff
changeset
 | 
2521  | 
end  | 
| 60533 | 2522  | 
\<close>  | 
| 27456 | 2523  | 
|
| 69605 | 2524  | 
ML_file \<open>cooper_tac.ML\<close>  | 
| 47432 | 2525  | 
|
| 60533 | 2526  | 
method_setup cooper = \<open>  | 
| 53168 | 2527  | 
Scan.lift (Args.mode "no_quantify") >>  | 
| 47432 | 2528  | 
(fn q => fn ctxt => SIMPLE_METHOD' (Cooper_Tac.linz_tac ctxt (not q)))  | 
| 60533 | 2529  | 
\<close> "decision procedure for linear integer arithmetic"  | 
| 47432 | 2530  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2531  | 
|
| 70091 | 2532  | 
subsection \<open>Tests\<close>  | 
| 27456 | 2533  | 
|
| 55814 | 2534  | 
lemma "\<exists>(j::int). \<forall>x\<ge>j. \<exists>a b. x = 3*a+5*b"  | 
| 27456 | 2535  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2536  | 
|
| 55814 | 2537  | 
lemma "\<forall>(x::int) \<ge> 8. \<exists>i j. 5*i + 3*j = x"  | 
| 27456 | 2538  | 
by cooper  | 
2539  | 
||
| 55814 | 2540  | 
theorem "(\<forall>(y::int). 3 dvd y) \<Longrightarrow> \<forall>(x::int). b < x \<longrightarrow> a \<le> x"  | 
| 23274 | 2541  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2542  | 
|
| 55814 | 2543  | 
theorem "\<And>(y::int) (z::int) (n::int). 3 dvd z \<Longrightarrow> 2 dvd (y::int) \<Longrightarrow>  | 
2544  | 
(\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)"  | 
|
| 23274 | 2545  | 
by cooper  | 
2546  | 
||
| 55814 | 2547  | 
theorem "\<And>(y::int) (z::int) n. Suc n < 6 \<Longrightarrow> 3 dvd z \<Longrightarrow>  | 
2548  | 
2 dvd (y::int) \<Longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)"  | 
|
| 23274 | 2549  | 
by cooper  | 
2550  | 
||
| 55814 | 2551  | 
theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 \<longrightarrow> y = 5 + x"  | 
| 23274 | 2552  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2553  | 
|
| 55814 | 2554  | 
lemma "\<forall>(x::int) \<ge> 8. \<exists>i j. 5*i + 3*j = x"  | 
| 50313 | 2555  | 
by cooper  | 
| 27456 | 2556  | 
|
| 55814 | 2557  | 
lemma "\<forall>(y::int) (z::int) (n::int).  | 
2558  | 
3 dvd z \<longrightarrow> 2 dvd (y::int) \<longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)"  | 
|
| 27456 | 2559  | 
by cooper  | 
2560  | 
||
| 55814 | 2561  | 
lemma "\<forall>(x::int) y. x < y \<longrightarrow> 2 * x + 1 < 2 * y"  | 
| 27456 | 2562  | 
by cooper  | 
2563  | 
||
| 55814 | 2564  | 
lemma "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y"  | 
| 27456 | 2565  | 
by cooper  | 
2566  | 
||
| 55814 | 2567  | 
lemma "\<exists>(x::int) y. 0 < x \<and> 0 \<le> y \<and> 3 * x - 5 * y = 1"  | 
| 27456 | 2568  | 
by cooper  | 
2569  | 
||
| 55814 | 2570  | 
lemma "\<not> (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)"  | 
| 27456 | 2571  | 
by cooper  | 
2572  | 
||
| 55814 | 2573  | 
lemma "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)"  | 
| 27456 | 2574  | 
by cooper  | 
2575  | 
||
| 55814 | 2576  | 
lemma "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<exists>(y::int). x = 2*y)"  | 
| 27456 | 2577  | 
by cooper  | 
2578  | 
||
| 55814 | 2579  | 
lemma "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y + 1)"  | 
| 27456 | 2580  | 
by cooper  | 
2581  | 
||
| 55814 | 2582  | 
lemma "\<not> (\<forall>(x::int).  | 
| 55921 | 2583  | 
(2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y+1) \<or>  | 
| 55814 | 2584  | 
(\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17) \<longrightarrow> 0 < x \<or> (\<not> 3 dvd x \<and> x + 8 = 0)))"  | 
| 23274 | 2585  | 
by cooper  | 
| 27456 | 2586  | 
|
| 55814 | 2587  | 
lemma "\<not> (\<forall>(i::int). 4 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))"  | 
| 27456 | 2588  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2589  | 
|
| 55814 | 2590  | 
lemma "\<exists>j. \<forall>(x::int) \<ge> j. \<exists>i j. 5*i + 3*j = x"  | 
2591  | 
by cooper  | 
|
2592  | 
||
2593  | 
theorem "(\<forall>(y::int). 3 dvd y) \<Longrightarrow> \<forall>(x::int). b < x \<longrightarrow> a \<le> x"  | 
|
| 23274 | 2594  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2595  | 
|
| 55814 | 2596  | 
theorem "\<And>(y::int) (z::int) (n::int). 3 dvd z \<Longrightarrow> 2 dvd (y::int) \<Longrightarrow>  | 
2597  | 
(\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)"  | 
|
| 23274 | 2598  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2599  | 
|
| 55814 | 2600  | 
theorem "\<And>(y::int) (z::int) n. Suc n < 6 \<Longrightarrow> 3 dvd z \<Longrightarrow>  | 
2601  | 
2 dvd (y::int) \<Longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)"  | 
|
| 23274 | 2602  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2603  | 
|
| 55814 | 2604  | 
theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 \<longrightarrow> y = 5 + x"  | 
| 23274 | 2605  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2606  | 
|
| 55814 | 2607  | 
theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x \<or> x div 6 + 1 = 2"  | 
| 23274 | 2608  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2609  | 
|
| 23274 | 2610  | 
theorem "\<exists>(x::int). 0 < x"  | 
2611  | 
by cooper  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2612  | 
|
| 55814 | 2613  | 
theorem "\<forall>(x::int) y. x < y \<longrightarrow> 2 * x + 1 < 2 * y"  | 
| 23274 | 2614  | 
by cooper  | 
| 50313 | 2615  | 
|
| 23274 | 2616  | 
theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y"  | 
2617  | 
by cooper  | 
|
| 50313 | 2618  | 
|
| 67123 | 2619  | 
theorem "\<exists>(x::int) y. 0 < x \<and> 0 \<le> y \<and> 3 * x - 5 * y = 1"  | 
| 23274 | 2620  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2621  | 
|
| 55814 | 2622  | 
theorem "\<not> (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)"  | 
| 23274 | 2623  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2624  | 
|
| 55814 | 2625  | 
theorem "\<not> (\<exists>(x::int). False)"  | 
| 23274 | 2626  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2627  | 
|
| 55814 | 2628  | 
theorem "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)"  | 
| 50313 | 2629  | 
by cooper  | 
| 23274 | 2630  | 
|
| 55814 | 2631  | 
theorem "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)"  | 
| 50313 | 2632  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2633  | 
|
| 55814 | 2634  | 
theorem "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<exists>(y::int). x = 2*y)"  | 
| 50313 | 2635  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2636  | 
|
| 55814 | 2637  | 
theorem "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y + 1)"  | 
| 50313 | 2638  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2639  | 
|
| 55814 | 2640  | 
theorem  | 
2641  | 
"\<not> (\<forall>(x::int).  | 
|
2642  | 
(2 dvd x \<longleftrightarrow>  | 
|
2643  | 
(\<forall>(y::int). x \<noteq> 2*y+1) \<or>  | 
|
2644  | 
(\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17)  | 
|
2645  | 
\<longrightarrow> 0 < x \<or> (\<not> 3 dvd x \<and> x + 8 = 0)))"  | 
|
| 23274 | 2646  | 
by cooper  | 
| 50313 | 2647  | 
|
| 55814 | 2648  | 
theorem "\<not> (\<forall>(i::int). 4 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))"  | 
| 23274 | 2649  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2650  | 
|
| 55814 | 2651  | 
theorem "\<forall>(i::int). 8 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i)"  | 
| 23274 | 2652  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2653  | 
|
| 55814 | 2654  | 
theorem "\<exists>(j::int). \<forall>i. j \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i)"  | 
| 23274 | 2655  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2656  | 
|
| 55814 | 2657  | 
theorem "\<not> (\<forall>j (i::int). j \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))"  | 
| 23274 | 2658  | 
by cooper  | 
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2659  | 
|
| 55814 | 2660  | 
theorem "(\<exists>m::nat. n = 2 * m) \<longrightarrow> (n + 1) div 2 = n div 2"  | 
| 23274 | 2661  | 
by cooper  | 
| 17388 | 2662  | 
|
| 70091 | 2663  | 
|
2664  | 
subsection \<open>Variant for HOL-Main\<close>  | 
|
2665  | 
||
| 
70092
 
a19dd7006a3c
more explicit way to re-generate ~~/src/HOL/Tools/Qelim/cooper_procedure.ML
 
haftmann 
parents: 
70091 
diff
changeset
 | 
2666  | 
export_code pa T Bound nat_of_integer integer_of_nat int_of_integer integer_of_int  | 
| 
 
a19dd7006a3c
more explicit way to re-generate ~~/src/HOL/Tools/Qelim/cooper_procedure.ML
 
haftmann 
parents: 
70091 
diff
changeset
 | 
2667  | 
in Eval module_name Cooper_Procedure file_prefix cooper_procedure  | 
| 70091 | 2668  | 
|
| 
17378
 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 
chaieb 
parents:  
diff
changeset
 | 
2669  | 
end  |