| author | wenzelm | 
| Mon, 23 Mar 2009 22:38:02 +0100 | |
| changeset 30677 | df6ca2f50199 | 
| parent 30595 | c87a3350f5a9 | 
| child 30686 | 47a32dd1b86e | 
| permissions | -rw-r--r-- | 
| 24584 | 1  | 
(* Title: HOL/Tools/Qelim/cooper.ML  | 
| 23466 | 2  | 
Author: Amine Chaieb, TU Muenchen  | 
3  | 
*)  | 
|
4  | 
||
5  | 
signature COOPER =  | 
|
6  | 
sig  | 
|
| 23484 | 7  | 
val cooper_conv : Proof.context -> conv  | 
| 23466 | 8  | 
exception COOPER of string * exn  | 
9  | 
end;  | 
|
10  | 
||
11  | 
structure Cooper: COOPER =  | 
|
12  | 
struct  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
13  | 
|
| 23466 | 14  | 
open Conv;  | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
15  | 
open Normalizer;  | 
| 
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
16  | 
|
| 23466 | 17  | 
exception COOPER of string * exn;  | 
| 27018 | 18  | 
fun simp_thms_conv ctxt =  | 
19  | 
Simplifier.rewrite (Simplifier.context ctxt HOL_basic_ss addsimps simp_thms);  | 
|
| 23484 | 20  | 
val FWD = Drule.implies_elim_list;  | 
| 23466 | 21  | 
|
22  | 
val true_tm = @{cterm "True"};
 | 
|
23  | 
val false_tm = @{cterm "False"};
 | 
|
24  | 
val zdvd1_eq = @{thm "zdvd1_eq"};
 | 
|
25  | 
val presburger_ss = @{simpset} addsimps [zdvd1_eq];
 | 
|
| 
30595
 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 
wenzelm 
parents: 
30448 
diff
changeset
 | 
26  | 
val lin_ss = presburger_ss addsimps (@{thm dvd_eq_mod_eq_0} :: zdvd1_eq :: @{thms zadd_ac});
 | 
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
27  | 
|
| 23466 | 28  | 
val iT = HOLogic.intT  | 
29  | 
val bT = HOLogic.boolT;  | 
|
30  | 
val dest_numeral = HOLogic.dest_number #> snd;  | 
|
31  | 
||
32  | 
val [miconj, midisj, mieq, mineq, milt, mile, migt, mige, midvd, mindvd, miP] =  | 
|
33  | 
    map(instantiate' [SOME @{ctyp "int"}] []) @{thms "minf"};
 | 
|
34  | 
||
35  | 
val [infDconj, infDdisj, infDdvd,infDndvd,infDP] =  | 
|
36  | 
    map(instantiate' [SOME @{ctyp "int"}] []) @{thms "inf_period"};
 | 
|
37  | 
||
38  | 
val [piconj, pidisj, pieq,pineq,pilt,pile,pigt,pige,pidvd,pindvd,piP] =  | 
|
39  | 
    map (instantiate' [SOME @{ctyp "int"}] []) @{thms "pinf"};
 | 
|
40  | 
||
41  | 
val [miP, piP] = map (instantiate' [SOME @{ctyp "bool"}] []) [miP, piP];
 | 
|
42  | 
||
43  | 
val infDP = instantiate' (map SOME [@{ctyp "int"}, @{ctyp "bool"}]) [] infDP;
 | 
|
44  | 
||
45  | 
val [[asetconj, asetdisj, aseteq, asetneq, asetlt, asetle,  | 
|
46  | 
asetgt, asetge, asetdvd, asetndvd,asetP],  | 
|
47  | 
[bsetconj, bsetdisj, bseteq, bsetneq, bsetlt, bsetle,  | 
|
48  | 
      bsetgt, bsetge, bsetdvd, bsetndvd,bsetP]]  = [@{thms "aset"}, @{thms "bset"}];
 | 
|
49  | 
||
50  | 
val [miex, cpmi, piex, cppi] = [@{thm "minusinfinity"}, @{thm "cpmi"}, 
 | 
|
51  | 
                                @{thm "plusinfinity"}, @{thm "cppi"}];
 | 
|
52  | 
||
53  | 
val unity_coeff_ex = instantiate' [SOME @{ctyp "int"}] [] @{thm "unity_coeff_ex"};
 | 
|
54  | 
||
55  | 
val [zdvd_mono,simp_from_to,all_not_ex] =  | 
|
56  | 
     [@{thm "zdvd_mono"}, @{thm "simp_from_to"}, @{thm "all_not_ex"}];
 | 
|
57  | 
||
58  | 
val [dvd_uminus, dvd_uminus'] = @{thms "uminus_dvd_conv"};
 | 
|
59  | 
||
60  | 
val eval_ss = presburger_ss addsimps [simp_from_to] delsimps [insert_iff,bex_triv];  | 
|
61  | 
val eval_conv = Simplifier.rewrite eval_ss;  | 
|
62  | 
||
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
63  | 
(* recognising cterm without moving to terms *)  | 
| 23466 | 64  | 
|
65  | 
datatype fm = And of cterm*cterm| Or of cterm*cterm| Eq of cterm | NEq of cterm  | 
|
66  | 
| Lt of cterm | Le of cterm | Gt of cterm | Ge of cterm  | 
|
67  | 
| Dvd of cterm*cterm | NDvd of cterm*cterm | Nox  | 
|
68  | 
||
69  | 
fun whatis x ct =  | 
|
70  | 
( case (term_of ct) of  | 
|
71  | 
  Const("op &",_)$_$_ => And (Thm.dest_binop ct)
 | 
|
72  | 
| Const ("op |",_)$_$_ => Or (Thm.dest_binop ct)
 | 
|
73  | 
| Const ("op =",ty)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
 | 
|
| 25768 | 74  | 
| Const (@{const_name Not},_) $ (Const ("op =",_)$y$_) => 
 | 
| 23466 | 75  | 
if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox  | 
| 23881 | 76  | 
| Const (@{const_name HOL.less}, _) $ y$ z =>
 | 
| 23466 | 77  | 
if term_of x aconv y then Lt (Thm.dest_arg ct)  | 
78  | 
else if term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox  | 
|
| 23881 | 79  | 
| Const (@{const_name HOL.less_eq}, _) $ y $ z => 
 | 
| 23466 | 80  | 
if term_of x aconv y then Le (Thm.dest_arg ct)  | 
81  | 
else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox  | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
82  | 
| Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$y$_) =>
 | 
| 23466 | 83  | 
if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox  | 
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
84  | 
| Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name "HOL.plus"},_)$y$_)) =>
 | 
| 23466 | 85  | 
if term_of x aconv y then  | 
86  | 
NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox  | 
|
87  | 
| _ => Nox)  | 
|
88  | 
handle CTERM _ => Nox;  | 
|
89  | 
||
90  | 
fun get_pmi_term t =  | 
|
91  | 
let val (x,eq) =  | 
|
92  | 
(Thm.dest_abs NONE o Thm.dest_arg o snd o Thm.dest_abs NONE o Thm.dest_arg)  | 
|
93  | 
(Thm.dest_arg t)  | 
|
94  | 
in (Thm.cabs x o Thm.dest_arg o Thm.dest_arg) eq end;  | 
|
95  | 
||
96  | 
val get_pmi = get_pmi_term o cprop_of;  | 
|
97  | 
||
98  | 
val p_v' = @{cpat "?P' :: int => bool"}; 
 | 
