| author | paulson | 
| Thu, 09 Sep 2004 11:10:16 +0200 | |
| changeset 15194 | ddbbab501213 | 
| parent 15088 | b8a95eadbc14 | 
| child 15195 | 197e00ce3f20 | 
| permissions | -rw-r--r-- | 
| 10413 | 1 | (* Title: Pure/meta_simplifier.ML | 
| 2 | ID: $Id$ | |
| 11672 | 3 | Author: Tobias Nipkow and Stefan Berghofer | 
| 10413 | 4 | |
| 11672 | 5 | Meta-level Simplification. | 
| 10413 | 6 | *) | 
| 7 | ||
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 8 | infix 4 | 
| 15023 | 9 | addsimps delsimps addeqcongs deleqcongs addcongs delcongs addsimprocs delsimprocs | 
| 10 | setmksimps setmkcong setmksym setmkeqTrue settermless setsubgoaler | |
| 11 | setloop addloop delloop setSSolver addSSolver setSolver addSolver; | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 12 | |
| 11672 | 13 | signature BASIC_META_SIMPLIFIER = | 
| 14 | sig | |
| 15023 | 15 | val debug_simp: bool ref | 
| 11672 | 16 | val trace_simp: bool ref | 
| 13828 | 17 | val simp_depth_limit: int ref | 
| 15023 | 18 | type rrule | 
| 19 | type cong | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 20 | type solver | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 21 | val mk_solver: string -> (thm list -> int -> tactic) -> solver | 
| 15023 | 22 | type simpset | 
| 23 | type proc | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 24 | val rep_ss: simpset -> | 
| 15023 | 25 |    {rules: rrule Net.net,
 | 
| 26 | prems: thm list, | |
| 27 | bounds: string list, | |
| 28 | depth: int} * | |
| 29 |    {congs: (string * cong) list * string list,
 | |
| 30 | procs: proc Net.net, | |
| 31 | mk_rews: | |
| 32 |      {mk: thm -> thm list,
 | |
| 33 | mk_cong: thm -> thm, | |
| 34 | mk_sym: thm -> thm option, | |
| 35 | mk_eq_True: thm -> thm option}, | |
| 36 | termless: term * term -> bool, | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 37 | subgoal_tac: simpset -> int -> tactic, | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 38 | loop_tacs: (string * (int -> tactic)) list, | 
| 15023 | 39 | solvers: solver list * solver list} | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 40 | val print_ss: simpset -> unit | 
| 15023 | 41 | val empty_ss: simpset | 
| 42 | val merge_ss: simpset * simpset -> simpset | |
| 43 | type simproc | |
| 44 | val mk_simproc: string -> cterm list -> | |
| 45 | (Sign.sg -> simpset -> term -> thm option) -> simproc | |
| 46 | val add_prems: thm list -> simpset -> simpset | |
| 47 | val prems_of_ss: simpset -> thm list | |
| 48 | val addsimps: simpset * thm list -> simpset | |
| 49 | val delsimps: simpset * thm list -> simpset | |
| 50 | val addeqcongs: simpset * thm list -> simpset | |
| 51 | val deleqcongs: simpset * thm list -> simpset | |
| 52 | val addcongs: simpset * thm list -> simpset | |
| 53 | val delcongs: simpset * thm list -> simpset | |
| 54 | val addsimprocs: simpset * simproc list -> simpset | |
| 55 | val delsimprocs: simpset * simproc list -> simpset | |
| 56 | val setmksimps: simpset * (thm -> thm list) -> simpset | |
| 57 | val setmkcong: simpset * (thm -> thm) -> simpset | |
| 58 | val setmksym: simpset * (thm -> thm option) -> simpset | |
| 59 | val setmkeqTrue: simpset * (thm -> thm option) -> simpset | |
| 60 | val settermless: simpset * (term * term -> bool) -> simpset | |
| 61 | val setsubgoaler: simpset * (simpset -> int -> tactic) -> simpset | |
| 62 | val setloop: simpset * (int -> tactic) -> simpset | |
| 63 | val addloop: simpset * (string * (int -> tactic)) -> simpset | |
| 64 | val delloop: simpset * string -> simpset | |
| 65 | val setSSolver: simpset * solver -> simpset | |
| 66 | val addSSolver: simpset * solver -> simpset | |
| 67 | val setSolver: simpset * solver -> simpset | |
| 68 | val addSolver: simpset * solver -> simpset | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 69 | val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 70 | end; | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 71 | |
| 10413 | 72 | signature META_SIMPLIFIER = | 
| 73 | sig | |
| 11672 | 74 | include BASIC_META_SIMPLIFIER | 
| 10413 | 75 | exception SIMPLIFIER of string * thm | 
| 15023 | 76 | val clear_ss: simpset -> simpset | 
| 77 | exception SIMPROC_FAIL of string * exn | |
| 78 | val simproc_i: Sign.sg -> string -> term list | |
| 79 | -> (Sign.sg -> simpset -> term -> thm option) -> simproc | |
| 80 | val simproc: Sign.sg -> string -> string list | |
| 81 | -> (Sign.sg -> simpset -> term -> thm option) -> simproc | |
| 11672 | 82 | val rewrite_cterm: bool * bool * bool -> | 
| 15023 | 83 | (simpset -> thm -> thm option) -> simpset -> cterm -> thm | 
| 84 | val rewrite_aux: (simpset -> thm -> thm option) -> bool -> thm list -> cterm -> thm | |
| 85 | val simplify_aux: (simpset -> thm -> thm option) -> bool -> thm list -> thm -> thm | |
| 13196 | 86 | val rewrite_term: Sign.sg -> thm list -> (term -> term option) list -> term -> term | 
| 15023 | 87 | val rewrite_thm: bool * bool * bool -> | 
| 88 | (simpset -> thm -> thm option) -> simpset -> thm -> thm | |
| 89 | val rewrite_goals_rule_aux: (simpset -> thm -> thm option) -> thm list -> thm -> thm | |
| 90 | val rewrite_goal_rule: bool * bool * bool -> | |
| 91 | (simpset -> thm -> thm option) -> simpset -> int -> thm -> thm | |
| 92 | val asm_rewrite_goal_tac: bool * bool * bool -> | |
| 93 | (simpset -> tactic) -> simpset -> int -> tactic | |
| 94 | val simp_thm: bool * bool * bool -> simpset -> thm -> thm | |
| 95 | val simp_cterm: bool * bool * bool -> simpset -> cterm -> thm | |
| 10413 | 96 | end; | 
| 97 | ||
| 15023 | 98 | structure MetaSimplifier: META_SIMPLIFIER = | 
| 10413 | 99 | struct | 
| 100 | ||
| 15023 | 101 | |
| 10413 | 102 | (** diagnostics **) | 
| 103 | ||
| 104 | exception SIMPLIFIER of string * thm; | |
| 105 | ||
| 15023 | 106 | val debug_simp = ref false; | 
| 107 | val trace_simp = ref false; | |
| 11505 
a410fa8acfca
Implemented indentation schema for conditional rewrite trace.
 nipkow parents: 
11504diff
changeset | 108 | val simp_depth = ref 0; | 
| 13828 | 109 | val simp_depth_limit = ref 1000; | 
| 11505 
a410fa8acfca
Implemented indentation schema for conditional rewrite trace.
 nipkow parents: 
11504diff
changeset | 110 | |
| 12603 | 111 | local | 
| 112 | ||
| 113 | fun println a = | |
| 15023 | 114 | tracing (case ! simp_depth of 0 => a | n => enclose "[" "]" (string_of_int n) ^ a); | 
| 11505 
a410fa8acfca
Implemented indentation schema for conditional rewrite trace.
 nipkow parents: 
11504diff
changeset | 115 | |
| 
a410fa8acfca
Implemented indentation schema for conditional rewrite trace.
 nipkow parents: 
11504diff
changeset | 116 | fun prnt warn a = if warn then warning a else println a; | 
| 15023 | 117 | fun prtm warn a sg t = prnt warn (a ^ "\n" ^ Sign.string_of_term sg t); | 
| 12603 | 118 | fun prctm warn a t = prnt warn (a ^ "\n" ^ Display.string_of_cterm t); | 
| 10413 | 119 | |
| 12603 | 120 | in | 
| 10413 | 121 | |
| 15023 | 122 | fun debug warn a = if ! debug_simp then prnt warn a else (); | 
| 123 | fun trace warn a = if ! trace_simp then prnt warn a else (); | |
| 10413 | 124 | |
| 15023 | 125 | fun debug_term warn a sign t = if ! debug_simp then prtm warn a sign t else (); | 
| 126 | fun trace_term warn a sign t = if ! trace_simp then prtm warn a sign t else (); | |
| 127 | fun trace_cterm warn a ct = if ! trace_simp then prctm warn a ct else (); | |
| 128 | fun trace_thm a th = if ! trace_simp then prctm false a (Thm.cprop_of th) else (); | |
| 13569 | 129 | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 130 | fun trace_named_thm a (thm, name) = | 
| 15023 | 131 | if ! trace_simp then | 
| 132 | prctm false (if name = "" then a else a ^ " " ^ quote name ^ ":") (Thm.cprop_of thm) | |
| 133 | else (); | |
| 134 | ||
| 135 | fun warn_thm a = prctm true a o Thm.cprop_of; | |
| 10413 | 136 | |
| 12603 | 137 | end; | 
| 10413 | 138 | |
| 139 | ||
| 140 | ||
| 15023 | 141 | (** datatype simpset **) | 
| 142 | ||
| 143 | (* rewrite rules *) | |
| 10413 | 144 | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 145 | type rrule = {thm: thm, name: string, lhs: term, elhs: cterm, fo: bool, perm: bool};
 | 
