| author | wenzelm | 
| Fri, 01 Feb 2013 21:58:13 +0100 | |
| changeset 51075 | b0901f4a857a | 
| parent 50321 | df5553c4973f | 
| child 51143 | 0a2371e7ced3 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: HOL/Tools/Qelim/cooper.ML | 
| 23466 | 2 | Author: Amine Chaieb, TU Muenchen | 
| 36802 | 3 | |
| 4 | Presburger arithmetic by Cooper's algorithm. | |
| 23466 | 5 | *) | 
| 6 | ||
| 36799 | 7 | signature COOPER = | 
| 36798 | 8 | sig | 
| 9 | type entry | |
| 10 | val get: Proof.context -> entry | |
| 11 | val del: term list -> attribute | |
| 12 | val add: term list -> attribute | |
| 37117 
59cee8807c29
eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
 wenzelm parents: 
36945diff
changeset | 13 | exception COOPER of string | 
| 36804 | 14 | val conv: Proof.context -> conv | 
| 15 | val tac: bool -> thm list -> thm list -> Proof.context -> int -> tactic | |
| 36798 | 16 | val setup: theory -> theory | 
| 23466 | 17 | end; | 
| 18 | ||
| 36799 | 19 | structure Cooper: COOPER = | 
| 36798 | 20 | struct | 
| 21 | ||
| 36799 | 22 | type entry = simpset * term list; | 
| 36798 | 23 | |
| 24 | val allowed_consts = | |
| 25 |   [@{term "op + :: int => _"}, @{term "op + :: nat => _"},
 | |
| 26 |    @{term "op - :: int => _"}, @{term "op - :: nat => _"},
 | |
| 27 |    @{term "op * :: int => _"}, @{term "op * :: nat => _"},
 | |
| 28 |    @{term "op div :: int => _"}, @{term "op div :: nat => _"},
 | |
| 29 |    @{term "op mod :: int => _"}, @{term "op mod :: nat => _"},
 | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 30 |    @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, 
 | 
| 36798 | 31 |    @{term "op = :: int => _"}, @{term "op = :: nat => _"}, @{term "op = :: bool => _"},
 | 
| 32 |    @{term "op < :: int => _"}, @{term "op < :: nat => _"},
 | |
| 33 |    @{term "op <= :: int => _"}, @{term "op <= :: nat => _"},
 | |
| 34 |    @{term "op dvd :: int => _"}, @{term "op dvd :: nat => _"},
 | |
| 35 |    @{term "abs :: int => _"},
 | |
| 36 |    @{term "max :: int => _"}, @{term "max :: nat => _"},
 | |
| 37 |    @{term "min :: int => _"}, @{term "min :: nat => _"},
 | |
| 38 |    @{term "uminus :: int => _"}, (*@ {term "uminus :: nat => _"},*)
 | |
| 37388 | 39 |    @{term "Not"}, @{term Suc},
 | 
| 36798 | 40 |    @{term "Ex :: (int => _) => _"}, @{term "Ex :: (nat => _) => _"},
 | 
| 41 |    @{term "All :: (int => _) => _"}, @{term "All :: (nat => _) => _"},
 | |
| 42 |    @{term "nat"}, @{term "int"},
 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 43 |    @{term "Num.One"}, @{term "Num.Bit0"}, @{term "Num.Bit1"},
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 44 |    @{term "Num.numeral :: num => int"}, @{term "Num.numeral :: num => nat"},
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 45 |    @{term "Num.neg_numeral :: num => int"},
 | 
| 36798 | 46 |    @{term "0::int"}, @{term "1::int"}, @{term "0::nat"}, @{term "1::nat"},
 | 
| 47 |    @{term "True"}, @{term "False"}];
 | |