|
99  | 
val q_v' = @{cpat "?Q' :: int => bool"};
 | 
|
100  | 
val p_v = @{cpat "?P:: int => bool"};
 | 
|
101  | 
val q_v = @{cpat "?Q:: int => bool"};
 | 
|
102  | 
||
103  | 
fun myfwd (th1, th2, th3) p q  | 
|
104  | 
[(th_1,th_2,th_3), (th_1',th_2',th_3')] =  | 
|
105  | 
let  | 
|
106  | 
val (mp', mq') = (get_pmi th_1, get_pmi th_1')  | 
|
107  | 
val mi_th = FWD (instantiate ([],[(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1)  | 
|
108  | 
[th_1, th_1']  | 
|
109  | 
val infD_th = FWD (instantiate ([],[(p_v,mp'), (q_v, mq')]) th3) [th_3,th_3']  | 
|
110  | 
val set_th = FWD (instantiate ([],[(p_v,p), (q_v,q)]) th2) [th_2, th_2']  | 
|
111  | 
in (mi_th, set_th, infD_th)  | 
|
112  | 
end;  | 
|
113  | 
||
114  | 
val inst' = fn cts => instantiate' [] (map SOME cts);  | 
|
115  | 
val infDTrue = instantiate' [] [SOME true_tm] infDP;  | 
|
116  | 
val infDFalse = instantiate' [] [SOME false_tm] infDP;  | 
|
117  | 
||
118  | 
val cadd =  @{cterm "op + :: int => _"}
 | 
|
119  | 
val cmulC =  @{cterm "op * :: int => _"}
 | 
|
120  | 
val cminus =  @{cterm "op - :: int => _"}
 | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
121  | 
val cone =  @{cterm "1 :: int"}
 | 
| 23466 | 122  | 
val cneg = @{cterm "uminus :: int => _"}
 | 
123  | 
val [addC, mulC, subC, negC] = map term_of [cadd, cmulC, cminus, cneg]  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
124  | 
val [zero, one] = [@{term "0 :: int"}, @{term "1 :: int"}];
 | 
| 23466 | 125  | 
|
126  | 
val is_numeral = can dest_numeral;  | 
|
127  | 
||
128  | 
fun numeral1 f n = HOLogic.mk_number iT (f (dest_numeral n));  | 
|
129  | 
fun numeral2 f m n = HOLogic.mk_number iT (f (dest_numeral m) (dest_numeral n));  | 
|
130  | 
||
131  | 
val [minus1,plus1] =  | 
|
132  | 
map (fn c => fn t => Thm.capply (Thm.capply c t) cone) [cminus,cadd];  | 
|
133  | 
||
134  | 
fun decomp_pinf x dvd inS [aseteq, asetneq, asetlt, asetle,  | 
|
135  | 
asetgt, asetge,asetdvd,asetndvd,asetP,  | 
|
136  | 
infDdvd, infDndvd, asetconj,  | 
|
137  | 
asetdisj, infDconj, infDdisj] cp =  | 
|
138  | 
case (whatis x cp) of  | 
|
139  | 
And (p,q) => ([p,q], myfwd (piconj, asetconj, infDconj) (Thm.cabs x p) (Thm.cabs x q))  | 
|
140  | 
| Or (p,q) => ([p,q], myfwd (pidisj, asetdisj, infDdisj) (Thm.cabs x p) (Thm.cabs x q))  | 
|
141  | 
| Eq t => ([], K (inst' [t] pieq, FWD (inst' [t] aseteq) [inS (plus1 t)], infDFalse))  | 
|
142  | 
| NEq t => ([], K (inst' [t] pineq, FWD (inst' [t] asetneq) [inS t], infDTrue))  | 
|
143  | 
| Lt t => ([], K (inst' [t] pilt, FWD (inst' [t] asetlt) [inS t], infDFalse))  | 
|
144  | 
| Le t => ([], K (inst' [t] pile, FWD (inst' [t] asetle) [inS (plus1 t)], infDFalse))  | 
|
145  | 
| Gt t => ([], K (inst' [t] pigt, (inst' [t] asetgt), infDTrue))  | 
|
146  | 
| Ge t => ([], K (inst' [t] pige, (inst' [t] asetge), infDTrue))  | 
|
147  | 
| Dvd (d,s) =>  | 
|
148  | 
([],let val dd = dvd d  | 
|
149  | 
in K (inst' [d,s] pidvd, FWD (inst' [d,s] asetdvd) [dd],FWD (inst' [d,s] infDdvd) [dd]) end)  | 
|
150  | 
| NDvd(d,s) => ([],let val dd = dvd d  | 
|
151  | 
in K (inst' [d,s] pindvd, FWD (inst' [d,s] asetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)  | 
|
152  | 
| _ => ([], K (inst' [cp] piP, inst' [cp] asetP, inst' [cp] infDP));  | 
|
153  | 
||
154  | 
fun decomp_minf x dvd inS [bseteq,bsetneq,bsetlt, bsetle, bsetgt,  | 
|
155  | 
bsetge,bsetdvd,bsetndvd,bsetP,  | 
|
156  | 
infDdvd, infDndvd, bsetconj,  | 
|
157  | 
bsetdisj, infDconj, infDdisj] cp =  | 
|
158  | 
case (whatis x cp) of  | 
|
159  | 
And (p,q) => ([p,q], myfwd (miconj, bsetconj, infDconj) (Thm.cabs x p) (Thm.cabs x q))  | 
|
160  | 
| Or (p,q) => ([p,q], myfwd (midisj, bsetdisj, infDdisj) (Thm.cabs x p) (Thm.cabs x q))  | 
|
161  | 
| Eq t => ([], K (inst' [t] mieq, FWD (inst' [t] bseteq) [inS (minus1 t)], infDFalse))  | 
|
162  | 
| NEq t => ([], K (inst' [t] mineq, FWD (inst' [t] bsetneq) [inS t], infDTrue))  | 
|
163  | 
| Lt t => ([], K (inst' [t] milt, (inst' [t] bsetlt), infDTrue))  | 
|
164  | 
| Le t => ([], K (inst' [t] mile, (inst' [t] bsetle), infDTrue))  | 
|
165  | 
| Gt t => ([], K (inst' [t] migt, FWD (inst' [t] bsetgt) [inS t], infDFalse))  | 
|
166  | 
| Ge t => ([], K (inst' [t] mige,FWD (inst' [t] bsetge) [inS (minus1 t)], infDFalse))  | 
|
167  | 
| Dvd (d,s) => ([],let val dd = dvd d  | 
|
168  | 
in K (inst' [d,s] midvd, FWD (inst' [d,s] bsetdvd) [dd] , FWD (inst' [d,s] infDdvd) [dd]) end)  | 
|
169  | 
| NDvd (d,s) => ([],let val dd = dvd d  | 
|
170  | 
in K (inst' [d,s] mindvd, FWD (inst' [d,s] bsetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)  | 
|
171  | 
| _ => ([], K (inst' [cp] miP, inst' [cp] bsetP, inst' [cp] infDP))  | 
|
172  | 
||
173  | 
(* Canonical linear form for terms, formulae etc.. *)  | 
|
174  | 
fun provelin ctxt t = Goal.prove ctxt [] [] t  | 
|
| 24075 | 175  | 
(fn _ => EVERY [simp_tac lin_ss 1, TRY (simple_arith_tac ctxt 1)]);  | 
| 25768 | 176  | 
fun linear_cmul 0 tm = zero  | 
177  | 
| linear_cmul n tm = case tm of  | 
|
178  | 
      Const (@{const_name HOL.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
 | 
|
179  | 
    | Const (@{const_name HOL.times}, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x
 | 
|
180  | 
    | Const (@{const_name HOL.minus}, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b
 | 
|
181  | 
    | (m as Const (@{const_name HOL.uminus}, _)) $ a => m $ linear_cmul n a
 | 
|
182  | 
| _ => numeral1 (fn m => n * m) tm;  | 
|
| 23466 | 183  | 
fun earlier [] x y = false  | 
184  | 
| earlier (h::t) x y =  | 
|
185  | 
if h aconv y then false else if h aconv x then true else earlier t x y;  | 
|
186  | 
||
| 25768 | 187  | 
fun linear_add vars tm1 tm2 = case (tm1, tm2) of  | 
188  | 
    (Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c1 $ x1) $ r1,
 | 
|
189  | 
    Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c2 $ x2) $ r2) =>
 | 
|
| 23466 | 190  | 
if x1 = x2 then  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
191  | 
let val c = numeral2 (curry op +) c1 c2  | 
| 25768 | 192  | 
in if c = zero then linear_add vars r1 r2  | 
193  | 
else addC$(mulC$c$x1)$(linear_add vars r1 r2)  | 
|
| 23466 | 194  | 
end  | 
| 25768 | 195  | 
else if earlier vars x1 x2 then addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2  | 
196  | 
else addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2  | 
|
197  | 
 | (Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c1 $ x1) $ r1, _) =>
 | 
|
198  | 
addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2  | 
|
199  | 
 | (_, Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c2 $ x2) $ r2) => 
 | 
|
200  | 
addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2  | 
|
201  | 
| (_, _) => numeral2 (curry op +) tm1 tm2;  | 
|
| 23466 | 202  | 
|
203  | 
fun linear_neg tm = linear_cmul ~1 tm;  | 
|
204  | 
fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2);  | 
|
205  | 
||
206  | 
||
| 25768 | 207  | 
fun lint vars tm = if is_numeral tm then tm else case tm of  | 
208  | 
  Const (@{const_name HOL.uminus}, _) $ t => linear_neg (lint vars t)
 | 
|
209  | 
| Const (@{const_name HOL.plus}, _) $ s $ t => linear_add vars (lint vars s) (lint vars t)
 | 
|
210  | 
| Const (@{const_name HOL.minus}, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t)
 | 
|
211  | 
| Const (@{const_name HOL.times}, _) $ s $ t =>
 | 
|
| 23466 | 212  | 
let val s' = lint vars s  | 
213  | 
val t' = lint vars t  | 
|
214  | 
in if is_numeral s' then (linear_cmul (dest_numeral s') t')  | 
|
215  | 
else if is_numeral t' then (linear_cmul (dest_numeral t') s')  | 
|
216  | 
     else raise COOPER ("Cooper Failed", TERM ("lint: not linear",[tm]))
 | 
|
217  | 
end  | 
|
| 25768 | 218  | 
| _ => addC $ (mulC $ one $ tm) $ zero;  | 
| 23466 | 219  | 
|
| 25768 | 220  | 
fun lin (vs as x::_) (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.less}, T) $ s $ t)) = 
 | 
221  | 
    lin vs (Const (@{const_name HOL.less_eq}, T) $ t $ s)
 | 
|
222  | 
  | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name HOL.less_eq}, T) $ s $ t)) = 
 | 