| 15023 | 146 | |
| 147 | (*thm: the rewrite rule; | |
| 148 | name: name of theorem from which rewrite rule was extracted; | |
| 149 | lhs: the left-hand side; | |
| 150 | elhs: the etac-contracted lhs; | |
| 151 | fo: use first-order matching; | |
| 152 | perm: the rewrite rule is permutative; | |
| 153 | ||
| 12603 | 154 | Remarks: | 
| 10413 | 155 | - elhs is used for matching, | 
| 15023 | 156 | lhs only for preservation of bound variable names; | 
| 10413 | 157 | - fo is set iff | 
| 158 | either elhs is first-order (no Var is applied), | |
| 15023 | 159 | in which case fo-matching is complete, | 
| 10413 | 160 | or elhs is not a pattern, | 
| 15023 | 161 | in which case there is nothing better to do;*) | 
| 10413 | 162 | |
| 163 | fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) =
 | |
| 15023 | 164 | Drule.eq_thm_prop (thm1, thm2); | 
| 165 | ||
| 166 | ||
| 167 | (* congruences *) | |
| 168 | ||
| 169 | type cong = {thm: thm, lhs: cterm};
 | |
| 10413 | 170 | |
| 12603 | 171 | fun eq_cong ({thm = thm1, ...}: cong, {thm = thm2, ...}: cong) =
 | 
| 15023 | 172 | Drule.eq_thm_prop (thm1, thm2); | 
| 10413 | 173 | |
| 174 | ||
| 15023 | 175 | (* solvers *) | 
| 176 | ||
| 177 | datatype solver = | |
| 178 | Solver of | |
| 179 |    {name: string,
 | |
| 180 | solver: thm list -> int -> tactic, | |
| 181 | id: stamp}; | |
| 182 | ||
| 183 | fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()};
 | |
| 184 | ||
| 15034 | 185 | fun solver_name (Solver {name, ...}) = name;
 | 
| 15023 | 186 | fun solver ths (Solver {solver = tacf, ...}) = tacf ths;
 | 
| 187 | fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2);
 | |
| 188 | val merge_solvers = gen_merge_lists eq_solver; | |
| 189 | ||
| 190 | ||
| 191 | (* simplification sets and procedures *) | |
| 192 | ||
| 193 | (*A simpset contains data required during conversion: | |
| 10413 | 194 | rules: discrimination net of rewrite rules; | 
| 15023 | 195 | prems: current premises; | 
| 196 | bounds: names of bound variables already used | |
| 197 | (for generating new names when rewriting under lambda abstractions); | |
| 198 | depth: depth of conditional rewriting; | |
| 10413 | 199 | congs: association list of congruence rules and | 
| 200 | a list of `weak' congruence constants. | |
| 201 | A congruence is `weak' if it avoids normalization of some argument. | |
| 202 | procs: discrimination net of simplification procedures | |
| 203 | (functions that prove rewrite rules on the fly); | |
| 15023 | 204 | mk_rews: | 
| 205 | mk: turn simplification thms into rewrite rules; | |
| 206 | mk_cong: prepare congruence rules; | |
| 207 | mk_sym: turn == around; | |
| 208 | mk_eq_True: turn P into P == True; | |
| 209 | termless: relation for ordered rewriting;*) | |
| 15011 | 210 | |
| 15023 | 211 | type mk_rews = | 
| 212 |  {mk: thm -> thm list,
 | |
| 213 | mk_cong: thm -> thm, | |
| 214 | mk_sym: thm -> thm option, | |
| 215 | mk_eq_True: thm -> thm option}; | |
| 216 | ||
| 217 | datatype simpset = | |
| 218 | Simpset of | |
| 219 |    {rules: rrule Net.net,
 | |
| 10413 | 220 | prems: thm list, | 
| 15023 | 221 | bounds: string list, | 
| 222 | depth: int} * | |
| 223 |    {congs: (string * cong) list * string list,
 | |
| 224 | procs: proc Net.net, | |
| 225 | mk_rews: mk_rews, | |
| 11504 | 226 | termless: term * term -> bool, | 
| 15011 | 227 | subgoal_tac: simpset -> int -> tactic, | 
| 228 | loop_tacs: (string * (int -> tactic)) list, | |
| 15023 | 229 | solvers: solver list * solver list} | 
| 230 | and proc = | |
| 231 | Proc of | |
| 232 |    {name: string,
 | |
| 233 | lhs: cterm, | |
| 234 | proc: Sign.sg -> simpset -> term -> thm option, | |
| 235 | id: stamp}; | |
| 236 | ||
| 237 | fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = (id1 = id2);
 | |
| 238 | ||
| 239 | fun rep_ss (Simpset args) = args; | |
| 10413 | 240 | |
| 15023 | 241 | fun make_ss1 (rules, prems, bounds, depth) = | 
| 242 |   {rules = rules, prems = prems, bounds = bounds, depth = depth};
 | |
| 243 | ||
| 244 | fun map_ss1 f {rules, prems, bounds, depth} =
 | |
| 245 | make_ss1 (f (rules, prems, bounds, depth)); | |
| 10413 | 246 | |
| 15023 | 247 | fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) = | 
| 248 |   {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless,
 | |
| 249 | subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; | |
| 250 | ||
| 251 | fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} =
 | |
| 252 | make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); | |
| 253 | ||
| 254 | fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); | |
| 10413 | 255 | |
| 15023 | 256 | fun map_simpset f (Simpset ({rules, prems, bounds, depth},
 | 
| 257 |     {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers})) =
 | |
| 258 | make_simpset (f ((rules, prems, bounds, depth), | |
| 259 | (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers))); | |
| 10413 | 260 | |
| 15023 | 261 | fun map_simpset1 f (Simpset (r1, r2)) = Simpset (map_ss1 f r1, r2); | 
| 262 | fun map_simpset2 f (Simpset (r1, r2)) = Simpset (r1, map_ss2 f r2); | |
| 263 | ||
| 264 | ||
| 265 | (* print simpsets *) | |
| 10413 | 266 | |
| 15023 | 267 | fun print_ss ss = | 
| 268 | let | |
| 15034 | 269 | val pretty_thms = map Display.pretty_thm; | 
| 15023 | 270 | |
| 15034 | 271 | fun pretty_cong (name, th) = | 
| 272 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, Display.pretty_thm th]; | |
| 15023 | 273 | fun pretty_proc (name, lhss) = | 
| 274 | Pretty.big_list (name ^ ":") (map Display.pretty_cterm lhss); | |
| 15034 | 275 | |
| 276 |     val Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...}) = ss;
 | |
