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