|
223  | 
    lin vs (Const (@{const_name HOL.less}, T) $ t $ s)
 | 
|
224  | 
  | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
 | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
225  | 
  | lin (vs as x::_) (Const(@{const_name Ring_and_Field.dvd},_)$d$t) = 
 | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
226  | 
    HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} (numeral1 abs d, lint vs t)
 | 
| 23466 | 227  | 
  | lin (vs as x::_) ((b as Const("op =",_))$s$t) = 
 | 
228  | 
(case lint vs (subC$t$s) of  | 
|
229  | 
(t as a$(m$c$y)$r) =>  | 
|
230  | 
if x <> y then b$zero$t  | 
|
231  | 
else if dest_numeral c < 0 then b$(m$(numeral1 ~ c)$y)$r  | 
|
232  | 
else b$(m$c$y)$(linear_neg r)  | 
|
233  | 
| t => b$zero$t)  | 
|
234  | 
| lin (vs as x::_) (b$s$t) =  | 
|
235  | 
(case lint vs (subC$t$s) of  | 
|
236  | 
(t as a$(m$c$y)$r) =>  | 
|
237  | 
if x <> y then b$zero$t  | 
|
238  | 
else if dest_numeral c < 0 then b$(m$(numeral1 ~ c)$y)$r  | 
|
239  | 
else b$(linear_neg r)$(m$c$y)  | 
|
240  | 
| t => b$zero$t)  | 
|
241  | 
| lin vs fm = fm;  | 
|
242  | 
||
243  | 
fun lint_conv ctxt vs ct =  | 
|
244  | 
let val t = term_of ct  | 
|
245  | 
in (provelin ctxt ((HOLogic.eq_const iT)$t$(lint vs t) |> HOLogic.mk_Trueprop))  | 
|
246  | 
RS eq_reflection  | 
|
247  | 
end;  | 
|
248  | 
||
249  | 
fun is_intrel (b$_$_) = domain_type (fastype_of b) = HOLogic.intT  | 
|
250  | 
  | is_intrel (@{term "Not"}$(b$_$_)) = domain_type (fastype_of b) = HOLogic.intT
 | 