| 277 | val smps = map (#thm o #2) (Net.dest rules); | |
| 278 |     val cngs = map (fn (name, {thm, ...}) => (name, thm)) (#1 congs);
 | |
| 279 | val prcs = Net.dest procs |> | |
| 280 |       map (fn (_, Proc {name, lhs, id, ...}) => ((name, lhs), id))
 | |
| 281 | |> partition_eq eq_snd | |
| 282 | |> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps)) | |
| 283 | |> Library.sort_wrt #1; | |
| 15023 | 284 | in | 
| 15034 | 285 | [Pretty.big_list "simplification rules:" (pretty_thms smps), | 
| 286 | Pretty.big_list "simplification procedures:" (map pretty_proc prcs), | |
| 287 | Pretty.big_list "congruences:" (map pretty_cong cngs), | |
| 15088 | 288 |       Pretty.strs ("loopers:" :: map (quote o #1) loop_tacs),
 | 
| 289 |       Pretty.strs ("unsafe solvers:" :: map (quote o solver_name) (#1 solvers)),
 | |
| 290 |       Pretty.strs ("safe solvers:" :: map (quote o solver_name) (#2 solvers))]
 | |
| 15023 | 291 | |> Pretty.chunks |> Pretty.writeln | 
| 13828 | 292 | end; | 
| 10413 | 293 | |
| 15023 | 294 | |
| 295 | (* empty simpsets *) | |
| 296 | ||
| 297 | local | |
| 298 | ||
| 299 | fun init_ss mk_rews termless subgoal_tac solvers = | |
| 300 | make_simpset ((Net.empty, [], [], 0), | |
| 301 | (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers)); | |
| 302 | ||
| 303 | val basic_mk_rews: mk_rews = | |
| 304 |  {mk = fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [],
 | |
| 305 | mk_cong = I, | |
| 306 | mk_sym = Some o Drule.symmetric_fun, | |
| 307 | mk_eq_True = K None}; | |
| 308 | ||
| 309 | in | |
| 310 | ||
| 311 | val empty_ss = init_ss basic_mk_rews Term.termless (K (K no_tac)) ([], []); | |
| 312 | ||
| 313 | fun clear_ss (Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...})) =
 | |
| 314 | init_ss mk_rews termless subgoal_tac solvers; | |
| 315 | ||
| 316 | end; | |
| 317 | ||
| 318 | ||
| 319 | (* merge simpsets *) (*NOTE: ignores some fields of 2nd simpset*) | |
| 15011 | 320 | |
| 15023 | 321 | fun merge_ss (ss1, ss2) = | 
| 322 | let | |
| 323 |     val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth},
 | |
| 324 |      {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac,
 | |
| 325 | loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; | |
| 326 |     val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = _},
 | |
| 327 |      {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _,
 | |
| 328 | loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; | |
| 15011 | 329 | |
| 15023 | 330 | val rules' = Net.merge (rules1, rules2, eq_rrule); | 
| 331 | val prems' = gen_merge_lists Drule.eq_thm_prop prems1 prems2; | |
| 332 | val bounds' = merge_lists bounds1 bounds2; | |
| 333 | val congs' = gen_merge_lists (eq_cong o pairself #2) congs1 congs2; | |
| 334 | val weak' = merge_lists weak1 weak2; | |
| 335 | val procs' = Net.merge (procs1, procs2, eq_proc); | |
| 336 | val loop_tacs' = merge_alists loop_tacs1 loop_tacs2; | |
| 337 | val unsafe_solvers' = merge_solvers unsafe_solvers1 unsafe_solvers2; | |
| 338 | val solvers' = merge_solvers solvers1 solvers2; | |
| 339 | in | |
| 340 | make_simpset ((rules', prems', bounds', depth), ((congs', weak'), procs', | |
| 341 | mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) | |
| 342 | end; | |
| 343 | ||
| 344 | ||
| 345 | (* simprocs *) | |
| 346 | ||
| 347 | exception SIMPROC_FAIL of string * exn; | |
| 348 | ||
| 349 | datatype simproc = Simproc of proc list; | |
| 350 | ||
| 351 | fun mk_simproc name lhss proc = | |
| 352 | let val id = stamp () in | |
| 353 | Simproc (lhss |> map (fn lhs => | |
| 354 |       Proc {name = name, lhs = lhs, proc = proc, id = id}))
 | |
| 355 | end; | |
| 356 | ||
| 357 | fun simproc_i sg name = mk_simproc name o map (Thm.cterm_of sg o Logic.varify); | |
| 358 | fun simproc sg name = simproc_i sg name o map (Sign.simple_read_term sg TypeInfer.logicT); | |
| 359 | ||
| 15011 | 360 | |
| 10413 | 361 | |
| 362 | (** simpset operations **) | |
| 363 | ||
| 15023 | 364 | (* bounds and prems *) | 
| 10413 | 365 | |
| 15023 | 366 | fun add_bound b = map_simpset1 (fn (rules, prems, bounds, depth) => | 
| 367 | (rules, prems, b :: bounds, depth)); | |
| 10413 | 368 | |
| 15023 | 369 | fun add_prems ths = map_simpset1 (fn (rules, prems, bounds, depth) => | 
| 370 | (rules, ths @ prems, bounds, depth)); | |
| 371 | ||
| 372 | fun prems_of_ss (Simpset ({prems, ...}, _)) = prems;
 | |
| 10413 | 373 | |
| 374 | ||
| 15023 | 375 | (* addsimps *) | 
| 10413 | 376 | |
| 15023 | 377 | fun mk_rrule2 {thm, name, lhs, elhs, perm} =
 | 
| 378 | let | |
| 379 | val fo = Pattern.first_order (term_of elhs) orelse not (Pattern.pattern (term_of elhs)) | |
| 380 |   in {thm = thm, name = name, lhs = lhs, elhs = elhs, fo = fo, perm = perm} end;
 | |
| 10413 | 381 | |
| 15023 | 382 | fun insert_rrule quiet (ss, rrule as {thm, name, lhs, elhs, perm}) =
 | 
| 383 | (trace_named_thm "Adding rewrite rule" (thm, name); | |
| 384 | ss |> map_simpset1 (fn (rules, prems, bounds, depth) => | |
| 385 | let | |
| 386 |       val rrule2 as {elhs, ...} = mk_rrule2 rrule;
 | |
| 387 | val rules' = Net.insert_term ((term_of elhs, rrule2), rules, eq_rrule); | |
| 388 | in (rules', prems, bounds, depth) end) | |
| 389 | handle Net.INSERT => | |
| 390 | (if quiet then () else warn_thm "Ignoring duplicate rewrite rule:" thm; ss)); | |
| 10413 | 391 | |
| 392 | fun vperm (Var _, Var _) = true | |
| 393 | | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) | |
| 394 | | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) | |
| 395 | | vperm (t, u) = (t = u); | |
| 396 | ||
| 397 | fun var_perm (t, u) = | |
| 398 | vperm (t, u) andalso eq_set (term_varnames t, term_varnames u); | |
| 399 | ||
| 400 | (* FIXME: it seems that the conditions on extra variables are too liberal if | |
| 401 | prems are nonempty: does solving the prems really guarantee instantiation of | |
| 402 | all its Vars? Better: a dynamic check each time a rule is applied. | |
| 403 | *) | |
| 404 | fun rewrite_rule_extra_vars prems elhs erhs = | |
| 405 | not (term_varnames erhs subset foldl add_term_varnames (term_varnames elhs, prems)) | |
| 406 | orelse | |
| 15023 | 407 | not (term_tvars erhs subset (term_tvars elhs union List.concat (map term_tvars prems))); | 
| 10413 | 408 | |
| 15023 | 409 | (*simple test for looping rewrite rules and stupid orientations*) | 
| 10413 | 410 | fun reorient sign prems lhs rhs = | 
| 15023 | 411 | rewrite_rule_extra_vars prems lhs rhs | 
| 412 | orelse | |
| 413 | is_Var (head_of lhs) | |
| 414 | orelse | |
| 415 | exists (apl (lhs, Logic.occs)) (rhs :: prems) | |
| 416 | orelse | |
| 417 | null prems andalso Pattern.matches (Sign.tsig_of sign) (lhs, rhs) | |
| 10413 | 418 | (*the condition "null prems" is necessary because conditional rewrites | 
| 419 | with extra variables in the conditions may terminate although | |
| 15023 | 420 | the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*) | 
| 421 | orelse | |
| 422 | is_Const lhs andalso not (is_Const rhs); | |
| 10413 | 423 | |
| 424 | fun decomp_simp thm = | |
| 15023 | 425 | let | 
| 426 |     val {sign, prop, ...} = Thm.rep_thm thm;
 | |
| 427 | val prems = Logic.strip_imp_prems prop; | |
| 428 | val concl = Drule.strip_imp_concl (Thm.cprop_of thm); | |
| 429 | val (lhs, rhs) = Drule.dest_equals concl handle TERM _ => | |
| 430 |       raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm);
 | |
| 431 | val (_, elhs) = Drule.dest_equals (Thm.cprop_of (Thm.eta_conversion lhs)); | |
| 432 | val elhs = if elhs = lhs then lhs else elhs; (*share identical copies*) | |
| 433 | val erhs = Pattern.eta_contract (term_of rhs); | |
| 434 | val perm = | |
| 435 | var_perm (term_of elhs, erhs) andalso | |
| 436 | not (term_of elhs aconv erhs) andalso | |
| 437 | not (is_Var (term_of elhs)); | |
| 10413 | 438 | in (sign, prems, term_of lhs, elhs, term_of rhs, perm) end; | 
| 439 | ||
| 12783 | 440 | fun decomp_simp' thm = | 
| 12979 
4c76bce4ce39
decomp_simp': use lhs instead of elhs (preserves more bound variable names);
 wenzelm parents: 
12783diff
changeset | 441 | let val (_, _, lhs, _, rhs, _) = decomp_simp thm in | 
| 12783 | 442 |     if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", thm)
 | 
| 12979 
4c76bce4ce39
decomp_simp': use lhs instead of elhs (preserves more bound variable names);
 wenzelm parents: 
12783diff
changeset | 443 | else (lhs, rhs) | 
| 12783 | 444 | end; | 
| 445 | ||
| 15023 | 446 | fun mk_eq_True (Simpset (_, {mk_rews = {mk_eq_True, ...}, ...})) (thm, name) =
 | 
| 447 | (case mk_eq_True thm of | |
| 10413 | 448 | None => [] | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 449 | | Some eq_True => | 
| 15023 | 450 | let val (_, _, lhs, elhs, _, _) = decomp_simp eq_True | 
| 451 |       in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end);
 | |
| 10413 | 452 | |
| 15023 | 453 | (*create the rewrite rule and possibly also the eq_True variant, | 
| 454 | in case there are extra vars on the rhs*) | |
| 455 | fun rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm2) = | |
| 456 |   let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in
 | |
| 457 | if term_varnames rhs subset term_varnames lhs andalso | |
| 458 | term_tvars rhs subset term_tvars lhs then [rrule] | |
| 459 | else mk_eq_True ss (thm2, name) @ [rrule] | |
| 10413 | 460 | end; | 
| 461 | ||
| 15023 | 462 | fun mk_rrule ss (thm, name) = | 
| 463 | let val (_, prems, lhs, elhs, rhs, perm) = decomp_simp thm in | |
| 464 |     if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
 | |
| 465 | else | |
| 466 | (*weak test for loops*) | |
| 467 | if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs) | |
| 468 | then mk_eq_True ss (thm, name) | |
| 469 | else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) | |
| 10413 | 470 | end; | 
| 471 | ||
| 15023 | 472 | fun orient_rrule ss (thm, name) = | 
| 473 | let val (sign, prems, lhs, elhs, rhs, perm) = decomp_simp thm in | |
| 474 |     if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
 | |
| 475 | else if reorient sign prems lhs rhs then | |
| 476 | if reorient sign prems rhs lhs | |
| 477 | then mk_eq_True ss (thm, name) | |
| 478 | else | |
| 479 |         let val Simpset (_, {mk_rews = {mk_sym, ...}, ...}) = ss in
 | |
| 480 | (case mk_sym thm of | |
| 481 | None => [] | |
| 482 | | Some thm' => | |
| 483 | let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm' | |
| 484 | in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end) | |
| 485 | end | |
| 486 | else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) | |
| 10413 | 487 | end; | 
| 488 | ||
| 15023 | 489 | fun extract_rews (Simpset (_, {mk_rews = {mk, ...}, ...}), thms) =
 | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 490 | flat (map (fn thm => map (rpair (Thm.name_of_thm thm)) (mk thm)) thms); | 
| 10413 | 491 | |
| 15023 | 492 | fun orient_comb_simps comb mk_rrule (ss, thms) = | 
| 493 | let | |
| 494 | val rews = extract_rews (ss, thms); | |
| 495 | val rrules = flat (map mk_rrule rews); | |
| 496 | in foldl comb (ss, rrules) end; | |
| 10413 | 497 | |
| 15023 | 498 | fun extract_safe_rrules (ss, thm) = | 
| 499 | flat (map (orient_rrule ss) (extract_rews (ss, [thm]))); | |
| 10413 | 500 | |
| 15023 | 501 | (*add rewrite rules explicitly; do not reorient!*) | 
| 502 | fun ss addsimps thms = | |
| 503 | orient_comb_simps (insert_rrule false) (mk_rrule ss) (ss, thms); | |
| 10413 | 504 | |
| 505 | ||
| 15023 | 506 | (* delsimps *) | 
| 10413 | 507 | |
| 15023 | 508 | fun del_rrule (ss, rrule as {thm, elhs, ...}) =
 | 
| 509 | ss |> map_simpset1 (fn (rules, prems, bounds, depth) => | |
| 510 | (Net.delete_term ((term_of elhs, rrule), rules, eq_rrule), prems, bounds, depth)) | |
| 511 | handle Net.DELETE => (warn_thm "Rewrite rule not in simpset:" thm; ss); | |
| 10413 | 512 | |
| 15023 | 513 | fun ss delsimps thms = | 
| 514 | orient_comb_simps del_rrule (map mk_rrule2 o mk_rrule ss) (ss, thms); | |
| 515 | ||
| 516 | ||
| 517 | (* congs *) | |
| 10413 | 518 | |
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 519 | fun cong_name (Const (a, _)) = Some a | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 520 |   | cong_name (Free (a, _)) = Some ("Free: " ^ a)
 | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 521 | | cong_name _ = None; | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 522 | |
| 15023 | 523 | local | 
| 524 | ||
| 525 | fun is_full_cong_prems [] [] = true | |
| 526 | | is_full_cong_prems [] _ = false | |
| 527 | | is_full_cong_prems (p :: prems) varpairs = | |
| 528 | (case Logic.strip_assums_concl p of | |
| 529 |         Const ("==", _) $ lhs $ rhs =>
 | |
| 530 | let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in | |
| 531 | is_Var x andalso forall is_Bound xs andalso | |
| 532 | null (findrep xs) andalso xs = ys andalso | |
| 533 | (x, y) mem varpairs andalso | |
| 534 | is_full_cong_prems prems (varpairs \ (x, y)) | |
| 535 | end | |
| 536 | | _ => false); | |
| 537 | ||
| 538 | fun is_full_cong thm = | |
| 10413 | 539 | let | 
| 15023 | 540 | val prems = prems_of thm and concl = concl_of thm; | 
| 541 | val (lhs, rhs) = Logic.dest_equals concl; | |
| 542 | val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs; | |
| 10413 | 543 | in | 
| 15023 | 544 | f = g andalso null (findrep (xs @ ys)) andalso length xs = length ys andalso | 
| 545 | is_full_cong_prems prems (xs ~~ ys) | |
| 10413 | 546 | end; | 
| 547 | ||
| 15023 | 548 | fun add_cong (ss, thm) = ss |> | 
| 549 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 550 | let | |
| 551 | val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm)) | |
| 552 |         handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | |
| 553 | (*val lhs = Pattern.eta_contract lhs;*) | |
| 554 | val a = the (cong_name (head_of (term_of lhs))) handle Library.OPTION => | |
| 555 |         raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm);
 | |