| 48 | ||
| 49 | structure Data = Generic_Data | |
| 50 | ( | |
| 51 | type T = simpset * term list; | |
| 52 | val empty = (HOL_ss, allowed_consts); | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
39159diff
changeset | 53 | val extend = I; | 
| 36798 | 54 | fun merge ((ss1, ts1), (ss2, ts2)) = | 
| 55 | (merge_ss (ss1, ss2), Library.merge (op aconv) (ts1, ts2)); | |
| 56 | ); | |
| 57 | ||
| 58 | val get = Data.get o Context.Proof; | |
| 59 | ||
| 60 | fun add ts = Thm.declaration_attribute (fn th => fn context => | |
| 61 | context |> Data.map (fn (ss,ts') => | |
| 62 | (ss addsimps [th], merge (op aconv) (ts',ts) ))) | |
| 63 | ||
| 64 | fun del ts = Thm.declaration_attribute (fn th => fn context => | |
| 65 | context |> Data.map (fn (ss,ts') => | |
| 66 | (ss delsimps [th], subtract (op aconv) ts' ts ))) | |
| 67 | ||
| 27018 | 68 | fun simp_thms_conv ctxt = | 
| 35410 | 69 |   Simplifier.rewrite (Simplifier.context ctxt HOL_basic_ss addsimps @{thms simp_thms});
 | 
| 23484 | 70 | val FWD = Drule.implies_elim_list; | 
| 23466 | 71 | |
| 72 | val true_tm = @{cterm "True"};
 | |
| 73 | val false_tm = @{cterm "False"};
 | |
| 74 | val zdvd1_eq = @{thm "zdvd1_eq"};
 | |
| 75 | val presburger_ss = @{simpset} addsimps [zdvd1_eq];
 | |
| 45196 
78478d938cb8
inlined @{thms} (ML compile-time) allows to get rid of legacy zadd_ac as well (cf. 49e305100097);
 wenzelm parents: 
44121diff
changeset | 76 | val lin_ss = | 
| 
78478d938cb8
inlined @{thms} (ML compile-time) allows to get rid of legacy zadd_ac as well (cf. 49e305100097);
 wenzelm parents: 
44121diff
changeset | 77 |   presburger_ss addsimps (@{thm dvd_eq_mod_eq_0} :: zdvd1_eq :: @{thms add_ac [where 'a=int]});
 | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 78 | |
| 23466 | 79 | val iT = HOLogic.intT | 
| 80 | val bT = HOLogic.boolT; | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 81 | val dest_number = HOLogic.dest_number #> snd; | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 82 | val perhaps_number = try dest_number; | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 83 | val is_number = can dest_number; | 
| 23466 | 84 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 85 | val [miconj, midisj, mieq, mineq, milt, mile, migt, mige, midvd, mindvd, miP] = | 
| 23466 | 86 |     map(instantiate' [SOME @{ctyp "int"}] []) @{thms "minf"};
 | 
| 87 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 88 | val [infDconj, infDdisj, infDdvd,infDndvd,infDP] = | 
| 23466 | 89 |     map(instantiate' [SOME @{ctyp "int"}] []) @{thms "inf_period"};
 | 
| 90 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 91 | val [piconj, pidisj, pieq,pineq,pilt,pile,pigt,pige,pidvd,pindvd,piP] = | 
| 23466 | 92 |     map (instantiate' [SOME @{ctyp "int"}] []) @{thms "pinf"};
 | 
| 93 | ||
| 94 | val [miP, piP] = map (instantiate' [SOME @{ctyp "bool"}] []) [miP, piP];
 | |
| 95 | ||
| 96 | val infDP = instantiate' (map SOME [@{ctyp "int"}, @{ctyp "bool"}]) [] infDP;
 | |
| 97 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 98 | val [[asetconj, asetdisj, aseteq, asetneq, asetlt, asetle, | 
| 23466 | 99 | asetgt, asetge, asetdvd, asetndvd,asetP], | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 100 | [bsetconj, bsetdisj, bseteq, bsetneq, bsetlt, bsetle, | 
| 23466 | 101 |       bsetgt, bsetge, bsetdvd, bsetndvd,bsetP]]  = [@{thms "aset"}, @{thms "bset"}];
 | 
| 102 | ||
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 103 | val [cpmi, cppi] = [@{thm "cpmi"}, @{thm "cppi"}];
 | 
| 23466 | 104 | |
| 105 | val unity_coeff_ex = instantiate' [SOME @{ctyp "int"}] [] @{thm "unity_coeff_ex"};
 | |
| 106 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 107 | val [zdvd_mono,simp_from_to,all_not_ex] = | 
| 23466 | 108 |      [@{thm "zdvd_mono"}, @{thm "simp_from_to"}, @{thm "all_not_ex"}];
 | 
| 109 | ||
| 110 | val [dvd_uminus, dvd_uminus'] = @{thms "uminus_dvd_conv"};
 | |
| 111 | ||
| 112 | val eval_ss = presburger_ss addsimps [simp_from_to] delsimps [insert_iff,bex_triv]; | |
| 113 | val eval_conv = Simplifier.rewrite eval_ss; | |
| 114 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 115 | (* recognising cterm without moving to terms *) | 
| 23466 | 116 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 117 | datatype fm = And of cterm*cterm| Or of cterm*cterm| Eq of cterm | NEq of cterm | 
| 23466 | 118 | | Lt of cterm | Le of cterm | Gt of cterm | Ge of cterm | 
| 119 | | Dvd of cterm*cterm | NDvd of cterm*cterm | Nox | |
| 120 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 121 | fun whatis x ct = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 122 | ( case (term_of ct) of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 123 |   Const(@{const_name HOL.conj},_)$_$_ => And (Thm.dest_binop ct)
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 124 | | Const (@{const_name HOL.disj},_)$_$_ => Or (Thm.dest_binop ct)
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38808diff
changeset | 125 | | Const (@{const_name HOL.eq},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
 | 
| 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38808diff
changeset | 126 | | Const (@{const_name Not},_) $ (Const (@{const_name HOL.eq},_)$y$_) =>
 | 
| 23466 | 127 | if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 128 | | Const (@{const_name Orderings.less}, _) $ y$ z =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 129 | if term_of x aconv y then Lt (Thm.dest_arg ct) | 
| 23466 | 130 | else if term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 131 | | Const (@{const_name Orderings.less_eq}, _) $ y $ z =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 132 | if term_of x aconv y then Le (Thm.dest_arg ct) | 
| 23466 | 133 | else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 134 | | Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_) =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 135 | if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 136 | | Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_)) =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 137 | if term_of x aconv y then | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 138 | NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox | 
| 23466 | 139 | | _ => Nox) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 140 | handle CTERM _ => Nox; | 
| 23466 | 141 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 142 | fun get_pmi_term t = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 143 | let val (x,eq) = | 
| 23466 | 144 | (Thm.dest_abs NONE o Thm.dest_arg o snd o Thm.dest_abs NONE o Thm.dest_arg) | 
| 145 | (Thm.dest_arg t) | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 146 | in (Thm.lambda x o Thm.dest_arg o Thm.dest_arg) eq end; | 
| 23466 | 147 | |
| 148 | val get_pmi = get_pmi_term o cprop_of; | |
| 149 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 150 | val p_v' = @{cpat "?P' :: int => bool"};
 | 
| 23466 | 151 | val q_v' = @{cpat "?Q' :: int => bool"};
 | 
| 152 | val p_v = @{cpat "?P:: int => bool"};
 | |
| 153 | val q_v = @{cpat "?Q:: int => bool"};
 | |
| 154 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 155 | fun myfwd (th1, th2, th3) p q | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 156 | [(th_1,th_2,th_3), (th_1',th_2',th_3')] = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 157 | let | 
| 23466 | 158 | val (mp', mq') = (get_pmi th_1, get_pmi th_1') | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42793diff
changeset | 159 | val mi_th = FWD (Drule.instantiate_normalize ([],[(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1) | 
| 23466 | 160 | [th_1, th_1'] | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42793diff
changeset | 161 | val infD_th = FWD (Drule.instantiate_normalize ([],[(p_v,mp'), (q_v, mq')]) th3) [th_3,th_3'] | 
| 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42793diff
changeset | 162 | val set_th = FWD (Drule.instantiate_normalize ([],[(p_v,p), (q_v,q)]) th2) [th_2, th_2'] | 
| 23466 | 163 | in (mi_th, set_th, infD_th) | 
| 164 | end; | |
| 165 | ||
| 166 | val inst' = fn cts => instantiate' [] (map SOME cts); | |
| 167 | val infDTrue = instantiate' [] [SOME true_tm] infDP; | |
| 168 | val infDFalse = instantiate' [] [SOME false_tm] infDP; | |
| 169 | ||
| 170 | val cadd =  @{cterm "op + :: int => _"}
 | |
| 171 | val cmulC =  @{cterm "op * :: int => _"}
 | |
| 172 | val cminus =  @{cterm "op - :: int => _"}
 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 173 | val cone =  @{cterm "1 :: int"}
 | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 174 | val [addC, mulC, subC] = map term_of [cadd, cmulC, cminus] | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 175 | val [zero, one] = [@{term "0 :: int"}, @{term "1 :: int"}];
 | 
| 23466 | 176 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 177 | fun numeral1 f n = HOLogic.mk_number iT (f (dest_number n)); | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 178 | fun numeral2 f m n = HOLogic.mk_number iT (f (dest_number m) (dest_number n)); | 
| 23466 | 179 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 180 | val [minus1,plus1] = | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 181 | map (fn c => fn t => Thm.apply (Thm.apply c t) cone) [cminus,cadd]; | 
| 23466 | 182 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 183 | fun decomp_pinf x dvd inS [aseteq, asetneq, asetlt, asetle, | 
| 23466 | 184 | asetgt, asetge,asetdvd,asetndvd,asetP, | 
| 185 | infDdvd, infDndvd, asetconj, | |
| 186 | asetdisj, infDconj, infDdisj] cp = | |
| 187 | case (whatis x cp) of | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 188 | And (p,q) => ([p,q], myfwd (piconj, asetconj, infDconj) (Thm.lambda x p) (Thm.lambda x q)) | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 189 | | Or (p,q) => ([p,q], myfwd (pidisj, asetdisj, infDdisj) (Thm.lambda x p) (Thm.lambda x q)) | 
| 23466 | 190 | | Eq t => ([], K (inst' [t] pieq, FWD (inst' [t] aseteq) [inS (plus1 t)], infDFalse)) | 
| 191 | | NEq t => ([], K (inst' [t] pineq, FWD (inst' [t] asetneq) [inS t], infDTrue)) | |
| 192 | | Lt t => ([], K (inst' [t] pilt, FWD (inst' [t] asetlt) [inS t], infDFalse)) | |
| 193 | | Le t => ([], K (inst' [t] pile, FWD (inst' [t] asetle) [inS (plus1 t)], infDFalse)) | |
| 194 | | Gt t => ([], K (inst' [t] pigt, (inst' [t] asetgt), infDTrue)) | |
| 195 | | Ge t => ([], K (inst' [t] pige, (inst' [t] asetge), infDTrue)) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 196 | | Dvd (d,s) => | 
| 23466 | 197 | ([],let val dd = dvd d | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 198 | in K (inst' [d,s] pidvd, FWD (inst' [d,s] asetdvd) [dd],FWD (inst' [d,s] infDdvd) [dd]) end) | 
| 23466 | 199 | | NDvd(d,s) => ([],let val dd = dvd d | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 200 | in K (inst' [d,s] pindvd, FWD (inst' [d,s] asetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end) | 
| 23466 | 201 | | _ => ([], K (inst' [cp] piP, inst' [cp] asetP, inst' [cp] infDP)); | 
| 202 | ||
| 203 | fun decomp_minf x dvd inS [bseteq,bsetneq,bsetlt, bsetle, bsetgt, | |
| 204 | bsetge,bsetdvd,bsetndvd,bsetP, | |
| 205 | infDdvd, infDndvd, bsetconj, | |
| 206 | bsetdisj, infDconj, infDdisj] cp = | |
| 207 | case (whatis x cp) of | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 208 | And (p,q) => ([p,q], myfwd (miconj, bsetconj, infDconj) (Thm.lambda x p) (Thm.lambda x q)) | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 209 | | Or (p,q) => ([p,q], myfwd (midisj, bsetdisj, infDdisj) (Thm.lambda x p) (Thm.lambda x q)) | 
| 23466 | 210 | | Eq t => ([], K (inst' [t] mieq, FWD (inst' [t] bseteq) [inS (minus1 t)], infDFalse)) | 
| 211 | | NEq t => ([], K (inst' [t] mineq, FWD (inst' [t] bsetneq) [inS t], infDTrue)) | |
| 212 | | Lt t => ([], K (inst' [t] milt, (inst' [t] bsetlt), infDTrue)) | |
| 213 | | Le t => ([], K (inst' [t] mile, (inst' [t] bsetle), infDTrue)) | |
| 214 | | Gt t => ([], K (inst' [t] migt, FWD (inst' [t] bsetgt) [inS t], infDFalse)) | |
| 215 | | Ge t => ([], K (inst' [t] mige,FWD (inst' [t] bsetge) [inS (minus1 t)], infDFalse)) | |
| 216 | | Dvd (d,s) => ([],let val dd = dvd d | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 217 | in K (inst' [d,s] midvd, FWD (inst' [d,s] bsetdvd) [dd] , FWD (inst' [d,s] infDdvd) [dd]) end) | 
| 23466 | 218 | | NDvd (d,s) => ([],let val dd = dvd d | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 219 | in K (inst' [d,s] mindvd, FWD (inst' [d,s] bsetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end) | 
| 23466 | 220 | | _ => ([], K (inst' [cp] miP, inst' [cp] bsetP, inst' [cp] infDP)) | 
| 221 | ||
| 222 | (* Canonical linear form for terms, formulae etc.. *) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 223 | fun provelin ctxt t = Goal.prove ctxt [] [] t | 
| 31101 
26c7bb764a38
qualified names for Lin_Arith tactics and simprocs
 haftmann parents: 
30686diff
changeset | 224 | (fn _ => EVERY [simp_tac lin_ss 1, TRY (Lin_Arith.tac ctxt 1)]); | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 225 | fun linear_cmul 0 tm = zero | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 226 | | linear_cmul n tm = case tm of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 227 |       Const (@{const_name Groups.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 228 |     | Const (@{const_name Groups.times}, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 229 |     | Const (@{const_name Groups.minus}, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 230 |     | (m as Const (@{const_name Groups.uminus}, _)) $ a => m $ linear_cmul n a
 | 
| 25768 | 231 | | _ => numeral1 (fn m => n * m) tm; | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 232 | fun earlier [] x y = false | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 233 | | earlier (h::t) x y = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 234 | if h aconv y then false else if h aconv x then true else earlier t x y; | 
| 23466 | 235 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 236 | fun linear_add vars tm1 tm2 = case (tm1, tm2) of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 237 |     (Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c1 $ x1) $ r1,
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 238 |     Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c2 $ x2) $ r2) =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 239 | if x1 = x2 then | 
| 33002 | 240 | let val c = numeral2 Integer.add c1 c2 | 
| 25768 | 241 | in if c = zero then linear_add vars r1 r2 | 
| 242 | else addC$(mulC$c$x1)$(linear_add vars r1 r2) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 243 | end | 
| 25768 | 244 | else if earlier vars x1 x2 then addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2 | 
| 245 | else addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 246 |  | (Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c1 $ x1) $ r1, _) =>
 | 
| 25768 | 247 | addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 248 |  | (_, Const (@{const_name Groups.plus}, _) $ (Const (@{const_name Groups.times}, _) $ c2 $ x2) $ r2) =>
 | 
| 25768 | 249 | addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2 | 
| 33002 | 250 | | (_, _) => numeral2 Integer.add tm1 tm2; | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 251 | |
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 252 | fun linear_neg tm = linear_cmul ~1 tm; | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 253 | fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); | 
| 23466 | 254 | |
| 36806 | 255 | exception COOPER of string; | 
| 23466 | 256 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 257 | fun lint vars tm = if is_number tm then tm else case tm of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 258 |   Const (@{const_name Groups.uminus}, _) $ t => linear_neg (lint vars t)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 259 | | Const (@{const_name Groups.plus}, _) $ s $ t => linear_add vars (lint vars s) (lint vars t)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 260 | | Const (@{const_name Groups.minus}, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t)
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 261 | | Const (@{const_name Groups.times}, _) $ s $ t =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 262 | let val s' = lint vars s | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 263 | val t' = lint vars t | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 264 | in case perhaps_number s' of SOME n => linear_cmul n t' | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 265 | | NONE => (case perhaps_number t' of SOME n => linear_cmul n s' | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 266 | | NONE => raise COOPER "lint: not linear") | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 267 | end | 
| 25768 | 268 | | _ => addC $ (mulC $ one $ tm) $ zero; | 
| 23466 | 269 | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 270 | fun lin (vs as x::_) (Const (@{const_name Not}, _) $ (Const (@{const_name Orderings.less}, T) $ s $ t)) =
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 271 |     lin vs (Const (@{const_name Orderings.less_eq}, T) $ t $ s)
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 272 |   | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name Orderings.less_eq}, T) $ s $ t)) =
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 273 |     lin vs (Const (@{const_name Orderings.less}, T) $ t $ s)
 | 
| 25768 | 274 |   | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
 | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 275 |   | lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) =
 | 
| 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 276 |     HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t)
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38808diff
changeset | 277 |   | lin (vs as x::_) ((b as Const(@{const_name HOL.eq},_))$s$t) =
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 278 | (case lint vs (subC$t$s) of | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 279 | (t as a$(m$c$y)$r) => | 
| 23466 | 280 | if x <> y then b$zero$t | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 281 | else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r | 
| 23466 | 282 | else b$(m$c$y)$(linear_neg r) | 
| 283 | | t => b$zero$t) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 284 | | lin (vs as x::_) (b$s$t) = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 285 | (case lint vs (subC$t$s) of | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 286 | (t as a$(m$c$y)$r) => | 
| 23466 | 287 | if x <> y then b$zero$t | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 288 | else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r | 
| 23466 | 289 | else b$(linear_neg r)$(m$c$y) | 
| 290 | | t => b$zero$t) | |
| 291 | | lin vs fm = fm; | |
| 292 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 293 | fun lint_conv ctxt vs ct = | 
| 23466 | 294 | let val t = term_of ct | 
| 295 | in (provelin ctxt ((HOLogic.eq_const iT)$t$(lint vs t) |> HOLogic.mk_Trueprop)) | |
| 296 | RS eq_reflection | |
| 297 | end; | |
| 298 | ||
| 32398 | 299 | fun is_intrel_type T = T = @{typ "int => int => bool"};
 | 
| 300 | ||
| 301 | fun is_intrel (b$_$_) = is_intrel_type (fastype_of b) | |
| 302 |   | is_intrel (@{term "Not"}$(b$_$_)) = is_intrel_type (fastype_of b)
 | |
| 23466 | 303 | | is_intrel _ = false; | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 304 | |
| 25768 | 305 | fun linearize_conv ctxt vs ct = case term_of ct of | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 306 |   Const(@{const_name Rings.dvd},_)$d$t =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 307 | let | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 308 | val th = Conv.binop_conv (lint_conv ctxt vs) ct | 
| 23466 | 309 | val (d',t') = Thm.dest_binop (Thm.rhs_of th) | 
| 310 | val (dt',tt') = (term_of d', term_of t') | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 311 | in if is_number dt' andalso is_number tt' | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 312 | then Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite presburger_ss)) th | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 313 | else | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 314 | let | 
| 50321 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 315 | val dth = | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 316 | case perhaps_number (term_of d') of | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 317 | SOME d => if d < 0 then | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 318 | (Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (lint_conv ctxt vs))) | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 319 | (Thm.transitive th (inst' [d',t'] dvd_uminus)) | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 320 | handle TERM _ => th) | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 321 | else th | 
| 
df5553c4973f
add check to Cooper's algorithm that left-hand of dvd is a numeral
 hoelzl parents: 
47476diff
changeset | 322 | | NONE => raise COOPER "linearize_conv: not linear" | 
| 23466 | 323 | val d'' = Thm.rhs_of dth |> Thm.dest_arg1 | 
| 324 | in | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 325 | case tt' of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 326 |         Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$_)$_ =>
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 327 | let val x = dest_number c | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 328 | in if x < 0 then Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (lint_conv ctxt vs))) | 
| 23466 | 329 | (Thm.transitive dth (inst' [d'',t'] dvd_uminus')) | 
| 330 | else dth end | |
| 331 | | _ => dth | |
| 332 | end | |
| 333 | end | |
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 334 | | Const (@{const_name Not},_)$(Const(@{const_name Rings.dvd},_)$_$_) => Conv.arg_conv (linearize_conv ctxt vs) ct
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 335 | | t => if is_intrel t | 
| 23466 | 336 | then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop)) | 
| 337 | RS eq_reflection | |
| 36945 | 338 | else Thm.reflexive ct; | 
| 23466 | 339 | |
| 340 | val dvdc = @{cterm "op dvd :: int => _"};
 | |
| 341 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 342 | fun unify ctxt q = | 
| 23466 | 343 | let | 
| 344 | val (e,(cx,p)) = q |> Thm.dest_comb ||> Thm.dest_abs NONE | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 345 | val x = term_of cx | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 346 | val ins = insert (op = : int * int -> bool) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 347 | fun h (acc,dacc) t = | 
| 23466 | 348 | case (term_of t) of | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 349 |     Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
 | 
| 23881 | 350 | if x aconv y andalso member (op =) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38808diff
changeset | 351 |       [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 352 | then (ins (dest_number c) acc,dacc) else (acc,dacc) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 353 |   | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
 | 
| 23881 | 354 | if x aconv y andalso member (op =) | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 355 |        [@{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 356 | then (ins (dest_number c) acc, dacc) else (acc,dacc) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 357 |   | Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$y)$_) =>
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 358 | if x aconv y then (acc,ins (dest_number c) dacc) else (acc,dacc) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 359 |   | Const(@{const_name HOL.conj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 360 |   | Const(@{const_name HOL.disj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
 | 
| 25768 | 361 |   | Const (@{const_name Not},_)$_ => h (acc,dacc) (Thm.dest_arg t)
 | 
| 23466 | 362 | | _ => (acc, dacc) | 
| 363 | val (cs,ds) = h ([],[]) p | |
| 33042 | 364 | val l = Integer.lcms (union (op =) cs ds) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 365 | fun cv k ct = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 366 | let val (tm as b$s$t) = term_of ct | 
| 23466 | 367 | in ((HOLogic.eq_const bT)$tm$(b$(linear_cmul k s)$(linear_cmul k t)) | 
| 368 | |> HOLogic.mk_Trueprop |> provelin ctxt) RS eq_reflection end | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 369 | fun nzprop x = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 370 | let | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 371 | val th = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 372 | Simplifier.rewrite lin_ss | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 373 |       (Thm.apply @{cterm Trueprop} (Thm.apply @{cterm "Not"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 374 |            (Thm.apply (Thm.apply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x))
 | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 375 |            @{cterm "0::int"})))
 | 
| 36945 | 376 | in Thm.equal_elim (Thm.symmetric th) TrueI end; | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 377 | val notz = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 378 | let val tab = fold Inttab.update | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 379 | (ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 380 | in | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 381 | fn ct => the (Inttab.lookup tab (ct |> term_of |> dest_number)) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 382 | handle Option => | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 383 |           (writeln ("noz: Theorems-Table contains no entry for " ^
 | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 384 | Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option) | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 385 | end | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 386 | fun unit_conv t = | 
| 23466 | 387 | case (term_of t) of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 388 |    Const(@{const_name HOL.conj},_)$_$_ => Conv.binop_conv unit_conv t
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 389 |   | Const(@{const_name HOL.disj},_)$_$_ => Conv.binop_conv unit_conv t
 | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 390 |   | Const (@{const_name Not},_)$_ => Conv.arg_conv unit_conv t
 | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 391 |   | Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
 | 
| 23881 | 392 | if x=y andalso member (op =) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38808diff
changeset | 393 |       [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 394 | then cv (l div dest_number c) t else Thm.reflexive t | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 395 |   | Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
 | 
| 23881 | 396 | if x=y andalso member (op =) | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35050diff
changeset | 397 |       [@{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 398 | then cv (l div dest_number c) t else Thm.reflexive t | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 399 |   | Const(@{const_name Rings.dvd},_)$d$(r as (Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$y)$_)) =>
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 400 | if x=y then | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 401 | let | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 402 | val k = l div dest_number c | 
| 23466 | 403 | val kt = HOLogic.mk_number iT k | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 404 | val th1 = inst' [Thm.dest_arg1 t, Thm.dest_arg t] | 
| 23466 | 405 | ((Thm.dest_arg t |> funpow 2 Thm.dest_arg1 |> notz) RS zdvd_mono) | 
| 406 | val (d',t') = (mulC$kt$d, mulC$kt$r) | |
| 407 | val thc = (provelin ctxt ((HOLogic.eq_const iT)$d'$(lint [] d') |> HOLogic.mk_Trueprop)) | |
| 408 | RS eq_reflection | |
| 409 | val tht = (provelin ctxt ((HOLogic.eq_const iT)$t'$(linear_cmul k r) |> HOLogic.mk_Trueprop)) | |
| 410 | RS eq_reflection | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 411 | in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule dvdc thc) tht) end | 
| 23466 | 412 | else Thm.reflexive t | 
| 413 | | _ => Thm.reflexive t | |
| 414 | val uth = unit_conv p | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 415 |   val clt =  Numeral.mk_cnumber @{ctyp "int"} l
 | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 416 | val ltx = Thm.apply (Thm.apply cmulC clt) cx | 
| 23466 | 417 | val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth) | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 418 | val th' = inst' [Thm.lambda ltx (Thm.rhs_of uth), clt] unity_coeff_ex | 
| 36945 | 419 | val thf = Thm.transitive th | 
| 420 | (Thm.transitive (Thm.symmetric (Thm.beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th') | |
| 23466 | 421 | val (lth,rth) = Thm.dest_comb (cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true | 
| 36945 | 422 | ||> Thm.beta_conversion true |>> Thm.symmetric | 
| 423 | in Thm.transitive (Thm.transitive lth thf) rth end; | |
| 23466 | 424 | |
| 425 | ||
| 426 | val emptyIS = @{cterm "{}::int set"};
 | |
| 427 | val insert_tm = @{cterm "insert :: int => _"};
 | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 428 | fun mkISet cts = fold_rev (Thm.apply insert_tm #> Thm.apply) cts emptyIS; | 
| 39159 | 429 | val eqelem_imp_imp = @{thm eqelem_imp_iff} RS iffD1;
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 430 | 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 | 
| 23466 | 431 | |> Thm.dest_abs NONE |> snd |> Thm.dest_fun |> Thm.dest_arg) | 
| 432 | [asetP,bsetP]; | |
| 433 | ||
| 434 | val D_tm = @{cpat "?D::int"};
 | |
| 435 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 436 | fun cooperex_conv ctxt vs q = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 437 | let | 
| 23466 | 438 | |
| 439 | val uth = unify ctxt q | |
| 440 | val (x,p) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of uth)) | |
| 441 | val ins = insert (op aconvc) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 442 | fun h t (bacc,aacc,dacc) = | 
| 23466 | 443 | case (whatis x t) of | 
| 444 | And (p,q) => h q (h p (bacc,aacc,dacc)) | |
| 445 | | Or (p,q) => h q (h p (bacc,aacc,dacc)) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 446 | | Eq t => (ins (minus1 t) bacc, | 
| 23466 | 447 | ins (plus1 t) aacc,dacc) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 448 | | NEq t => (ins t bacc, | 
| 23466 | 449 | ins t aacc, dacc) | 
| 450 | | Lt t => (bacc, ins t aacc, dacc) | |
| 451 | | Le t => (bacc, ins (plus1 t) aacc,dacc) | |
| 452 | | Gt t => (ins t bacc, aacc,dacc) | |
| 453 | | Ge t => (ins (minus1 t) bacc, aacc,dacc) | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 454 | | Dvd (d,_) => (bacc,aacc,insert (op =) (term_of d |> dest_number) dacc) | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 455 | | NDvd (d,_) => (bacc,aacc,insert (op =) (term_of d|> dest_number) dacc) | 
| 23466 | 456 | | _ => (bacc, aacc, dacc) | 
| 457 | val (b0,a0,ds) = h p ([],[],[]) | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 458 | val d = Integer.lcms ds | 
| 23582 | 459 |  val cd = Numeral.mk_cnumber @{ctyp "int"} d
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 460 | fun divprop x = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 461 | let | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 462 | val th = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 463 | Simplifier.rewrite lin_ss | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 464 |       (Thm.apply @{cterm Trueprop}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 465 |            (Thm.apply (Thm.apply dvdc (Numeral.mk_cnumber @{ctyp "int"} x)) cd))
 | 
| 36945 | 466 | in Thm.equal_elim (Thm.symmetric th) TrueI end; | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 467 | val dvd = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 468 | let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 469 | fn ct => the (Inttab.lookup tab (term_of ct |> dest_number)) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 470 | handle Option => | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 471 |         (writeln ("dvd: Theorems-Table contains no entry for" ^
 | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 472 | Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option) | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 473 | end | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 474 | val dp = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 475 | let val th = Simplifier.rewrite lin_ss | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 476 |       (Thm.apply @{cterm Trueprop}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 477 |            (Thm.apply (Thm.apply @{cterm "op < :: int => _"} @{cterm "0::int"}) cd))
 | 
| 36945 | 478 | in Thm.equal_elim (Thm.symmetric th) TrueI end; | 
| 23466 | 479 | (* A and B set *) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 480 | local | 
| 23466 | 481 |      val insI1 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI1"}
 | 
| 482 |      val insI2 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI2"}
 | |
| 483 | in | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 484 | fun provein x S = | 
| 23466 | 485 | case term_of S of | 
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
31101diff
changeset | 486 |         Const(@{const_name Orderings.bot}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 487 |       | Const(@{const_name insert}, _) $ y $ _ =>
 | 
| 23466 | 488 | let val (cy,S') = Thm.dest_binop S | 
| 489 | in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1 | |
| 36945 | 490 | else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2) | 
| 23466 | 491 | (provein x S') | 
| 492 | end | |
| 493 | end | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 494 | |
| 23466 | 495 | val al = map (lint vs o term_of) a0 | 
| 496 | val bl = map (lint vs o term_of) b0 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 497 | val (sl,s0,f,abths,cpth) = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 498 | if length (distinct (op aconv) bl) <= length (distinct (op aconv) al) | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 499 | then | 
| 23466 | 500 | (bl,b0,decomp_minf, | 
| 36945 | 501 | fn B => (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp) | 
| 23466 | 502 | [bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@ | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 503 | (map (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)])) | 
| 23466 | 504 | [bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj, | 
| 505 | bsetdisj,infDconj, infDdisj]), | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 506 | cpmi) | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 507 | else (al,a0,decomp_pinf,fn A => | 
| 36945 | 508 | (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]) th) dp) | 
| 23466 | 509 | [aseteq,asetneq,asetlt, asetle, asetgt,asetge])@ | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 510 | (map (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)])) | 
| 23466 | 511 | [asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj, | 
| 512 | asetdisj,infDconj, infDdisj]),cppi) | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 513 | val cpth = | 
| 23466 | 514 | let | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 515 | val sths = map (fn (tl,t0) => | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 516 | if tl = term_of t0 | 
| 23466 | 517 |                       then instantiate' [SOME @{ctyp "int"}] [SOME t0] refl
 | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 518 | else provelin ctxt ((HOLogic.eq_const iT)$tl$(term_of t0) | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 519 | |> HOLogic.mk_Trueprop)) | 
| 23466 | 520 | (sl ~~ s0) | 
| 521 | val csl = distinct (op aconvc) (map (cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths) | |
| 522 | val S = mkISet csl | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 523 | val inStab = fold (fn ct => fn tab => Termtab.update (term_of ct, provein ct S) tab) | 
| 23466 | 524 | csl Termtab.empty | 
| 525 |    val eqelem_th = instantiate' [SOME @{ctyp "int"}] [NONE,NONE, SOME S] eqelem_imp_imp
 | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 526 | val inS = | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 527 | let | 
| 23466 | 528 | val tab = fold Termtab.update | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 529 | (map (fn eq => | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 530 | let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 531 | val th = if term_of s = term_of t | 
| 33035 | 532 | then the (Termtab.lookup inStab (term_of s)) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 533 | else FWD (instantiate' [] [SOME s, SOME t] eqelem_th) | 
| 33035 | 534 | [eq, the (Termtab.lookup inStab (term_of s))] | 
| 23466 | 535 | in (term_of t, th) end) | 
| 536 | sths) Termtab.empty | |
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 537 | in | 
| 33035 | 538 | fn ct => the (Termtab.lookup tab (term_of ct)) | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 539 | handle Option => | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 540 |               (writeln ("inS: No theorem for " ^ Syntax.string_of_term ctxt (Thm.term_of ct));
 | 
| 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 541 | raise Option) | 
| 23466 | 542 | end | 
| 543 | val (inf, nb, pd) = divide_and_conquer (f x dvd inS (abths S)) p | |
| 544 | in [dp, inf, nb, pd] MRS cpth | |
| 545 | end | |
| 546 | val cpth' = Thm.transitive uth (cpth RS eq_reflection) | |
| 27018 | 547 | in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv) (Thm.rhs_of cpth')) | 
| 23466 | 548 | end; | 
| 549 | ||
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 550 | fun literals_conv bops uops env cv = | 
| 23466 | 551 | let fun h t = | 
| 32429 
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
 wenzelm parents: 
32398diff
changeset | 552 | case (term_of t) of | 
| 36797 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 553 | b$_$_ => if member (op aconv) bops b then Conv.binop_conv h t else cv env t | 
| 
cb074cec7a30
dropped unused bindings; avoid open (documents dependency on generated code more explicitly)
 haftmann parents: 
36717diff
changeset | 554 | | u$_ => if member (op aconv) uops u then Conv.arg_conv h t else cv env t | 
| 23466 | 555 | | _ => cv env t | 
| 556 | in h end; | |
| 557 | ||
| 558 | fun integer_nnf_conv ctxt env = | |
| 559 | nnf_conv then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt); | |
| 560 | ||
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 561 | val conv_ss = HOL_basic_ss addsimps | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 562 |   (@{thms simp_thms} @ take 4 @{thms ex_simps} @ [not_all, all_not_ex, @{thm ex_disj_distrib}]);
 | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 563 | |
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 564 | fun conv ctxt p = | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 565 | Qelim.gen_qelim_conv (Simplifier.rewrite conv_ss) (Simplifier.rewrite presburger_ss) (Simplifier.rewrite conv_ss) | 
| 44121 | 566 | (cons o term_of) (Misc_Legacy.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 567 | (cooperex_conv ctxt) p | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 568 | handle CTERM s => raise COOPER "bad cterm" | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 569 | | THM s => raise COOPER "bad thm" | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 570 | | TYPE s => raise COOPER "bad type" | 
| 23466 | 571 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 572 | fun add_bools t = | 
| 36807 | 573 | let | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 574 |     val ops = [@{term "op = :: int => _"}, @{term "op < :: int => _"}, @{term "op <= :: int => _"},
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 575 |       @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, @{term "op = :: bool => _"},
 | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 576 |       @{term "Not"}, @{term "All :: (int => _) => _"},
 | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 577 |       @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}];
 | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 578 | val is_op = member (op =) ops; | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 579 | val skip = not (fastype_of t = HOLogic.boolT) | 
| 36807 | 580 | in case t of | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 581 | (l as f $ a) $ b => if skip orelse is_op f then add_bools b o add_bools l | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 582 | else insert (op aconv) t | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 583 | | f $ a => if skip orelse is_op f then add_bools a o add_bools f | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 584 | else insert (op aconv) t | 
| 42284 | 585 | | Abs p => add_bools (snd (Syntax_Trans.variant_abs p)) (* FIXME !? *) | 
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 586 | | _ => if skip orelse is_op t then I else insert (op aconv) t | 
| 36807 | 587 | end; | 
| 588 | ||
| 36832 | 589 | fun descend vs (abs as (_, xT, _)) = | 
| 590 | let | |
| 42284 | 591 | val (xn', p') = Syntax_Trans.variant_abs abs; (* FIXME !? *) | 
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 592 | in ((xn', xT) :: vs, p') end; | 
| 36832 | 593 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 594 | local structure Proc = Cooper_Procedure in | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 595 | |
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 596 | fun num_of_term vs (Free vT) = Proc.Bound (find_index (fn vT' => vT' = vT) vs) | 
| 36832 | 597 | | num_of_term vs (Term.Bound i) = Proc.Bound i | 
| 598 |   | num_of_term vs @{term "0::int"} = Proc.C 0
 | |
| 599 |   | num_of_term vs @{term "1::int"} = Proc.C 1
 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 600 |   | num_of_term vs (t as Const (@{const_name numeral}, _) $ _) =
 | 
| 36832 | 601 | Proc.C (dest_number t) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 602 |   | num_of_term vs (t as Const (@{const_name neg_numeral}, _) $ _) =
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 603 | Proc.Neg (Proc.C (dest_number t)) | 
| 36832 | 604 |   | num_of_term vs (Const (@{const_name Groups.uminus}, _) $ t') =
 | 
| 605 | Proc.Neg (num_of_term vs t') | |
| 606 |   | num_of_term vs (Const (@{const_name Groups.plus}, _) $ t1 $ t2) =
 | |
| 607 | Proc.Add (num_of_term vs t1, num_of_term vs t2) | |
| 608 |   | num_of_term vs (Const (@{const_name Groups.minus}, _) $ t1 $ t2) =
 | |
| 609 | Proc.Sub (num_of_term vs t1, num_of_term vs t2) | |
| 610 |   | num_of_term vs (Const (@{const_name Groups.times}, _) $ t1 $ t2) =
 | |
| 611 | (case perhaps_number t1 | |
| 612 | of SOME n => Proc.Mul (n, num_of_term vs t2) | |
| 613 | | NONE => (case perhaps_number t2 | |
| 614 | of SOME n => Proc.Mul (n, num_of_term vs t1) | |
| 615 | | NONE => raise COOPER "reification: unsupported kind of multiplication")) | |
| 616 | | num_of_term _ _ = raise COOPER "reification: bad term"; | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23582diff
changeset | 617 | |
| 36832 | 618 | fun fm_of_term ps vs (Const (@{const_name True}, _)) = Proc.T
 | 
| 619 |   | fm_of_term ps vs (Const (@{const_name False}, _)) = Proc.F
 | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 620 |   | fm_of_term ps vs (Const (@{const_name HOL.conj}, _) $ t1 $ t2) =
 | 
| 36832 | 621 | Proc.And (fm_of_term ps vs t1, fm_of_term ps vs t2) | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 622 |   | fm_of_term ps vs (Const (@{const_name HOL.disj}, _) $ t1 $ t2) =
 | 
| 36832 | 623 | Proc.Or (fm_of_term ps vs t1, fm_of_term ps vs t2) | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 624 |   | fm_of_term ps vs (Const (@{const_name HOL.implies}, _) $ t1 $ t2) =
 | 
| 36832 | 625 | Proc.Imp (fm_of_term ps vs t1, fm_of_term ps vs t2) | 
| 626 |   | fm_of_term ps vs (@{term "op = :: bool => _ "} $ t1 $ t2) =
 | |
| 627 | Proc.Iff (fm_of_term ps vs t1, fm_of_term ps vs t2) | |
| 628 |   | fm_of_term ps vs (Const (@{const_name Not}, _) $ t') =
 | |
| 629 | Proc.Not (fm_of_term ps vs t') | |
| 38558 | 630 |   | fm_of_term ps vs (Const (@{const_name Ex}, _) $ Abs abs) =
 | 
| 36832 | 631 | Proc.E (uncurry (fm_of_term ps) (descend vs abs)) | 
| 38558 | 632 |   | fm_of_term ps vs (Const (@{const_name All}, _) $ Abs abs) =
 | 
| 36832 | 633 | Proc.A (uncurry (fm_of_term ps) (descend vs abs)) | 
| 634 |   | fm_of_term ps vs (@{term "op = :: int => _"} $ t1 $ t2) =
 | |
| 635 | Proc.Eq (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | |
| 636 |   | fm_of_term ps vs (Const (@{const_name Orderings.less_eq}, _) $ t1 $ t2) =
 | |
| 637 | Proc.Le (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | |
| 638 |   | fm_of_term ps vs (Const (@{const_name Orderings.less}, _) $ t1 $ t2) =
 | |
| 639 | Proc.Lt (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | |
| 640 |   | fm_of_term ps vs (Const (@{const_name Rings.dvd}, _) $ t1 $ t2) =
 | |
| 641 | (case perhaps_number t1 | |
| 642 | of SOME n => Proc.Dvd (n, num_of_term vs t2) | |
| 643 | | NONE => raise COOPER "reification: unsupported dvd") | |
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 644 | | fm_of_term ps vs t = let val n = find_index (fn t' => t aconv t') ps | 
| 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 645 | in if n > 0 then Proc.Closed n else raise COOPER "reification: unknown term" end; | 
| 23466 | 646 | |
| 36832 | 647 | fun term_of_num vs (Proc.C i) = HOLogic.mk_number HOLogic.intT i | 
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 648 | | term_of_num vs (Proc.Bound n) = Free (nth vs n) | 
| 36832 | 649 | | term_of_num vs (Proc.Neg t') = | 
| 650 |       @{term "uminus :: int => _"} $ term_of_num vs t'
 | |
| 651 | | term_of_num vs (Proc.Add (t1, t2)) = | |
| 652 |       @{term "op + :: int => _"} $ term_of_num vs t1 $ term_of_num vs t2
 | |
| 653 | | term_of_num vs (Proc.Sub (t1, t2)) = | |
| 654 |       @{term "op - :: int => _"} $ term_of_num vs t1 $ term_of_num vs t2
 | |
| 655 | | term_of_num vs (Proc.Mul (i, t2)) = | |
| 656 |       @{term "op * :: int => _"} $ HOLogic.mk_number HOLogic.intT i $ term_of_num vs t2
 | |
| 657 | | term_of_num vs (Proc.Cn (n, i, t')) = | |
| 658 | term_of_num vs (Proc.Add (Proc.Mul (i, Proc.Bound n), t')); | |
| 23466 | 659 | |
| 45740 | 660 | fun term_of_fm ps vs Proc.T = @{term True}
 | 
| 661 |   | term_of_fm ps vs Proc.F = @{term False}
 | |
| 36832 | 662 | | term_of_fm ps vs (Proc.And (t1, t2)) = HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | 
| 663 | | term_of_fm ps vs (Proc.Or (t1, t2)) = HOLogic.disj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 664 | | term_of_fm ps vs (Proc.Imp (t1, t2)) = HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 665 |   | term_of_fm ps vs (Proc.Iff (t1, t2)) = @{term "op = :: bool => _"} $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
 | |
| 666 | | term_of_fm ps vs (Proc.Not t') = HOLogic.Not $ term_of_fm ps vs t' | |
| 667 |   | term_of_fm ps vs (Proc.Eq t') = @{term "op = :: int => _ "} $ term_of_num vs t'$ @{term "0::int"}
 | |
| 668 | | term_of_fm ps vs (Proc.NEq t') = term_of_fm ps vs (Proc.Not (Proc.Eq t')) | |
| 669 |   | term_of_fm ps vs (Proc.Lt t') = @{term "op < :: int => _ "} $ term_of_num vs t' $ @{term "0::int"}
 | |
| 670 |   | term_of_fm ps vs (Proc.Le t') = @{term "op <= :: int => _ "} $ term_of_num vs t' $ @{term "0::int"}
 | |
| 671 |   | term_of_fm ps vs (Proc.Gt t') = @{term "op < :: int => _ "} $ @{term "0::int"} $ term_of_num vs t'
 | |
| 672 |   | term_of_fm ps vs (Proc.Ge t') = @{term "op <= :: int => _ "} $ @{term "0::int"} $ term_of_num vs t'
 | |
| 673 |   | term_of_fm ps vs (Proc.Dvd (i, t')) = @{term "op dvd :: int => _ "} $
 | |
| 674 | HOLogic.mk_number HOLogic.intT i $ term_of_num vs t' | |
| 675 | | term_of_fm ps vs (Proc.NDvd (i, t')) = term_of_fm ps vs (Proc.Not (Proc.Dvd (i, t'))) | |
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 676 | | term_of_fm ps vs (Proc.Closed n) = nth ps n | 
| 36832 | 677 | | term_of_fm ps vs (Proc.NClosed n) = term_of_fm ps vs (Proc.Not (Proc.Closed n)); | 
| 23466 | 678 | |
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 679 | fun procedure t = | 
| 23713 | 680 | let | 
| 36833 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 681 | val vs = Term.add_frees t []; | 
| 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 682 | val ps = add_bools t []; | 
| 
9628f969d843
represent de-Bruin indices simply by position in list
 haftmann parents: 
36832diff
changeset | 683 | in (term_of_fm ps vs o Proc.pa o fm_of_term ps vs) t end; | 
| 23466 | 684 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 685 | end; | 
| 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 686 | |
| 38808 | 687 | val (_, oracle) = Context.>>> (Context.map_theory_result | 
| 688 |   (Thm.add_oracle (@{binding cooper},
 | |
| 689 | (fn (ctxt, t) => | |
| 42361 | 690 | (Thm.cterm_of (Proof_Context.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop) | 
| 38808 | 691 | (t, procedure t))))); | 
| 36802 | 692 | |
| 693 | val comp_ss = HOL_ss addsimps @{thms semiring_norm};
 | |
| 694 | ||
| 695 | fun strip_objimp ct = | |
| 696 | (case Thm.term_of ct of | |
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 697 |     Const (@{const_name HOL.implies}, _) $ _ $ _ =>
 | 
| 36802 | 698 | let val (A, B) = Thm.dest_binop ct | 
| 699 | in A :: strip_objimp B end | |
| 700 | | _ => [ct]); | |
| 701 | ||
| 702 | fun strip_objall ct = | |
| 703 | case term_of ct of | |
| 38558 | 704 |   Const (@{const_name All}, _) $ Abs (xn,xT,p) => 
 | 
| 36802 | 705 | let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct | 
| 706 | in apfst (cons (a,v)) (strip_objall t') | |
| 707 | end | |
| 708 | | _ => ([],ct); | |
| 709 | ||
| 710 | local | |
| 711 | val all_maxscope_ss = | |
| 712 |      HOL_basic_ss addsimps map (fn th => th RS sym) @{thms "all_simps"}
 | |
| 713 | in | |
| 42793 | 714 | fun thin_prems_tac ctxt P = simp_tac all_maxscope_ss THEN' | 
| 36802 | 715 | CSUBGOAL (fn (p', i) => | 
| 716 | let | |
| 717 | val (qvs, p) = strip_objall (Thm.dest_arg p') | |
| 718 | val (ps, c) = split_last (strip_objimp p) | |
| 719 | val qs = filter P ps | |
| 720 |      val q = if P c then c else @{cterm "False"}
 | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 721 | val ng = fold_rev (fn (a,v) => fn t => Thm.apply a (Thm.lambda v t)) qvs | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 722 |          (fold_rev (fn p => fn q => Thm.apply (Thm.apply @{cterm HOL.implies} p) q) qs q)
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 723 |      val g = Thm.apply (Thm.apply @{cterm "op ==>"} (Thm.apply @{cterm "Trueprop"} ng)) p'
 | 
| 36802 | 724 |      val ntac = (case qs of [] => q aconvc @{cterm "False"}
 | 
| 725 | | _ => false) | |
| 726 | in | |
| 47476 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 727 | if ntac then no_tac | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 728 | else | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 729 | (case try (fn () => | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 730 | Goal.prove_internal [] g (K (blast_tac (put_claset HOL_cs ctxt) 1))) () of | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 731 | NONE => no_tac | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 732 | | SOME r => rtac r i) | 
| 36802 | 733 | end) | 
| 23466 | 734 | end; | 
| 36802 | 735 | |
| 736 | local | |
| 737 | fun isnum t = case t of | |
| 738 |    Const(@{const_name Groups.zero},_) => true
 | |
| 739 |  | Const(@{const_name Groups.one},_) => true
 | |
| 37388 | 740 |  | @{term Suc}$s => isnum s
 | 
| 36802 | 741 |  | @{term "nat"}$s => isnum s
 | 
| 742 |  | @{term "int"}$s => isnum s
 | |
| 743 |  | Const(@{const_name Groups.uminus},_)$s => isnum s
 | |
| 744 |  | Const(@{const_name Groups.plus},_)$l$r => isnum l andalso isnum r
 | |
| 745 |  | Const(@{const_name Groups.times},_)$l$r => isnum l andalso isnum r
 | |
| 746 |  | Const(@{const_name Groups.minus},_)$l$r => isnum l andalso isnum r
 | |
| 747 |  | Const(@{const_name Power.power},_)$l$r => isnum l andalso isnum r
 | |
| 748 |  | Const(@{const_name Divides.mod},_)$l$r => isnum l andalso isnum r
 | |
| 749 |  | Const(@{const_name Divides.div},_)$l$r => isnum l andalso isnum r
 | |
| 36831 
3037d6810fca
tuned code; toward a tightended interface with generated code
 haftmann parents: 
36807diff
changeset | 750 | | _ => is_number t orelse can HOLogic.dest_nat t | 
| 36802 | 751 | |
| 752 | fun ty cts t = | |
| 753 | if not (member (op =) [HOLogic.intT, HOLogic.natT, HOLogic.boolT] (typ_of (ctyp_of_term t))) then false | |
| 754 | else case term_of t of | |
| 755 |       c$l$r => if member (op =) [@{term"op *::int => _"}, @{term"op *::nat => _"}] c
 | |
| 756 | then not (isnum l orelse isnum r) | |
| 757 | else not (member (op aconv) cts c) | |
| 758 | | c$_ => not (member (op aconv) cts c) | |
| 759 | | c => not (member (op aconv) cts c) | |
| 760 | ||
| 761 | val term_constants = | |
| 762 | let fun h acc t = case t of | |
| 763 | Const _ => insert (op aconv) t acc | |
| 764 | | a$b => h (h acc a) b | |
| 765 | | Abs (_,_,t) => h acc t | |
| 766 | | _ => acc | |
| 767 | in h [] end; | |
| 768 | in | |
| 769 | fun is_relevant ctxt ct = | |
| 770 | subset (op aconv) (term_constants (term_of ct) , snd (get ctxt)) | |
| 44121 | 771 |  andalso forall (fn Free (_,T) => member (op =) [@{typ int}, @{typ nat}] T) (Misc_Legacy.term_frees (term_of ct))
 | 
| 772 |  andalso forall (fn Var (_,T) => member (op =) [@{typ int}, @{typ nat}] T) (Misc_Legacy.term_vars (term_of ct));
 | |
| 36802 | 773 | |
| 774 | fun int_nat_terms ctxt ct = | |
| 775 | let | |
| 776 | val cts = snd (get ctxt) | |
| 777 | fun h acc t = if ty cts t then insert (op aconvc) t acc else | |
| 778 | case (term_of t) of | |
| 779 | _$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | |
| 780 | | Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc)) | |
| 781 | | _ => acc | |
| 782 | in h [] ct end | |
| 783 | end; | |
| 784 | ||
| 785 | fun generalize_tac f = CSUBGOAL (fn (p, i) => PRIMITIVE (fn st => | |
| 786 | let | |
| 787 |    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
 | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45740diff
changeset | 788 | fun gen x t = Thm.apply (all (ctyp_of_term x)) (Thm.lambda x t) | 
| 36802 | 789 | val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p) | 
| 790 | val p' = fold_rev gen ts p | |
| 36945 | 791 | in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); | 
| 36802 | 792 | |
| 793 | local | |
| 794 | val ss1 = comp_ss | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 795 |   addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm "zdvd_int"}] 
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 796 | @ map (fn r => r RS sym) | 
| 36802 | 797 |         [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
 | 
| 798 |          @{thm "zmult_int"}]
 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 799 |   |> Splitter.add_split @{thm "zdiff_int_split"}
 | 
| 36802 | 800 | |
| 801 | val ss2 = HOL_basic_ss | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 802 |   addsimps [@{thm "nat_0_le"}, @{thm "int_numeral"},
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 803 |             @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, 
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 804 |             @{thm "le_numeral_extra"(3)}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}]
 | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 805 |   |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
 | 
| 36802 | 806 | val div_mod_ss = HOL_basic_ss addsimps @{thms simp_thms}
 | 
| 36945 | 807 | @ map (Thm.symmetric o mk_meta_eq) | 
| 36802 | 808 |     [@{thm "dvd_eq_mod_eq_0"},
 | 
| 809 |      @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"}, 
 | |
| 810 |      @{thm "mod_add_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}]
 | |
| 47142 | 811 |   @ [@{thm "mod_self"}, @{thm "mod_by_0"}, @{thm "div_by_0"},
 | 
| 36802 | 812 |      @{thm "div_0"}, @{thm "mod_0"}, @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, 
 | 
| 813 |      @{thm "mod_1"}, @{thm "Suc_eq_plus1"}]
 | |
| 814 |   @ @{thms add_ac}
 | |
| 43594 | 815 |  addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}]
 | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 816 | val splits_ss = | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 817 |   comp_ss addsimps [@{thm "mod_div_equality'"}]
 | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 818 | |> fold Splitter.add_split | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 819 |     [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, 
 | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45196diff
changeset | 820 |      @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}]
 | 
| 36802 | 821 | in | 
| 822 | fun nat_to_int_tac ctxt = | |
| 823 | simp_tac (Simplifier.context ctxt ss1) THEN_ALL_NEW | |
| 824 | simp_tac (Simplifier.context ctxt ss2) THEN_ALL_NEW | |
| 825 | simp_tac (Simplifier.context ctxt comp_ss); | |
| 826 | ||
| 827 | fun div_mod_tac ctxt i = simp_tac (Simplifier.context ctxt div_mod_ss) i; | |
| 828 | fun splits_tac ctxt i = simp_tac (Simplifier.context ctxt splits_ss) i; | |
| 829 | end; | |
| 830 | ||
| 36804 | 831 | fun core_tac ctxt = CSUBGOAL (fn (p, i) => | 
| 36805 | 832 | let | 
| 36802 | 833 | val cpth = | 
| 834 | if !quick_and_dirty | |
| 36805 | 835 | then oracle (ctxt, Envir.beta_norm (Pattern.eta_long [] (term_of (Thm.dest_arg p)))) | 
| 36804 | 836 | else Conv.arg_conv (conv ctxt) p | 
| 36802 | 837 | val p' = Thm.rhs_of cpth | 
| 36945 | 838 | val th = Thm.implies_intr p' (Thm.equal_elim (Thm.symmetric cpth) (Thm.assume p')) | 
| 36802 | 839 | in rtac th i end | 
| 840 | handle COOPER _ => no_tac); | |
| 841 | ||
| 842 | fun finish_tac q = SUBGOAL (fn (_, i) => | |
| 843 | (if q then I else TRY) (rtac TrueI i)); | |
| 844 | ||
| 47476 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 845 | fun tac elim add_ths del_ths = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} =>
 | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 846 | let | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 847 | val ss = Simplifier.context ctxt (fst (get ctxt)) delsimps del_ths addsimps add_ths | 
| 36802 | 848 | val aprems = Arith_Data.get_arith_facts ctxt | 
| 47476 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 849 | in | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 850 | Method.insert_tac aprems | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 851 | THEN_ALL_NEW Object_Logic.full_atomize_tac | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 852 | THEN_ALL_NEW CONVERSION Thm.eta_long_conversion | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 853 | THEN_ALL_NEW simp_tac ss | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 854 | THEN_ALL_NEW (TRY o generalize_tac (int_nat_terms ctxt)) | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 855 | THEN_ALL_NEW Object_Logic.full_atomize_tac | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 856 | THEN_ALL_NEW (thin_prems_tac ctxt (is_relevant ctxt)) | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 857 | THEN_ALL_NEW Object_Logic.full_atomize_tac | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 858 | THEN_ALL_NEW div_mod_tac ctxt | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 859 | THEN_ALL_NEW splits_tac ctxt | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 860 | THEN_ALL_NEW simp_tac ss | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 861 | THEN_ALL_NEW CONVERSION Thm.eta_long_conversion | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 862 | THEN_ALL_NEW nat_to_int_tac ctxt | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 863 | THEN_ALL_NEW (core_tac ctxt) | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 864 | THEN_ALL_NEW finish_tac elim | 
| 
92d1c566ebbf
refined Cooper.tac / "presburger" method: Subgoal.FOCUS_PARAMS allows to solve more problems with outer quantifiers, e.g "!!x. [| 0 <= (x::int); x div 2 < f x |] ==> x < f x * 2";
 wenzelm parents: 
47432diff
changeset | 865 | end 1); | 
| 36802 | 866 | |
| 867 | ||
| 868 | (* theory setup *) | |
| 869 | ||
| 870 | local | |
| 871 | ||
| 872 | fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); | |
| 873 | ||
| 874 | val constsN = "consts"; | |
| 875 | val any_keyword = keyword constsN | |
| 876 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; | |
| 877 | val terms = thms >> map (term_of o Drule.dest_term); | |
| 878 | ||
| 879 | fun optional scan = Scan.optional scan []; | |
| 880 | ||
| 881 | in | |
| 882 | ||
| 883 | val setup = | |
| 884 |   Attrib.setup @{binding presburger}
 | |
| 885 | ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del || | |
| 886 | optional (keyword constsN |-- terms) >> add) "data for Cooper's algorithm" | |
| 36804 | 887 | #> Arith_Data.add_tactic "Presburger arithmetic" (K (tac true [] [])); | 
| 36802 | 888 | |
| 889 | end; | |
| 890 | ||
| 891 | end; |