|
251  | 
| is_intrel _ = false;  | 
|
252  | 
||
| 25768 | 253  | 
fun linearize_conv ctxt vs ct = case term_of ct of  | 
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
254  | 
  Const(@{const_name Ring_and_Field.dvd},_)$d$t => 
 | 
| 23466 | 255  | 
let  | 
256  | 
val th = binop_conv (lint_conv ctxt vs) ct  | 
|
257  | 
val (d',t') = Thm.dest_binop (Thm.rhs_of th)  | 
|
258  | 
val (dt',tt') = (term_of d', term_of t')  | 
|
259  | 
in if is_numeral dt' andalso is_numeral tt'  | 
|
260  | 
then Conv.fconv_rule (arg_conv (Simplifier.rewrite presburger_ss)) th  | 
|
261  | 
else  | 
|
262  | 
let  | 
|
263  | 
val dth =  | 
|
264  | 
((if dest_numeral (term_of d') < 0 then  | 
|
265  | 
Conv.fconv_rule (arg_conv (arg1_conv (lint_conv ctxt vs)))  | 
|
266  | 
(Thm.transitive th (inst' [d',t'] dvd_uminus))  | 
|
267  | 
else th) handle TERM _ => th)  | 
|
268  | 
val d'' = Thm.rhs_of dth |> Thm.dest_arg1  | 
|
269  | 
in  | 
|
270  | 
case tt' of  | 
|
| 25768 | 271  | 
        Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$_)$_ => 
 | 
| 23466 | 272  | 
let val x = dest_numeral c  | 
273  | 
in if x < 0 then Conv.fconv_rule (arg_conv (arg_conv (lint_conv ctxt vs)))  | 
|
274  | 
(Thm.transitive dth (inst' [d'',t'] dvd_uminus'))  | 
|
275  | 
else dth end  | 
|
276  | 
| _ => dth  | 
|
277  | 
end  | 
|
278  | 
end  | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
279  | 
| Const (@{const_name Not},_)$(Const(@{const_name Ring_and_Field.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct
 | 
| 23466 | 280  | 
| t => if is_intrel t  | 
281  | 
then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop))  | 
|
282  | 
RS eq_reflection  | 
|
283  | 
else reflexive ct;  | 
|
284  | 
||
285  | 
val dvdc = @{cterm "op dvd :: int => _"};
 | 
|
286  | 
||
287  | 
fun unify ctxt q =  | 
|
288  | 
let  | 
|
289  | 
val (e,(cx,p)) = q |> Thm.dest_comb ||> Thm.dest_abs NONE  | 
|
290  | 
val x = term_of cx  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
291  | 
val ins = insert (op = : int * int -> bool)  | 
| 23466 | 292  | 
fun h (acc,dacc) t =  | 
293  | 
case (term_of t) of  | 
|
| 25768 | 294  | 
    Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ => 
 | 
| 23881 | 295  | 
if x aconv y andalso member (op =)  | 
296  | 
      ["op =", @{const_name HOL.less}, @{const_name HOL.less_eq}] s
 | 
|
| 23466 | 297  | 
then (ins (dest_numeral c) acc,dacc) else (acc,dacc)  | 
| 25768 | 298  | 
  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) => 
 | 
| 23881 | 299  | 
if x aconv y andalso member (op =)  | 
300  | 
       [@{const_name HOL.less}, @{const_name HOL.less_eq}] s 
 | 
|
| 23466 | 301  | 
then (ins (dest_numeral c) acc, dacc) else (acc,dacc)  | 
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
302  | 
  | Const(@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_) => 
 | 
| 23466 | 303  | 
if x aconv y then (acc,ins (dest_numeral c) dacc) else (acc,dacc)  | 
304  | 
  | Const("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
 | 
|
305  | 
  | Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
 | 
|
| 25768 | 306  | 
  | Const (@{const_name Not},_)$_ => h (acc,dacc) (Thm.dest_arg t)
 | 
| 23466 | 307  | 
| _ => (acc, dacc)  | 
308  | 
val (cs,ds) = h ([],[]) p  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
309  | 
val l = Integer.lcms (cs union ds)  | 
| 23466 | 310  | 
fun cv k ct =  | 
311  | 
let val (tm as b$s$t) = term_of ct  | 
|
312  | 
in ((HOLogic.eq_const bT)$tm$(b$(linear_cmul k s)$(linear_cmul k t))  | 
|
313  | 
|> HOLogic.mk_Trueprop |> provelin ctxt) RS eq_reflection end  | 
|
314  | 
fun nzprop x =  | 
|
315  | 
let  | 
|
316  | 
val th =  | 
|
317  | 
Simplifier.rewrite lin_ss  | 
|
318  | 
      (Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} 
 | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
319  | 
           (Thm.capply (Thm.capply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x)) 
 | 
| 
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
320  | 
           @{cterm "0::int"})))
 | 
| 23466 | 321  | 
in equal_elim (Thm.symmetric th) TrueI end;  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
322  | 
val notz = let val tab = fold Inttab.update  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
323  | 
(ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty  | 
| 23466 | 324  | 
in  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
325  | 
(fn ct => (valOf (Inttab.lookup tab (ct |> term_of |> dest_numeral))  | 
| 23466 | 326  | 
handle Option => (writeln "noz: Theorems-Table contains no entry for";  | 
| 26928 | 327  | 
Display.print_cterm ct ; raise Option)))  | 
| 23466 | 328  | 
end  | 
329  | 
fun unit_conv t =  | 
|
330  | 
case (term_of t) of  | 
|
331  | 
   Const("op &",_)$_$_ => binop_conv unit_conv t
 | 
|
332  | 
  | Const("op |",_)$_$_ => binop_conv unit_conv t
 | 
|
| 25768 | 333  | 
  | Const (@{const_name Not},_)$_ => arg_conv unit_conv t
 | 
334  | 
  | Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ => 
 | 
|
| 23881 | 335  | 
if x=y andalso member (op =)  | 
336  | 
      ["op =", @{const_name HOL.less}, @{const_name HOL.less_eq}] s
 | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
337  | 
then cv (l div dest_numeral c) t else Thm.reflexive t  | 
| 25768 | 338  | 
  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) => 
 | 
| 23881 | 339  | 
if x=y andalso member (op =)  | 
340  | 
      [@{const_name HOL.less}, @{const_name HOL.less_eq}] s
 | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
341  | 
then cv (l div dest_numeral c) t else Thm.reflexive t  | 
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
342  | 
  | Const(@{const_name Ring_and_Field.dvd},_)$d$(r as (Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_)) => 
 | 
| 23466 | 343  | 
if x=y then  | 
344  | 
let  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
345  | 
val k = l div dest_numeral c  | 
| 23466 | 346  | 
val kt = HOLogic.mk_number iT k  | 
347  | 
val th1 = inst' [Thm.dest_arg1 t, Thm.dest_arg t]  | 
|
348  | 
((Thm.dest_arg t |> funpow 2 Thm.dest_arg1 |> notz) RS zdvd_mono)  | 
|
349  | 
val (d',t') = (mulC$kt$d, mulC$kt$r)  | 
|
350  | 
val thc = (provelin ctxt ((HOLogic.eq_const iT)$d'$(lint [] d') |> HOLogic.mk_Trueprop))  | 
|
351  | 
RS eq_reflection  | 
|
352  | 
val tht = (provelin ctxt ((HOLogic.eq_const iT)$t'$(linear_cmul k r) |> HOLogic.mk_Trueprop))  | 
|
353  | 
RS eq_reflection  | 
|
354  | 
in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule dvdc thc) tht) end  | 
|
355  | 
else Thm.reflexive t  | 
|
356  | 
| _ => Thm.reflexive t  | 
|
357  | 
val uth = unit_conv p  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
358  | 
  val clt =  Numeral.mk_cnumber @{ctyp "int"} l
 | 