| 556 | val (alist, weak) = congs; | |
| 557 |       val alist2 = overwrite_warn (alist, (a, {lhs = lhs, thm = thm}))
 | |
| 558 |         ("Overwriting congruence rule for " ^ quote a);
 | |
| 559 | val weak2 = if is_full_cong thm then weak else a :: weak; | |
| 560 | in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); | |
| 10413 | 561 | |
| 15023 | 562 | fun del_cong (ss, thm) = ss |> | 
| 563 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 564 | let | |
| 565 | val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => | |
| 566 |         raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | |
| 567 | (*val lhs = Pattern.eta_contract lhs;*) | |
| 568 | val a = the (cong_name (head_of lhs)) handle Library.OPTION => | |
| 569 |         raise SIMPLIFIER ("Congruence must start with a constant", thm);
 | |
| 570 | val (alist, _) = congs; | |
| 571 | val alist2 = filter (fn (x, _) => x <> a) alist; | |
| 572 |       val weak2 = alist2 |> mapfilter (fn (a, {thm, ...}) =>
 | |
| 573 | if is_full_cong thm then None else Some a); | |
| 574 | in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); | |
| 10413 | 575 | |
| 15023 | 576 | fun mk_cong (Simpset (_, {mk_rews = {mk_cong = f, ...}, ...})) = f;
 | 
| 577 | ||
| 578 | in | |
| 579 | ||
| 580 | val (op addeqcongs) = foldl add_cong; | |
| 581 | val (op deleqcongs) = foldl del_cong; | |
| 582 | ||
| 583 | fun ss addcongs congs = ss addeqcongs map (mk_cong ss) congs; | |
| 584 | fun ss delcongs congs = ss deleqcongs map (mk_cong ss) congs; | |
| 585 | ||
| 586 | end; | |
| 10413 | 587 | |
| 588 | ||
| 15023 | 589 | (* simprocs *) | 
| 590 | ||
| 591 | local | |
| 10413 | 592 | |
| 15023 | 593 | fun add_proc (ss, proc as Proc {name, lhs, ...}) =
 | 