| 23466 | 359  | 
val ltx = Thm.capply (Thm.capply cmulC clt) cx  | 
360  | 
val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth)  | 
|
361  | 
val th' = inst' [Thm.cabs ltx (Thm.rhs_of uth), clt] unity_coeff_ex  | 
|
362  | 
val thf = transitive th  | 
|
363  | 
(transitive (symmetric (beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th')  | 
|
364  | 
val (lth,rth) = Thm.dest_comb (cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true  | 
|
365  | 
||> beta_conversion true |>> Thm.symmetric  | 
|
366  | 
in transitive (transitive lth thf) rth end;  | 
|
367  | 
||
368  | 
||
369  | 
val emptyIS = @{cterm "{}::int set"};
 | 
|
370  | 
val insert_tm = @{cterm "insert :: int => _"};
 | 
|
371  | 
val mem_tm = Const("op :",[iT , HOLogic.mk_setT iT] ---> bT);
 | 
|
372  | 
fun mkISet cts = fold_rev (Thm.capply insert_tm #> Thm.capply) cts emptyIS;  | 
|
373  | 
val cTrp = @{cterm "Trueprop"};
 | 
|
374  | 
val eqelem_imp_imp = (thm"eqelem_imp_iff") RS iffD1;  | 
|
375  | 
val [A_tm,B_tm] = map (fn th => cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg  | 
|
376  | 
|> Thm.dest_abs NONE |> snd |> Thm.dest_fun |> Thm.dest_arg)  | 
|
377  | 
[asetP,bsetP];  | 
|
378  | 
||
379  | 
val D_tm = @{cpat "?D::int"};
 | 
|
380  | 
||
381  | 
fun cooperex_conv ctxt vs q =  | 
|
382  | 
let  | 
|
383  | 
||
384  | 
val uth = unify ctxt q  | 
|
385  | 
val (x,p) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of uth))  | 
|
386  | 
val ins = insert (op aconvc)  | 
|
387  | 
fun h t (bacc,aacc,dacc) =  | 
|
388  | 
case (whatis x t) of  | 
|
389  | 
And (p,q) => h q (h p (bacc,aacc,dacc))  | 
|
390  | 
| Or (p,q) => h q (h p (bacc,aacc,dacc))  | 
|
391  | 
| Eq t => (ins (minus1 t) bacc,  | 
|
392  | 
ins (plus1 t) aacc,dacc)  | 
|
393  | 
| NEq t => (ins t bacc,  | 
|
394  | 
ins t aacc, dacc)  | 
|
395  | 
| Lt t => (bacc, ins t aacc, dacc)  | 
|
396  | 
| Le t => (bacc, ins (plus1 t) aacc,dacc)  | 
|
397  | 
| Gt t => (ins t bacc, aacc,dacc)  | 
|
398  | 
| Ge t => (ins (minus1 t) bacc, aacc,dacc)  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
399  | 
| Dvd (d,s) => (bacc,aacc,insert (op =) (term_of d |> dest_numeral) dacc)  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
400  | 
| NDvd (d,s) => (bacc,aacc,insert (op =) (term_of d|> dest_numeral) dacc)  | 
| 23466 | 401  | 
| _ => (bacc, aacc, dacc)  | 
402  | 
val (b0,a0,ds) = h p ([],[],[])  | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
403  | 
val d = Integer.lcms ds  | 
| 23582 | 404  | 
 val cd = Numeral.mk_cnumber @{ctyp "int"} d
 | 
| 23466 | 405  | 
val dt = term_of cd  | 
406  | 
fun divprop x =  | 
|
407  | 
let  | 
|
408  | 
val th =  | 
|
409  | 
Simplifier.rewrite lin_ss  | 
|
410  | 
      (Thm.capply @{cterm Trueprop} 
 | 
|
| 23582 | 411  | 
           (Thm.capply (Thm.capply dvdc (Numeral.mk_cnumber @{ctyp "int"} x)) cd))
 | 
| 23466 | 412  | 
in equal_elim (Thm.symmetric th) TrueI end;  | 
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
413  | 
val dvd = let val tab = fold Inttab.update  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
414  | 
(ds ~~ (map divprop ds)) Inttab.empty in  | 
| 
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
415  | 
(fn ct => (valOf (Inttab.lookup tab (term_of ct |> dest_numeral))  | 
| 23466 | 416  | 
handle Option => (writeln "dvd: Theorems-Table contains no entry for";  | 
| 26928 | 417  | 
Display.print_cterm ct ; raise Option)))  | 
| 23466 | 418  | 
end  | 
419  | 
val dp =  | 
|
420  | 
let val th = Simplifier.rewrite lin_ss  | 
|
421  | 
      (Thm.capply @{cterm Trueprop} 
 | 
|
422  | 
           (Thm.capply (Thm.capply @{cterm "op < :: int => _"} @{cterm "0::int"}) cd))
 | 
|
423  | 
in equal_elim (Thm.symmetric th) TrueI end;  | 
|
424  | 
(* A and B set *)  | 
|
425  | 
local  | 
|
426  | 
     val insI1 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI1"}
 | 
|
427  | 
     val insI2 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI2"}
 | 
|
428  | 
in  | 
|
429  | 
fun provein x S =  | 
|
430  | 
case term_of S of  | 
|
| 30448 | 431  | 
        Const(@{const_name Set.empty}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
 | 
| 
30304
 
d8e4cd2ac2a1
set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
 
haftmann 
parents: 
29787 
diff
changeset
 | 
432  | 
      | Const(@{const_name insert}, _) $ y $ _ => 
 | 
| 23466 | 433  | 
let val (cy,S') = Thm.dest_binop S  | 
434  | 
in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1  | 
|
435  | 
else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)  | 
|
436  | 
(provein x S')  | 
|
437  | 
end  | 
|
438  | 
end  | 
|
439  | 
||
440  | 
val al = map (lint vs o term_of) a0  | 
|
441  | 
val bl = map (lint vs o term_of) b0  | 
|
442  | 
val (sl,s0,f,abths,cpth) =  | 
|
443  | 
if length (distinct (op aconv) bl) <= length (distinct (op aconv) al)  | 
|
444  | 
then  | 
|
445  | 
(bl,b0,decomp_minf,  | 
|
446  | 
fn B => (map (fn th => implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp)  | 
|
447  | 
[bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@  | 
|
448  | 
(map (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]))  | 
|
449  | 
[bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj,  | 
|
450  | 
bsetdisj,infDconj, infDdisj]),  | 
|
451  | 
cpmi)  | 
|
452  | 
else (al,a0,decomp_pinf,fn A =>  | 
|
453  | 
(map (fn th => implies_elim (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]) th) dp)  | 
|
454  | 
[aseteq,asetneq,asetlt, asetle, asetgt,asetge])@  | 
|
455  | 
(map (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]))  | 
|
456  | 
[asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj,  | 
|
457  | 
asetdisj,infDconj, infDdisj]),cppi)  | 
|
458  | 
val cpth =  | 
|
459  | 
let  | 
|
460  | 
val sths = map (fn (tl,t0) =>  | 
|
461  | 
if tl = term_of t0  | 
|
462  | 
                      then instantiate' [SOME @{ctyp "int"}] [SOME t0] refl
 | 
|
463  | 
else provelin ctxt ((HOLogic.eq_const iT)$tl$(term_of t0)  | 
|
464  | 
|> HOLogic.mk_Trueprop))  | 
|
465  | 
(sl ~~ s0)  | 
|
466  | 
val csl = distinct (op aconvc) (map (cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths)  | 
|
467  | 
val S = mkISet csl  | 
|
468  | 
val inStab = fold (fn ct => fn tab => Termtab.update (term_of ct, provein ct S) tab)  | 
|
469  | 
csl Termtab.empty  | 
|
470  | 
   val eqelem_th = instantiate' [SOME @{ctyp "int"}] [NONE,NONE, SOME S] eqelem_imp_imp
 | 
|
471  | 
val inS =  | 
|
472  | 
let  | 
|
473  | 
fun transmem th0 th1 =  | 
|
474  | 
Thm.equal_elim  | 
|
475  | 
(Drule.arg_cong_rule cTrp (Drule.fun_cong_rule (Drule.arg_cong_rule  | 
|
476  | 
((Thm.dest_fun o Thm.dest_fun o Thm.dest_arg o cprop_of) th1) th0) S)) th1  | 
|
477  | 
val tab = fold Termtab.update  | 
|
478  | 
(map (fn eq =>  | 
|
479  | 
let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop  | 
|
480  | 
val th = if term_of s = term_of t  | 
|
481  | 
then valOf(Termtab.lookup inStab (term_of s))  | 
|
482  | 
else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)  | 
|
483  | 
[eq, valOf(Termtab.lookup inStab (term_of s))]  | 
|
484  | 
in (term_of t, th) end)  | 
|
485  | 
sths) Termtab.empty  | 
|
486  | 
in fn ct =>  | 
|
487  | 
(valOf (Termtab.lookup tab (term_of ct))  | 
|
| 26928 | 488  | 
handle Option => (writeln "inS: No theorem for " ; Display.print_cterm ct ; raise Option))  | 
| 23466 | 489  | 
end  | 
490  | 
val (inf, nb, pd) = divide_and_conquer (f x dvd inS (abths S)) p  | 
|
491  | 
in [dp, inf, nb, pd] MRS cpth  | 
|
492  | 
end  | 
|
493  | 
val cpth' = Thm.transitive uth (cpth RS eq_reflection)  | 
|
| 27018 | 494  | 
in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv) (Thm.rhs_of cpth'))  | 
| 23466 | 495  | 
end;  | 
496  | 
||
497  | 
fun literals_conv bops uops env cv =  | 
|
498  | 
let fun h t =  | 
|
499  | 
case (term_of t) of  | 
|
500  | 
b$_$_ => if member (op aconv) bops b then binop_conv h t else cv env t  | 
|
501  | 
| u$_ => if member (op aconv) uops u then arg_conv h t else cv env t  | 
|
502  | 
| _ => cv env t  | 
|
503  | 
in h end;  | 
|
504  | 
||
505  | 
fun integer_nnf_conv ctxt env =  | 
|
506  | 
nnf_conv then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt);  | 
|
507  | 
||
508  | 
local  | 
|
509  | 
val pcv = Simplifier.rewrite  | 
|
510  | 
(HOL_basic_ss addsimps (simp_thms @ (List.take(ex_simps,4))  | 
|
511  | 
@ [not_all,all_not_ex, ex_disj_distrib]))  | 
|
512  | 
val postcv = Simplifier.rewrite presburger_ss  | 
|
513  | 
fun conv ctxt p =  | 
|
| 24298 | 514  | 
let val _ = ()  | 
| 23466 | 515  | 
in  | 
| 23523 | 516  | 
Qelim.gen_qelim_conv pcv postcv pcv (cons o term_of)  | 
| 
29265
 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 
wenzelm 
parents: 
28397 
diff
changeset
 | 
517  | 
(OldTerm.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt)  | 
| 23466 | 518  | 
(cooperex_conv ctxt) p  | 
519  | 
end  | 
|
520  | 
  handle  CTERM s => raise COOPER ("Cooper Failed", CTERM s)
 | 
|
521  | 
        | THM s => raise COOPER ("Cooper Failed", THM s) 
 | 
|
| 23523 | 522  | 
        | TYPE s => raise COOPER ("Cooper Failed", TYPE s) 
 | 
| 23466 | 523  | 
in val cooper_conv = conv  | 
524  | 
end;  | 
|
525  | 
end;  | 
|
526  | 
||
527  | 
||
528  | 
||
529  | 
structure Coopereif =  | 
|
530  | 
struct  | 
|
531  | 
||
| 23713 | 532  | 
open GeneratedCooper;  | 
533  | 
||
534  | 
fun cooper s = raise Cooper.COOPER ("Cooper oracle failed", ERROR s);
 | 
|
535  | 
fun i_of_term vs t = case t  | 
|
536  | 
of Free (xn, xT) => (case AList.lookup (op aconv) vs t  | 
|
537  | 
of NONE => cooper "Variable not found in the list!"  | 
|
538  | 
| SOME n => Bound n)  | 
|
539  | 
  | @{term "0::int"} => C 0
 | 