| 594 |  (trace_cterm false ("Adding simplification procedure " ^ quote name ^ " for") lhs;
 | |
| 595 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 596 | (congs, Net.insert_term ((term_of lhs, proc), procs, eq_proc), | |
| 597 | mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss | |
| 598 | handle Net.INSERT => | |
| 599 |     (warning ("Ignoring duplicate simplification procedure " ^ quote name); ss));
 | |
| 10413 | 600 | |
| 15023 | 601 | fun del_proc (ss, proc as Proc {name, lhs, ...}) =
 | 
| 602 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 603 | (congs, Net.delete_term ((term_of lhs, proc), procs, eq_proc), | |
| 604 | mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss | |
| 605 | handle Net.DELETE => | |
| 606 |     (warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss);
 | |
| 10413 | 607 | |
| 15023 | 608 | in | 
| 10413 | 609 | |
| 15023 | 610 | val (op addsimprocs) = foldl (fn (ss, Simproc procs) => foldl add_proc (ss, procs)); | 
| 611 | val (op delsimprocs) = foldl (fn (ss, Simproc procs) => foldl del_proc (ss, procs)); | |
| 10413 | 612 | |
| 15023 | 613 | end; | 
| 10413 | 614 | |
| 615 | ||
| 616 | (* mk_rews *) | |
| 617 | ||
| 15023 | 618 | local | 
| 619 | ||
| 620 | fun map_mk_rews f = map_simpset2 (fn (congs, procs, {mk, mk_cong, mk_sym, mk_eq_True},
 | |
| 621 | termless, subgoal_tac, loop_tacs, solvers) => | |
| 622 | let val (mk', mk_cong', mk_sym', mk_eq_True') = f (mk, mk_cong, mk_sym, mk_eq_True) in | |
| 623 |     (congs, procs, {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True'},
 | |
| 624 | termless, subgoal_tac, loop_tacs, solvers) | |
| 625 | end); | |
| 626 | ||
| 627 | in | |
| 10413 | 628 | |
| 15023 | 629 | fun ss setmksimps mk = ss |> map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True) => | 
| 630 | (mk, mk_cong, mk_sym, mk_eq_True)); | |
| 631 | ||
| 632 | fun ss setmkcong mk_cong = ss |> map_mk_rews (fn (mk, _, mk_sym, mk_eq_True) => | |
| 633 | (mk, mk_cong, mk_sym, mk_eq_True)); | |
| 10413 | 634 | |
| 15023 | 635 | fun ss setmksym mk_sym = ss |> map_mk_rews (fn (mk, mk_cong, _, mk_eq_True) => | 
| 636 | (mk, mk_cong, mk_sym, mk_eq_True)); | |
| 10413 | 637 | |
| 15023 | 638 | fun ss setmkeqTrue mk_eq_True = ss |> map_mk_rews (fn (mk, mk_cong, mk_sym, _) => | 
| 639 | (mk, mk_cong, mk_sym, mk_eq_True)); | |
| 640 | ||
| 641 | end; | |
| 642 | ||
| 14242 
ec70653a02bf
Added access to the mk_rews field (and friends).
 skalberg parents: 
14040diff
changeset | 643 | |
| 10413 | 644 | (* termless *) | 
| 645 | ||
| 15023 | 646 | fun ss settermless termless = ss |> | 
| 647 | map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) => | |
| 648 | (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 649 | |
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 650 | |
| 15023 | 651 | (* tactics *) | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 652 | |
| 15023 | 653 | fun ss setsubgoaler subgoal_tac = ss |> | 
| 654 | map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) => | |
| 655 | (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 656 | |
| 15023 | 657 | fun ss setloop tac = ss |> | 
| 658 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) => | |
| 659 |    (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers));
 | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 660 | |
| 15023 | 661 | fun ss addloop (name, tac) = ss |> | 
| 662 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 663 | (congs, procs, mk_rews, termless, subgoal_tac, | |
| 664 |       overwrite_warn (loop_tacs, (name, tac)) ("Overwriting looper " ^ quote name),
 | |
| 665 | solvers)); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 666 | |
| 15023 | 667 | fun ss delloop name = ss |> | 
| 668 | map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => | |
| 15034 | 669 | let val loop_tacs' = filter_out (equal name o #1) loop_tacs in | 
| 670 | if length loop_tacs <> length loop_tacs' then () | |
| 671 |       else warning ("No such looper in simpset: " ^ quote name);
 | |