|
540  | 
  | @{term "1::int"} => C 1
 | 
|
| 
24630
 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 
wenzelm 
parents: 
24584 
diff
changeset
 | 
541  | 
| Term.Bound i => Bound i  | 
| 23713 | 542  | 
  | Const(@{const_name HOL.uminus},_)$t' => Neg (i_of_term vs t')
 | 
543  | 
  | Const(@{const_name HOL.plus},_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2)
 | 
|
544  | 
  | Const(@{const_name HOL.minus},_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2)
 | 
|
545  | 
  | Const(@{const_name HOL.times},_)$t1$t2 => 
 | 
|
546  | 
(Mul (HOLogic.dest_number t1 |> snd, i_of_term vs t2)  | 
|
547  | 
handle TERM _ =>  | 
|
548  | 
(Mul (HOLogic.dest_number t2 |> snd, i_of_term vs t1)  | 
|
549  | 
handle TERM _ => cooper "Reification: Unsupported kind of multiplication"))  | 
|
550  | 
| _ => (C (HOLogic.dest_number t |> snd)  | 
|
551  | 
handle TERM _ => cooper "Reification: unknown term");  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
552  | 
|
| 23713 | 553  | 
fun qf_of_term ps vs t = case t  | 
554  | 
 of Const("True",_) => T
 | 
|
555  | 
  | Const("False",_) => F
 | 
|
| 23881 | 556  | 
  | Const(@{const_name HOL.less},_)$t1$t2 => Lt (Sub (i_of_term vs t1,i_of_term vs t2))
 | 
557  | 
  | Const(@{const_name HOL.less_eq},_)$t1$t2 => Le (Sub(i_of_term vs t1,i_of_term vs t2))
 | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
27330 
diff
changeset
 | 
558  | 
  | Const(@{const_name Ring_and_Field.dvd},_)$t1$t2 => 
 | 
| 
28397
 
389c5e494605
handle _ should be avoided (spurious Interrupt will spoil the game);
 
wenzelm 
parents: 
28290 
diff
changeset
 | 
559  | 
(Dvd(HOLogic.dest_number t1 |> snd, i_of_term vs t2) handle _ => cooper "Reification: unsupported dvd") (* FIXME avoid handle _ *)  | 
| 23713 | 560  | 
  | @{term "op = :: int => _"}$t1$t2 => Eq (Sub (i_of_term vs t1,i_of_term vs t2))
 | 
| 29787 | 561  | 
  | @{term "op = :: bool => _ "}$t1$t2 => Iff(qf_of_term ps vs t1,qf_of_term ps vs t2)
 | 
| 23713 | 562  | 
  | Const("op &",_)$t1$t2 => And(qf_of_term ps vs t1,qf_of_term ps vs t2)
 | 
563  | 
  | Const("op |",_)$t1$t2 => Or(qf_of_term ps vs t1,qf_of_term ps vs t2)
 | 
|
| 29787 | 564  | 
  | Const("op -->",_)$t1$t2 => Imp(qf_of_term ps vs t1,qf_of_term ps vs t2)
 | 
565  | 
  | Const (@{const_name Not},_)$t' => Not(qf_of_term ps vs t')
 | 
|
| 23713 | 566  | 
  | Const("Ex",_)$Abs(xn,xT,p) => 
 | 
567  | 
let val (xn',p') = variant_abs (xn,xT,p)  | 
|
568  | 
val vs' = (Free (xn',xT), 0) :: (map (fn(v,n) => (v,1+ n)) vs)  | 
|
569  | 
in E (qf_of_term ps vs' p')  | 
|
570  | 
end  | 
|
571  | 
  | Const("All",_)$Abs(xn,xT,p) => 
 | 
|
572  | 
let val (xn',p') = variant_abs (xn,xT,p)  | 
|
573  | 
val vs' = (Free (xn',xT), 0) :: (map (fn(v,n) => (v,1+ n)) vs)  | 
|
574  | 
in A (qf_of_term ps vs' p')  | 
|
575  | 
end  | 
|
576  | 
| _ =>(case AList.lookup (op aconv) ps t of  | 
|
577  | 
NONE => cooper "Reification: unknown term!"  | 
|
578  | 
| SOME n => Closed n);  | 
|
| 23466 | 579  | 
|
580  | 
local  | 
|
581  | 
 val ops = [@{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
 | 
|
582  | 
             @{term "op = :: int => _"}, @{term "op < :: int => _"}, 
 | 
|
583  | 
             @{term "op <= :: int => _"}, @{term "Not"}, @{term "All:: (int => _) => _"}, 
 | 
|
584  | 
             @{term "Ex:: (int => _) => _"}, @{term "True"}, @{term "False"}]
 | 
|
585  | 
fun ty t = Bool.not (fastype_of t = HOLogic.boolT)  | 
|
586  | 
in  | 
|
587  | 
fun term_bools acc t =  | 
|
588  | 
case t of  | 
|
589  | 
(l as f $ a) $ b => if ty t orelse f mem ops then term_bools (term_bools acc l)b  | 
|
590  | 
else insert (op aconv) t acc  | 
|
591  | 
| f $ a => if ty t orelse f mem ops then term_bools (term_bools acc f) a  | 
|
592  | 
else insert (op aconv) t acc  | 
|
593  | 
| Abs p => term_bools acc (snd (variant_abs p))  | 
|
594  | 
| _ => if ty t orelse t mem ops then acc else insert (op aconv) t acc  | 
|
595  | 
end;  | 
|
596  | 
||
597  | 
fun myassoc2 l v =  | 
|
598  | 
case l of  | 
|
599  | 
[] => NONE  | 
|
| 
23689
 
0410269099dc
replaced code generator framework for reflected cooper
 
haftmann 
parents: 
23582 
diff
changeset
 | 
600  | 
| (x,v')::xs => if v = v' then SOME x  | 
| 23466 | 601  | 
else myassoc2 xs v;  | 
602  | 
||
| 23713 | 603  | 
fun term_of_i vs t = case t  | 
604  | 
of C i => HOLogic.mk_number HOLogic.intT i  | 
|
605  | 
| Bound n => the (myassoc2 vs n)  | 
|
606  | 
  | Neg t' => @{term "uminus :: int => _"} $ term_of_i vs t'
 | 
|
607  | 
  | Add (t1, t2) => @{term "op + :: int => _"} $ term_of_i vs t1 $ term_of_i vs t2
 | 
|
608  | 
  | Sub (t1, t2) => @{term "op - :: int => _"} $ term_of_i vs t1 $ term_of_i vs t2
 | 
|
609  | 
  | Mul (i, t2) => @{term "op * :: int => _"} $
 | 
|
610  | 
HOLogic.mk_number HOLogic.intT i $ term_of_i vs t2  | 
|
| 29787 | 611  | 
| Cn (n, i, t') => term_of_i vs (Add (Mul (i, Bound n), t'));  | 
| 23466 | 612  | 
|
613  | 
fun term_of_qf ps vs t =  | 
|
614  | 
case t of  | 
|
615  | 
T => HOLogic.true_const  | 
|
616  | 
| F => HOLogic.false_const  | 
|
617  | 
 | Lt t' => @{term "op < :: int => _ "}$ term_of_i vs t'$ @{term "0::int"}
 | 
|
618  | 
 | Le t' => @{term "op <= :: int => _ "}$ term_of_i vs t' $ @{term "0::int"}
 | 
|
619  | 
 | Gt t' => @{term "op < :: int => _ "}$ @{term "0::int"}$ term_of_i vs t'
 | 
|
620  | 
 | Ge t' => @{term "op <= :: int => _ "}$ @{term "0::int"}$ term_of_i vs t'
 | 
|
621  | 
 | Eq t' => @{term "op = :: int => _ "}$ term_of_i vs t'$ @{term "0::int"}
 | 
|
| 29787 | 622  | 
| NEq t' => term_of_qf ps vs (Not (Eq t'))  | 
| 23713 | 623  | 
 | Dvd(i,t') => @{term "op dvd :: int => _ "} $ 
 | 
624  | 
HOLogic.mk_number HOLogic.intT i $ term_of_i vs t'  | 
|
| 29787 | 625  | 
| NDvd(i,t')=> term_of_qf ps vs (Not(Dvd(i,t')))  | 
626  | 
| Not t' => HOLogic.Not$(term_of_qf ps vs t')  | 
|
| 23466 | 627  | 
| And(t1,t2) => HOLogic.conj$(term_of_qf ps vs t1)$(term_of_qf ps vs t2)  | 
628  | 
| Or(t1,t2) => HOLogic.disj$(term_of_qf ps vs t1)$(term_of_qf ps vs t2)  | 
|
| 29787 | 629  | 
| Imp(t1,t2) => HOLogic.imp$(term_of_qf ps vs t1)$(term_of_qf ps vs t2)  | 
630  | 
 | Iff(t1,t2) => @{term "op = :: bool => _"} $ term_of_qf ps vs t1 $ term_of_qf ps vs t2
 | 
|
| 23713 | 631  | 
| Closed n => the (myassoc2 ps n)  | 
| 29787 | 632  | 
| NClosed n => term_of_qf ps vs (Not (Closed n))  | 
633  | 
| _ => cooper "If this is raised, Isabelle/HOL or code generator is inconsistent!";  | 
|
| 23466 | 634  | 
|
| 28290 | 635  | 
fun cooper_oracle ct =  | 
| 23713 | 636  | 
let  | 
| 28290 | 637  | 
val thy = Thm.theory_of_cterm ct;  | 
638  | 
val t = Thm.term_of ct;  | 
|
| 
29265
 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 
wenzelm 
parents: 
28397 
diff
changeset
 | 
639  | 
val (vs, ps) = pairself (map_index swap) (OldTerm.term_frees t, term_bools [] t);  | 
| 23713 | 640  | 
in  | 
| 28290 | 641  | 
Thm.cterm_of thy (Logic.mk_equals (HOLogic.mk_Trueprop t,  | 
642  | 
HOLogic.mk_Trueprop (term_of_qf ps vs (pa (qf_of_term ps vs t)))))  | 
|
| 23713 | 643  | 
end;  | 
| 23466 | 644  | 
|
645  | 
end;  |