| 672 | (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs', solvers) | |
| 15023 | 673 | end); | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 674 | |
| 15023 | 675 | fun ss setSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, | 
| 676 | subgoal_tac, loop_tacs, (unsafe_solvers, _)) => | |
| 677 | (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver]))); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 678 | |
| 15023 | 679 | fun ss addSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, | 
| 680 | subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, | |
| 681 | subgoal_tac, loop_tacs, (unsafe_solvers, merge_solvers solvers [solver]))); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 682 | |
| 15023 | 683 | fun ss setSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, | 
| 684 | subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless, | |
| 685 | subgoal_tac, loop_tacs, ([solver], solvers))); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 686 | |
| 15023 | 687 | fun ss addSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, | 
| 688 | subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, | |
| 689 | subgoal_tac, loop_tacs, (merge_solvers unsafe_solvers [solver], solvers))); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 690 | |
| 15023 | 691 | fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, termless, | 
| 692 | subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, termless, | |
| 693 | subgoal_tac, loop_tacs, (solvers, solvers))); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 694 | |
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 695 | |
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 696 | |
| 10413 | 697 | (** rewriting **) | 
| 698 | ||
| 699 | (* | |
| 700 | Uses conversions, see: | |
| 701 | L C Paulson, A higher-order implementation of rewriting, | |
| 702 | Science of Computer Programming 3 (1983), pages 119-149. | |
| 703 | *) | |
| 704 | ||
| 15023 | 705 | val dest_eq = Drule.dest_equals o Thm.cprop_of; | 
| 706 | val lhs_of = #1 o dest_eq; | |
| 707 | val rhs_of = #2 o dest_eq; | |
| 10413 | 708 | |
| 709 | fun check_conv msg thm thm' = | |
| 710 | let | |
| 711 | val thm'' = transitive thm (transitive | |
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 712 | (symmetric (Drule.beta_eta_conversion (lhs_of thm'))) thm') | 
| 15023 | 713 | in if msg then trace_thm "SUCCEEDED" thm' else (); Some thm'' end | 
| 10413 | 714 | handle THM _ => | 
| 15023 | 715 |     let val {sign, prop = _ $ _ $ prop0, ...} = Thm.rep_thm thm in
 | 
| 716 | trace_thm "Proved wrong thm (Check subgoaler?)" thm'; | |
| 717 | trace_term false "Should have proved:" sign prop0; | |
| 718 | None | |
| 10413 | 719 | end; | 
| 720 | ||
| 721 | ||
| 722 | (* mk_procrule *) | |
| 723 | ||
| 724 | fun mk_procrule thm = | |
| 15023 | 725 | let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in | 
| 726 | if rewrite_rule_extra_vars prems lhs rhs | |
| 727 | then (warn_thm "Extra vars on rhs:" thm; []) | |
| 728 |     else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}]
 | |
| 10413 | 729 | end; | 
| 730 | ||
| 731 | ||
| 15023 | 732 | (* rewritec: conversion to apply the meta simpset to a term *) | 
| 10413 | 733 | |
| 15023 | 734 | (*Since the rewriting strategy is bottom-up, we avoid re-normalizing already | 
| 735 | normalized terms by carrying around the rhs of the rewrite rule just | |
| 736 | applied. This is called the `skeleton'. It is decomposed in parallel | |
| 737 | with the term. Once a Var is encountered, the corresponding term is | |
| 738 | already in normal form. | |
| 739 | skel0 is a dummy skeleton that is to enforce complete normalization.*) | |
| 740 | ||
| 10413 | 741 | val skel0 = Bound 0; | 
| 742 | ||
| 15023 | 743 | (*Use rhs as skeleton only if the lhs does not contain unnormalized bits. | 
| 744 | The latter may happen iff there are weak congruence rules for constants | |
| 745 | in the lhs.*) | |
| 10413 | 746 | |
| 15023 | 747 | fun uncond_skel ((_, weak), (lhs, rhs)) = | 
| 748 | if null weak then rhs (*optimization*) | |
| 749 | else if exists_Const (fn (c, _) => c mem weak) lhs then skel0 | |
| 750 | else rhs; | |
| 751 | ||
| 752 | (*Behaves like unconditional rule if rhs does not contain vars not in the lhs. | |
| 753 | Otherwise those vars may become instantiated with unnormalized terms | |
| 754 | while the premises are solved.*) | |
| 755 | ||
| 756 | fun cond_skel (args as (congs, (lhs, rhs))) = | |
| 757 | if term_varnames rhs subset term_varnames lhs then uncond_skel args | |
| 10413 | 758 | else skel0; | 
| 759 | ||
| 15023 | 760 | fun incr_depth ss = | 
| 761 | let | |
| 762 | val ss' = ss |> map_simpset1 (fn (rules, prems, bounds, depth) => | |
| 763 | (rules, prems, bounds, depth + 1)); | |
| 764 |     val Simpset ({depth = depth', ...}, _) = ss';
 | |
| 765 | in | |
| 766 | if depth' > ! simp_depth_limit | |
| 767 | then (warning "simp_depth_limit exceeded - giving up"; None) | |
| 768 | else | |
| 769 | (if depth' mod 10 = 0 | |
| 770 |       then warning ("Simplification depth " ^ string_of_int depth')
 | |
| 771 | else (); | |
| 772 | Some ss') | |
| 773 | end; | |
| 774 | ||
| 10413 | 775 | (* | 
| 15023 | 776 | Rewriting -- we try in order: | 
| 10413 | 777 | (1) beta reduction | 
| 778 | (2) unconditional rewrite rules | |
| 779 | (3) conditional rewrite rules | |
| 780 | (4) simplification procedures | |
| 781 | ||
| 782 | IMPORTANT: rewrite rules must not introduce new Vars or TVars! | |
| 783 | *) | |
| 784 | ||
| 15023 | 785 | fun rewritec (prover, signt, maxt) ss t = | 
| 10413 | 786 | let | 
| 15023 | 787 |     val Simpset ({rules, ...}, {congs, procs, termless, ...}) = ss;
 | 
| 10413 | 788 | val eta_thm = Thm.eta_conversion t; | 
| 789 | val eta_t' = rhs_of eta_thm; | |
| 790 | val eta_t = term_of eta_t'; | |
| 791 | val tsigt = Sign.tsig_of signt; | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 792 |     fun rew {thm, name, lhs, elhs, fo, perm} =
 | 
| 10413 | 793 | let | 
| 794 |         val {sign, prop, maxidx, ...} = rep_thm thm;
 | |
| 795 | val _ = if Sign.subsig (sign, signt) then () | |
| 15023 | 796 | else (warn_thm "Ignoring rewrite rule from different theory:" thm; | 
| 10413 | 797 | raise Pattern.MATCH); | 
| 798 | val (rthm, elhs') = if maxt = ~1 then (thm, elhs) | |
| 799 | else (Thm.incr_indexes (maxt+1) thm, Thm.cterm_incr_indexes (maxt+1) elhs); | |
| 800 | val insts = if fo then Thm.cterm_first_order_match (elhs', eta_t') | |
| 801 | else Thm.cterm_match (elhs', eta_t'); | |
| 802 | val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm); | |
| 14643 | 803 | val prop' = Thm.prop_of thm'; | 
| 10413 | 804 | val unconditional = (Logic.count_prems (prop',0) = 0); | 
| 805 | val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop') | |
| 806 | in | |
| 11295 | 807 | if perm andalso not (termless (rhs', lhs')) | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 808 | then (trace_named_thm "Cannot apply permutative rewrite rule" (thm, name); | 
| 13569 | 809 | trace_thm "Term does not become smaller:" thm'; None) | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 810 | else (trace_named_thm "Applying instance of rewrite rule" (thm, name); | 
| 10413 | 811 | if unconditional | 
| 812 | then | |
| 13569 | 813 | (trace_thm "Rewriting:" thm'; | 
| 10413 | 814 | let val lr = Logic.dest_equals prop; | 
| 815 | val Some thm'' = check_conv false eta_thm thm' | |
| 816 | in Some (thm'', uncond_skel (congs, lr)) end) | |
| 817 | else | |
| 13569 | 818 | (trace_thm "Trying to rewrite:" thm'; | 
| 15023 | 819 | case incr_depth ss of | 
| 13828 | 820 | None => (trace_thm "FAILED - reached depth limit" thm'; None) | 
| 15023 | 821 | | Some ss' => | 
| 822 | (case prover ss' thm' of | |
| 823 | None => (trace_thm "FAILED" thm'; None) | |
| 10413 | 824 | | Some thm2 => | 
| 825 | (case check_conv true eta_thm thm2 of | |
| 826 | None => None | | |
| 827 | Some thm2' => | |
| 828 | let val concl = Logic.strip_imp_concl prop | |
| 829 | val lr = Logic.dest_equals concl | |
| 13828 | 830 | in Some (thm2', cond_skel (congs, lr)) end)))) | 
| 10413 | 831 | end | 
| 832 | ||
| 833 | fun rews [] = None | |
| 834 | | rews (rrule :: rrules) = | |
| 835 | let val opt = rew rrule handle Pattern.MATCH => None | |
| 836 | in case opt of None => rews rrules | some => some end; | |
| 837 | ||
| 838 | fun sort_rrules rrs = let | |
| 14643 | 839 |       fun is_simple({thm, ...}:rrule) = case Thm.prop_of thm of
 | 
| 10413 | 840 |                                       Const("==",_) $ _ $ _ => true
 | 
| 12603 | 841 | | _ => false | 
| 10413 | 842 | fun sort [] (re1,re2) = re1 @ re2 | 
| 12603 | 843 | | sort (rr::rrs) (re1,re2) = if is_simple rr | 
| 10413 | 844 | then sort rrs (rr::re1,re2) | 
| 845 | else sort rrs (re1,rr::re2) | |
| 846 | in sort rrs ([],[]) end | |
| 847 | ||
| 15023 | 848 | fun proc_rews [] = None | 
| 849 |       | proc_rews (Proc {name, proc, lhs, ...} :: ps) =
 | |
| 850 | if Pattern.matches tsigt (Thm.term_of lhs, Thm.term_of t) then | |
| 10413 | 851 |             (debug_term false ("Trying procedure " ^ quote name ^ " on:") signt eta_t;
 | 
| 13486 
54464ea94d6f
exception SIMPROC_FAIL: solid error reporting of simprocs;
 wenzelm parents: 
13458diff
changeset | 852 | case transform_failure (curry SIMPROC_FAIL name) | 
| 15023 | 853 | (fn () => proc signt ss eta_t) () of | 
| 13486 
54464ea94d6f
exception SIMPROC_FAIL: solid error reporting of simprocs;
 wenzelm parents: 
13458diff
changeset | 854 | None => (debug false "FAILED"; proc_rews ps) | 
| 
54464ea94d6f
exception SIMPROC_FAIL: solid error reporting of simprocs;
 wenzelm parents: 
13458diff
changeset | 855 | | Some raw_thm => | 
| 13569 | 856 |                  (trace_thm ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm;
 | 
| 10413 | 857 | (case rews (mk_procrule raw_thm) of | 
| 13486 
54464ea94d6f
exception SIMPROC_FAIL: solid error reporting of simprocs;
 wenzelm parents: 
13458diff
changeset | 858 |                     None => (trace_cterm true ("IGNORED result of simproc " ^ quote name ^
 | 
| 
54464ea94d6f
exception SIMPROC_FAIL: solid error reporting of simprocs;
 wenzelm parents: 
13458diff
changeset | 859 | " -- does not match") t; proc_rews ps) | 
| 10413 | 860 | | some => some))) | 
| 861 | else proc_rews ps; | |
| 862 | in case eta_t of | |
| 863 | Abs _ $ _ => Some (transitive eta_thm | |
| 12155 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 864 | (beta_conversion false eta_t'), skel0) | 
| 10413 | 865 | | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of | 
| 866 | None => proc_rews (Net.match_term procs eta_t) | |
| 867 | | some => some) | |
| 868 | end; | |
| 869 | ||
| 870 | ||
| 871 | (* conversion to apply a congruence rule to a term *) | |
| 872 | ||
| 873 | fun congc (prover,signt,maxt) {thm=cong,lhs=lhs} t =
 | |
| 14643 | 874 | let val sign = Thm.sign_of_thm cong | 
| 10413 | 875 | val _ = if Sign.subsig (sign, signt) then () | 
| 876 |                  else error("Congruence rule from different theory")
 | |
| 877 | val rthm = if maxt = ~1 then cong else Thm.incr_indexes (maxt+1) cong; | |
| 878 | val rlhs = fst (Drule.dest_equals (Drule.strip_imp_concl (cprop_of rthm))); | |
| 879 | val insts = Thm.cterm_match (rlhs, t) | |
| 880 | (* Pattern.match can raise Pattern.MATCH; | |
| 881 | is handled when congc is called *) | |
| 882 | val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm); | |
| 13569 | 883 | val unit = trace_thm "Applying congruence rule:" thm'; | 
| 13932 
0eb3d91b519a
Simplifier no longer aborts on failed congruence proof.
 ballarin parents: 
13835diff
changeset | 884 | fun err (msg, thm) = (trace_thm msg thm; None) | 
| 10413 | 885 | in case prover thm' of | 
| 13932 
0eb3d91b519a
Simplifier no longer aborts on failed congruence proof.
 ballarin parents: 
13835diff
changeset | 886 |        None => err ("Congruence proof failed.  Could not prove", thm')
 | 
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 887 | | Some thm2 => (case check_conv true (Drule.beta_eta_conversion t) thm2 of | 
| 13932 
0eb3d91b519a
Simplifier no longer aborts on failed congruence proof.
 ballarin parents: 
13835diff
changeset | 888 |           None => err ("Congruence proof failed.  Should not have proved", thm2)
 | 
| 12155 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 889 | | Some thm2' => | 
| 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 890 | if op aconv (pairself term_of (dest_equals (cprop_of thm2'))) | 
| 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 891 | then None else Some thm2') | 
| 10413 | 892 | end; | 
| 893 | ||
| 894 | val (cA, (cB, cC)) = | |
| 895 | apsnd dest_equals (dest_implies (hd (cprems_of Drule.imp_cong))); | |
| 896 | ||
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 897 | fun transitive1 None None = None | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 898 | | transitive1 (Some thm1) None = Some thm1 | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 899 | | transitive1 None (Some thm2) = Some thm2 | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 900 | | transitive1 (Some thm1) (Some thm2) = Some (transitive thm1 thm2) | 
| 10413 | 901 | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 902 | fun transitive2 thm = transitive1 (Some thm); | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 903 | fun transitive3 thm = transitive1 thm o Some; | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 904 | |
| 15023 | 905 | fun bottomc ((simprem, useprem, mutsimp), prover, sign, maxidx) = | 
| 10413 | 906 | let | 
| 15023 | 907 | fun botc skel ss t = | 
| 10413 | 908 | if is_Var skel then None | 
| 909 | else | |
| 15023 | 910 | (case subc skel ss t of | 
| 10413 | 911 | some as Some thm1 => | 
| 15023 | 912 | (case rewritec (prover, sign, maxidx) ss (rhs_of thm1) of | 
| 10413 | 913 | Some (thm2, skel2) => | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 914 | transitive2 (transitive thm1 thm2) | 
| 15023 | 915 | (botc skel2 ss (rhs_of thm2)) | 
| 10413 | 916 | | None => some) | 
| 917 | | None => | |
| 15023 | 918 | (case rewritec (prover, sign, maxidx) ss t of | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 919 | Some (thm2, skel2) => transitive2 thm2 | 
| 15023 | 920 | (botc skel2 ss (rhs_of thm2)) | 
| 10413 | 921 | | None => None)) | 
| 922 | ||
| 15023 | 923 | and try_botc ss t = | 
| 924 | (case botc skel0 ss t of | |
| 10413 | 925 | Some trec1 => trec1 | None => (reflexive t)) | 
| 926 | ||
| 15023 | 927 |     and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 =
 | 
| 10413 | 928 | (case term_of t0 of | 
| 929 | Abs (a, T, t) => | |
| 15023 | 930 | let | 
| 931 | val b = variant bounds a; | |
| 932 |                  val (v, t') = Thm.dest_abs (Some ("." ^ b)) t0;
 | |
| 933 | val ss' = add_bound b ss; | |
| 934 | val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0; | |
| 935 | in case botc skel' ss' t' of | |
| 10413 | 936 | Some thm => Some (abstract_rule a v thm) | 
| 937 | | None => None | |
| 938 | end | |
| 939 | | t $ _ => (case t of | |
| 15023 | 940 |              Const ("==>", _) $ _  => impc t0 ss
 | 
| 10413 | 941 | | Abs _ => | 
| 942 | let val thm = beta_conversion false t0 | |
| 15023 | 943 | in case subc skel0 ss (rhs_of thm) of | 
| 10413 | 944 | None => Some thm | 
| 945 | | Some thm' => Some (transitive thm thm') | |
| 946 | end | |
| 947 | | _ => | |
| 948 | let fun appc () = | |
| 949 | let | |
| 950 | val (tskel, uskel) = case skel of | |
| 951 | tskel $ uskel => (tskel, uskel) | |
| 952 | | _ => (skel0, skel0); | |
| 10767 
8fa4aafa7314
Thm: dest_comb, dest_abs, capply, cabs no longer global;
 wenzelm parents: 
10413diff
changeset | 953 | val (ct, cu) = Thm.dest_comb t0 | 
| 10413 | 954 | in | 
| 15023 | 955 | (case botc tskel ss ct of | 
| 10413 | 956 | Some thm1 => | 
| 15023 | 957 | (case botc uskel ss cu of | 
| 10413 | 958 | Some thm2 => Some (combination thm1 thm2) | 
| 959 | | None => Some (combination thm1 (reflexive cu))) | |
| 960 | | None => | |
| 15023 | 961 | (case botc uskel ss cu of | 
| 10413 | 962 | Some thm1 => Some (combination (reflexive ct) thm1) | 
| 963 | | None => None)) | |
| 964 | end | |
| 965 | val (h, ts) = strip_comb t | |
| 13835 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 966 | in case cong_name h of | 
| 
12b2ffbe543a
Change to meta simplifier: congruence rules may now have frees as head of term.
 ballarin parents: 
13828diff
changeset | 967 | Some a => | 
| 10413 | 968 | (case assoc_string (fst congs, a) of | 
| 969 | None => appc () | |
| 970 | | Some cong => | |
| 15023 | 971 | (*post processing: some partial applications h t1 ... tj, j <= length ts, | 
| 972 | may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*) | |
| 10413 | 973 | (let | 
| 15023 | 974 | val thm = congc (prover ss, sign, maxidx) cong t0; | 
| 12155 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 975 | val t = if_none (apsome rhs_of thm) t0; | 
| 10767 
8fa4aafa7314
Thm: dest_comb, dest_abs, capply, cabs no longer global;
 wenzelm parents: 
10413diff
changeset | 976 | val (cl, cr) = Thm.dest_comb t | 
| 10413 | 977 |                              val dVar = Var(("", 0), dummyT)
 | 
| 978 | val skel = | |
| 979 | list_comb (h, replicate (length ts) dVar) | |
| 15023 | 980 | in case botc skel ss cl of | 
| 12155 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 981 | None => thm | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 982 | | Some thm' => transitive3 thm | 
| 12155 
13c5469b4bb3
congc now returns None if congruence rule has no effect.
 berghofe parents: 
11886diff
changeset | 983 | (combination thm' (reflexive cr)) | 
| 10413 | 984 | end handle TERM _ => error "congc result" | 
| 985 | | Pattern.MATCH => appc ())) | |
| 986 | | _ => appc () | |
| 987 | end) | |
| 988 | | _ => None) | |
| 989 | ||
| 15023 | 990 | and impc ct ss = | 
| 991 | if mutsimp then mut_impc0 [] ct [] [] ss else nonmut_impc ct ss | |
| 10413 | 992 | |
| 15023 | 993 | and rules_of_prem ss prem = | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 994 | if maxidx_of_term (term_of prem) <> ~1 | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 995 | then (trace_cterm true | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 996 | "Cannot add premise as rewrite rule because it contains (type) unknowns:" prem; ([], None)) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 997 | else | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 998 | let val asm = assume prem | 
| 15023 | 999 | in (extract_safe_rrules (ss, asm), Some asm) end | 
| 10413 | 1000 | |
| 15023 | 1001 | and add_rrules (rrss, asms) ss = | 
| 1002 | foldl (insert_rrule true) (ss, flat rrss) |> add_prems (mapfilter I asms) | |
| 10413 | 1003 | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1004 | and disch r (prem, eq) = | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1005 | let | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1006 | val (lhs, rhs) = dest_eq eq; | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1007 | val eq' = implies_elim (Thm.instantiate | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1008 | ([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1009 | (implies_intr prem eq) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1010 | in if not r then eq' else | 
| 10413 | 1011 | let | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1012 | val (prem', concl) = dest_implies lhs; | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1013 | val (prem'', _) = dest_implies rhs | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1014 | in transitive (transitive | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1015 | (Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)]) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1016 | Drule.swap_prems_eq) eq') | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1017 | (Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)]) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1018 | Drule.swap_prems_eq) | 
| 10413 | 1019 | end | 
| 1020 | end | |
| 1021 | ||
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1022 | and rebuild [] _ _ _ _ eq = eq | 
| 15023 | 1023 | | rebuild (prem :: prems) concl (rrs :: rrss) (asm :: asms) ss eq = | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1024 | let | 
| 15023 | 1025 | val ss' = add_rrules (rev rrss, rev asms) ss; | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1026 | val concl' = | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1027 | Drule.mk_implies (prem, if_none (apsome rhs_of eq) concl); | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1028 | val dprem = apsome (curry (disch false) prem) | 
| 15023 | 1029 | in case rewritec (prover, sign, maxidx) ss' concl' of | 
| 1030 | None => rebuild prems concl' rrss asms ss (dprem eq) | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1031 | | Some (eq', _) => transitive2 (foldl (disch false o swap) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1032 | (the (transitive3 (dprem eq) eq'), prems)) | 
| 15023 | 1033 | (mut_impc0 (rev prems) (rhs_of eq') (rev rrss) (rev asms) ss) | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1034 | end | 
| 15023 | 1035 | |
| 1036 | and mut_impc0 prems concl rrss asms ss = | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1037 | let | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1038 | val prems' = strip_imp_prems concl; | 
| 15023 | 1039 | val (rrss', asms') = split_list (map (rules_of_prem ss) prems') | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1040 | in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') | 
| 15023 | 1041 | (asms @ asms') [] [] [] [] ss ~1 ~1 | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1042 | end | 
| 15023 | 1043 | |
| 1044 | and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k = | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1045 | transitive1 (foldl (fn (eq2, (eq1, prem)) => transitive1 eq1 | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1046 | (apsome (curry (disch false) prem) eq2)) (None, eqns ~~ prems')) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1047 | (if changed > 0 then | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1048 | mut_impc (rev prems') concl (rev rrss') (rev asms') | 
| 15023 | 1049 | [] [] [] [] ss ~1 changed | 
| 1050 | else rebuild prems' concl rrss' asms' ss | |
| 1051 | (botc skel0 (add_rrules (rev rrss', rev asms') ss) concl)) | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1052 | |
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1053 | | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) | 
| 15023 | 1054 | prems' rrss' asms' eqns ss changed k = | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1055 | case (if k = 0 then None else botc skel0 (add_rrules | 
| 15023 | 1056 | (rev rrss' @ rrss, rev asms' @ asms) ss) prem) of | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1057 | None => mut_impc prems concl rrss asms (prem :: prems') | 
| 15023 | 1058 | (rrs :: rrss') (asm :: asms') (None :: eqns) ss changed | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1059 | (if k = 0 then 0 else k - 1) | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1060 | | Some eqn => | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1061 | let | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1062 | val prem' = rhs_of eqn; | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1063 | val tprems = map term_of prems; | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1064 | val i = 1 + foldl Int.max (~1, map (fn p => | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1065 | find_index_eq p tprems) (#hyps (rep_thm eqn))); | 
| 15023 | 1066 | val (rrs', asm') = rules_of_prem ss prem' | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1067 | in mut_impc prems concl rrss asms (prem' :: prems') | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1068 | (rrs' :: rrss') (asm' :: asms') (Some (foldr (disch true) | 
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 1069 | (take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies | 
| 15023 | 1070 | (drop (i, prems), concl))))) :: eqns) ss (length prems') ~1 | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1071 | end | 
| 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1072 | |
| 15023 | 1073 | (*legacy code - only for backwards compatibility*) | 
| 1074 | and nonmut_impc ct ss = | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1075 | let val (prem, conc) = dest_implies ct; | 
| 15023 | 1076 | val thm1 = if simprem then botc skel0 ss prem else None; | 
| 10413 | 1077 | val prem1 = if_none (apsome rhs_of thm1) prem; | 
| 15023 | 1078 | val ss1 = if not useprem then ss else add_rrules | 
| 1079 | (apsnd single (apfst single (rules_of_prem ss prem1))) ss | |
| 1080 | in (case botc skel0 ss1 conc of | |
| 10413 | 1081 | None => (case thm1 of | 
| 1082 | None => None | |
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 1083 | | Some thm1' => Some (Drule.imp_cong' thm1' (reflexive conc))) | 
| 10413 | 1084 | | Some thm2 => | 
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1085 | let val thm2' = disch false (prem1, thm2) | 
| 10413 | 1086 | in (case thm1 of | 
| 1087 | None => Some thm2' | |
| 13607 
6908230623a3
Completely reimplemented mutual simplification of premises.
 berghofe parents: 
13569diff
changeset | 1088 | | Some thm1' => | 
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 1089 | Some (transitive (Drule.imp_cong' thm1' (reflexive conc)) thm2')) | 
| 10413 | 1090 | end) | 
| 1091 | end | |
| 1092 | ||
| 15023 | 1093 | in try_botc end; | 
| 10413 | 1094 | |
| 1095 | ||
| 15023 | 1096 | (* Meta-rewriting: rewrites t to u and returns the theorem t==u *) | 
| 10413 | 1097 | |
| 1098 | (* | |
| 1099 | Parameters: | |
| 1100 | mode = (simplify A, | |
| 1101 | use A in simplifying B, | |
| 1102 | use prems of B (if B is again a meta-impl.) to simplify A) | |
| 1103 | when simplifying A ==> B | |
| 1104 | prover: how to solve premises in conditional rewrites and congruences | |
| 1105 | *) | |
| 1106 | ||
| 15023 | 1107 | fun rewrite_cterm mode prover ss ct = | 
| 1108 | let | |
| 1109 |     val Simpset ({depth, ...}, _) = ss;
 | |
| 1110 |     val {sign, t, maxidx, ...} = Thm.rep_cterm ct;
 | |
| 1111 | in | |
| 1112 | trace_cterm false "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" ct; | |
| 1113 | simp_depth := depth; | |
| 1114 | bottomc (mode, prover, sign, maxidx) ss ct | |
| 1115 | end handle THM (s, _, thms) => | |
| 10413 | 1116 |     error ("Exception THM was raised in simplifier:\n" ^ s ^ "\n" ^
 | 
| 11886 | 1117 | Pretty.string_of (Display.pretty_thms thms)); | 
| 10413 | 1118 | |
| 11760 | 1119 | (*Rewrite a cterm*) | 
| 11767 | 1120 | fun rewrite_aux _ _ [] = (fn ct => Thm.reflexive ct) | 
| 15023 | 1121 | | rewrite_aux prover full thms = | 
| 1122 | rewrite_cterm (full, false, false) prover (empty_ss addsimps thms); | |
| 11672 | 1123 | |
| 10413 | 1124 | (*Rewrite a theorem*) | 
| 11767 | 1125 | fun simplify_aux _ _ [] = (fn th => th) | 
| 1126 | | simplify_aux prover full thms = | |
| 15023 | 1127 | Drule.fconv_rule (rewrite_cterm (full, false, false) prover (empty_ss addsimps thms)); | 
| 10413 | 1128 | |
| 15023 | 1129 | (*simple term rewriting -- no proof*) | 
| 1130 | fun rewrite_term sg rules procs = | |
| 1131 | Pattern.rewrite_term (Sign.tsig_of sg) (map decomp_simp' rules) procs; | |
| 1132 | ||
| 1133 | fun rewrite_thm mode prover ss = Drule.fconv_rule (rewrite_cterm mode prover ss); | |
| 10413 | 1134 | |
| 1135 | (*Rewrite the subgoals of a proof state (represented by a theorem) *) | |
| 1136 | fun rewrite_goals_rule_aux _ [] th = th | |
| 1137 | | rewrite_goals_rule_aux prover thms th = | |
| 15001 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 skalberg parents: 
14981diff
changeset | 1138 | Drule.fconv_rule (Drule.goals_conv (K true) (rewrite_cterm (true, true, false) prover | 
| 15023 | 1139 | (empty_ss addsimps thms))) th; | 
| 10413 | 1140 | |
| 15023 | 1141 | (*Rewrite the subgoal of a proof state (represented by a theorem)*) | 
| 15011 | 1142 | fun rewrite_goal_rule mode prover ss i thm = | 
| 10413 | 1143 | if 0 < i andalso i <= nprems_of thm | 
| 15011 | 1144 | then Drule.fconv_rule (Drule.goals_conv (fn j => j=i) (rewrite_cterm mode prover ss)) thm | 
| 10413 | 1145 |   else raise THM("rewrite_goal_rule",i,[thm]);
 | 
| 1146 | ||
| 15023 | 1147 | (*Rewrite subgoal i only. SELECT_GOAL avoids inefficiencies in goals_conv.*) | 
| 1148 | fun asm_rewrite_goal_tac mode prover_tac ss = | |
| 1149 | SELECT_GOAL | |
| 1150 | (PRIMITIVE (rewrite_goal_rule mode (SINGLE o prover_tac) ss 1)); | |
| 12783 | 1151 | |
| 15023 | 1152 | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1153 | |
| 15023 | 1154 | (** simplification tactics and rules **) | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1155 | |
| 15023 | 1156 | fun solve_all_tac solvers ss = | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1157 | let | 
| 15023 | 1158 |     val Simpset (_, {subgoal_tac, ...}) = ss;
 | 
| 1159 | val solve_tac = subgoal_tac (set_solvers solvers ss) THEN_ALL_NEW (K no_tac); | |
| 1160 | in DEPTH_SOLVE (solve_tac 1) end; | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1161 | |
| 15023 | 1162 | (*NOTE: may instantiate unknowns that appear also in other subgoals*) | 
| 1163 | fun generic_simp_tac safe mode ss = | |
| 1164 | let | |
| 1165 |     val Simpset ({prems, ...}, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = ss;
 | |
| 1166 | val loop_tac = FIRST' (map #2 loop_tacs); | |
| 1167 | val solve_tac = FIRST' (map (solver prems) (if safe then solvers else unsafe_solvers)); | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1168 | |
| 15023 | 1169 | fun simp_loop_tac i = | 
| 1170 | asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN | |
| 1171 | (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); | |
| 1172 | in simp_loop_tac end; | |
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1173 | |
| 15023 | 1174 | fun simp rew mode ss thm = | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1175 | let | 
| 15023 | 1176 |     val Simpset (_, {solvers = (unsafe_solvers, _), ...}) = ss;
 | 
| 1177 | val tacf = solve_all_tac unsafe_solvers; | |
| 1178 | fun prover s th = apsome #1 (Seq.pull (tacf s th)); | |
| 15011 | 1179 | in rew mode prover ss thm end; | 
| 15006 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1180 | |
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1181 | val simp_thm = simp rewrite_thm; | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1182 | val simp_cterm = simp rewrite_cterm; | 
| 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 skalberg parents: 
15001diff
changeset | 1183 | |
| 10413 | 1184 | end; | 
| 1185 | ||
| 11672 | 1186 | structure BasicMetaSimplifier: BASIC_META_SIMPLIFIER = MetaSimplifier; | 
| 1187 | open BasicMetaSimplifier; |