| author | haftmann | 
| Sun, 05 Jun 2005 14:41:23 +0200 | |
| changeset 16276 | 3a50bf1f04d0 | 
| parent 15794 | 5de27a5fc5ed | 
| child 16428 | d2203a276b07 | 
| permissions | -rw-r--r-- | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 1 | (* Title: HOL/Import/shuffler.ML | 
| 14516 | 2 | ID: $Id$ | 
| 3 | Author: Sebastian Skalberg, TU Muenchen | |
| 4 | ||
| 5 | Package for proving two terms equal by normalizing (hence the | |
| 6 | "shuffler" name). Uses the simplifier for the normalization. | |
| 7 | *) | |
| 8 | ||
| 9 | signature Shuffler = | |
| 10 | sig | |
| 11 | val debug : bool ref | |
| 12 | ||
| 13 | val norm_term : theory -> term -> thm | |
| 14 | val make_equal : theory -> term -> term -> thm option | |
| 15 | val set_prop : theory -> term -> (string * thm) list -> (string * thm) option | |
| 16 | ||
| 17 | val find_potential: theory -> term -> (string * thm) list | |
| 18 | ||
| 19 | val gen_shuffle_tac: theory -> bool -> (string * thm) list -> int -> tactic | |
| 20 | ||
| 21 | val shuffle_tac: (string * thm) list -> int -> tactic | |
| 22 | val search_tac : (string * thm) list -> int -> tactic | |
| 23 | ||
| 24 | val print_shuffles: theory -> unit | |
| 25 | ||
| 26 | val add_shuffle_rule: thm -> theory -> theory | |
| 27 | val shuffle_attr: theory attribute | |
| 28 | ||
| 29 | val setup : (theory -> theory) list | |
| 30 | end | |
| 31 | ||
| 32 | structure Shuffler :> Shuffler = | |
| 33 | struct | |
| 34 | ||
| 35 | val debug = ref false | |
| 36 | ||
| 37 | fun if_debug f x = if !debug then f x else () | |
| 38 | val message = if_debug writeln | |
| 39 | ||
| 40 | (*Prints exceptions readably to users*) | |
| 41 | fun print_sign_exn_unit sign e = | |
| 42 | case e of | |
| 43 | THM (msg,i,thms) => | |
| 44 | 	 (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
 | |
| 15570 | 45 | List.app print_thm thms) | 
| 14516 | 46 | | THEORY (msg,thys) => | 
| 47 | 	 (writeln ("Exception THEORY raised:\n" ^ msg);
 | |
| 15570 | 48 | List.app (Pretty.writeln o Display.pretty_theory) thys) | 
| 14516 | 49 | | TERM (msg,ts) => | 
| 50 | 	 (writeln ("Exception TERM raised:\n" ^ msg);
 | |
| 15570 | 51 | List.app (writeln o Sign.string_of_term sign) ts) | 
| 14516 | 52 | | TYPE (msg,Ts,ts) => | 
| 53 | 	 (writeln ("Exception TYPE raised:\n" ^ msg);
 | |
| 15570 | 54 | List.app (writeln o Sign.string_of_typ sign) Ts; | 
| 55 | List.app (writeln o Sign.string_of_term sign) ts) | |
| 14516 | 56 | | e => raise e | 
| 57 | ||
| 58 | (*Prints an exception, then fails*) | |
| 59 | fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise e) | |
| 60 | ||
| 14848 | 61 | val string_of_thm = Library.setmp print_mode [] string_of_thm; | 
| 62 | val string_of_cterm = Library.setmp print_mode [] string_of_cterm; | |
| 14516 | 63 | |
| 64 | fun mk_meta_eq th = | |
| 65 | (case concl_of th of | |
| 66 | 	 Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection
 | |
| 67 |        | Const("==",_) $ _ $ _ => th
 | |
| 68 |        | _ => raise THM("Not an equality",0,[th]))
 | |
| 69 |     handle _ => raise THM("Couldn't make meta equality",0,[th])
 | |
| 70 | ||
| 71 | fun mk_obj_eq th = | |
| 72 | (case concl_of th of | |
| 73 | 	 Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th
 | |
| 74 |        | Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq
 | |
| 75 |        | _ => raise THM("Not an equality",0,[th]))
 | |
| 76 |     handle _ => raise THM("Couldn't make object equality",0,[th])
 | |
| 77 | ||
| 78 | structure ShuffleDataArgs: THEORY_DATA_ARGS = | |
| 79 | struct | |
| 80 | val name = "HOL/shuffles" | |
| 81 | type T = thm list | |
| 82 | val empty = [] | |
| 83 | val copy = I | |
| 84 | val prep_ext = I | |
| 85 | val merge = Library.gen_union Thm.eq_thm | |
| 86 | fun print sg thms = | |
| 87 | Pretty.writeln (Pretty.big_list "Shuffle theorems:" | |
| 88 | (map Display.pretty_thm thms)) | |
| 89 | end | |
| 90 | ||
| 91 | structure ShuffleData = TheoryDataFun(ShuffleDataArgs) | |
| 92 | ||
| 93 | val weaken = | |
| 94 | let | |
| 95 | val cert = cterm_of (sign_of ProtoPure.thy) | |
| 96 | 	val P = Free("P",propT)
 | |
| 97 | 	val Q = Free("Q",propT)
 | |
| 98 | val PQ = Logic.mk_implies(P,Q) | |
| 99 | val PPQ = Logic.mk_implies(P,PQ) | |
| 100 | val cP = cert P | |
| 101 | val cQ = cert Q | |
| 102 | val cPQ = cert PQ | |
| 103 | val cPPQ = cert PPQ | |
| 104 | val th1 = assume cPQ |> implies_intr_list [cPQ,cP] | |
| 105 | val th3 = assume cP | |
| 106 | val th4 = implies_elim_list (assume cPPQ) [th3,th3] | |
| 107 | |> implies_intr_list [cPPQ,cP] | |
| 108 | in | |
| 109 | equal_intr th4 th1 |> standard | |
| 110 | end | |
| 111 | ||
| 112 | val imp_comm = | |
| 113 | let | |
| 114 | val cert = cterm_of (sign_of ProtoPure.thy) | |
| 115 | 	val P = Free("P",propT)
 | |
| 116 | 	val Q = Free("Q",propT)
 | |
| 117 | 	val R = Free("R",propT)
 | |
| 118 | val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R)) | |
| 119 | val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R)) | |
| 120 | val cP = cert P | |
| 121 | val cQ = cert Q | |
| 122 | val cPQR = cert PQR | |
| 123 | val cQPR = cert QPR | |
| 124 | val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ] | |
| 125 | |> implies_intr_list [cPQR,cQ,cP] | |
| 126 | val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP] | |
| 127 | |> implies_intr_list [cQPR,cP,cQ] | |
| 128 | in | |
| 129 | equal_intr th1 th2 |> standard | |
| 130 | end | |
| 131 | ||
| 132 | val def_norm = | |
| 133 | let | |
| 134 | val cert = cterm_of (sign_of ProtoPure.thy) | |
| 14854 | 135 | 	val aT = TFree("'a",[])
 | 
| 136 | 	val bT = TFree("'b",[])
 | |
| 14516 | 137 | 	val v = Free("v",aT)
 | 
| 138 | 	val P = Free("P",aT-->bT)
 | |
| 139 | 	val Q = Free("Q",aT-->bT)
 | |
| 140 | 	val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0)))
 | |
| 141 | val cPQ = cert (Logic.mk_equals(P,Q)) | |
| 142 | val cv = cert v | |
| 143 | val rew = assume cvPQ | |
| 144 | |> forall_elim cv | |
| 145 | |> abstract_rule "v" cv | |
| 146 | val (lhs,rhs) = Logic.dest_equals(concl_of rew) | |
| 147 | val th1 = transitive (transitive | |
| 148 | (eta_conversion (cert lhs) |> symmetric) | |
| 149 | rew) | |
| 150 | (eta_conversion (cert rhs)) | |
| 151 | |> implies_intr cvPQ | |
| 152 | val th2 = combination (assume cPQ) (reflexive cv) | |
| 153 | |> forall_intr cv | |
| 154 | |> implies_intr cPQ | |
| 155 | in | |
| 156 | equal_intr th1 th2 |> standard | |
| 157 | end | |
| 158 | ||
| 159 | val all_comm = | |
| 160 | let | |
| 161 | val cert = cterm_of (sign_of ProtoPure.thy) | |
| 14854 | 162 | 	val xT = TFree("'a",[])
 | 
| 163 | 	val yT = TFree("'b",[])
 | |
| 14516 | 164 | 	val P = Free("P",xT-->yT-->propT)
 | 
| 165 | 	val lhs = all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))
 | |
| 166 | 	val rhs = all yT $ (Abs("y",yT,all xT $ (Abs("x",xT,P $ Bound 0 $ Bound 1))))
 | |
| 167 | val cl = cert lhs | |
| 168 | val cr = cert rhs | |
| 169 | 	val cx = cert (Free("x",xT))
 | |
| 170 | 	val cy = cert (Free("y",yT))
 | |
| 171 | val th1 = assume cr | |
| 172 | |> forall_elim_list [cy,cx] | |
| 173 | |> forall_intr_list [cx,cy] | |
| 174 | |> implies_intr cr | |
| 175 | val th2 = assume cl | |
| 176 | |> forall_elim_list [cx,cy] | |
| 177 | |> forall_intr_list [cy,cx] | |
| 178 | |> implies_intr cl | |
| 179 | in | |
| 180 | equal_intr th1 th2 |> standard | |
| 181 | end | |
| 182 | ||
| 183 | val equiv_comm = | |
| 184 | let | |
| 185 | val cert = cterm_of (sign_of ProtoPure.thy) | |
| 186 | 	val T    = TFree("'a",[])
 | |
| 187 | 	val t    = Free("t",T)
 | |
| 188 | 	val u    = Free("u",T)
 | |
| 189 | val ctu = cert (Logic.mk_equals(t,u)) | |
| 190 | val cut = cert (Logic.mk_equals(u,t)) | |
| 191 | val th1 = assume ctu |> symmetric |> implies_intr ctu | |
| 192 | val th2 = assume cut |> symmetric |> implies_intr cut | |
| 193 | in | |
| 194 | equal_intr th1 th2 |> standard | |
| 195 | end | |
| 196 | ||
| 197 | (* This simplification procedure rewrites !!x y. P x y | |
| 198 | deterministicly, in order for the normalization function, defined | |
| 199 | below, to handle nested quantifiers robustly *) | |
| 200 | ||
| 201 | local | |
| 202 | ||
| 203 | exception RESULT of int | |
| 204 | ||
| 205 | fun find_bound n (Bound i) = if i = n then raise RESULT 0 | |
| 206 | else if i = n+1 then raise RESULT 1 | |
| 207 | else () | |
| 208 | | find_bound n (t $ u) = (find_bound n t; find_bound n u) | |
| 209 | | find_bound n (Abs(_,_,t)) = find_bound (n+1) t | |
| 210 | | find_bound _ _ = () | |
| 211 | ||
| 212 | fun swap_bound n (Bound i) = if i = n then Bound (n+1) | |
| 213 | else if i = n+1 then Bound n | |
| 214 | else Bound i | |
| 215 | | swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u) | |
| 216 | | swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t) | |
| 217 | | swap_bound n t = t | |
| 218 | ||
| 219 | fun rew_th sg (xv as (x,xT)) (yv as (y,yT)) t = | |
| 220 | let | |
| 221 | val lhs = list_all ([xv,yv],t) | |
| 222 | val rhs = list_all ([yv,xv],swap_bound 0 t) | |
| 223 | val rew = Logic.mk_equals (lhs,rhs) | |
| 224 | val init = trivial (cterm_of sg rew) | |
| 225 | in | |
| 226 | (all_comm RS init handle e => (message "rew_th"; print_exn e)) | |
| 227 | end | |
| 228 | ||
| 229 | fun quant_rewrite sg assumes (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) =
 | |
| 230 | let | |
| 231 | val res = (find_bound 0 body;2) handle RESULT i => i | |
| 232 | in | |
| 233 | case res of | |
| 15531 | 234 | 0 => SOME (rew_th sg (x,xT) (y,yT) body) | 
| 14516 | 235 | | 1 => if string_ord(y,x) = LESS | 
| 236 | then | |
| 237 | let | |
| 238 | 			 val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
 | |
| 239 | val t_th = reflexive (cterm_of sg t) | |
| 240 | val newt_th = reflexive (cterm_of sg newt) | |
| 241 | in | |
| 15531 | 242 | SOME (transitive t_th newt_th) | 
| 14516 | 243 | end | 
| 15531 | 244 | else NONE | 
| 14516 | 245 | | _ => error "norm_term (quant_rewrite) internal error" | 
| 246 | end | |
| 15531 | 247 | | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE) | 
| 14516 | 248 | |
| 249 | fun freeze_thaw_term t = | |
| 250 | let | |
| 251 | val tvars = term_tvars t | |
| 252 | val tfree_names = add_term_tfree_names(t,[]) | |
| 253 | val (type_inst,_) = | |
| 15570 | 254 | Library.foldl (fn ((inst,used),(w as (v,_),S)) => | 
| 14516 | 255 | let | 
| 256 | val v' = variant used v | |
| 257 | in | |
| 258 | ((w,TFree(v',S))::inst,v'::used) | |
| 259 | end) | |
| 260 | (([],tfree_names),tvars) | |
| 261 | val t' = subst_TVars type_inst t | |
| 262 | in | |
| 14518 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 263 | (t',map (fn (w,TFree(v,S)) => (v,TVar(w,S)) | 
| 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 264 | | _ => error "Internal error in Shuffler.freeze_thaw") type_inst) | 
| 14516 | 265 | end | 
| 266 | ||
| 267 | fun inst_tfrees sg [] thm = thm | |
| 268 | | inst_tfrees sg ((name,U)::rest) thm = | |
| 269 | let | |
| 270 | val cU = ctyp_of sg U | |
| 15794 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15661diff
changeset | 271 | val tfrees = add_term_tfrees (prop_of thm,[]) | 
| 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15661diff
changeset | 272 | val (thm',rens) = varifyT' | 
| 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15661diff
changeset | 273 | (gen_rem (op = o apfst fst) (tfrees, name)) thm | 
| 14516 | 274 | val mid = | 
| 275 | case rens of | |
| 276 | [] => thm' | |
| 15794 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15661diff
changeset | 277 | | [((_, S), idx)] => instantiate | 
| 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15661diff
changeset | 278 | ([(ctyp_of sg (TVar (idx, S)), cU)], []) thm' | 
| 14516 | 279 | | _ => error "Shuffler.inst_tfrees internal error" | 
| 280 | in | |
| 281 | inst_tfrees sg rest mid | |
| 282 | end | |
| 283 | ||
| 284 | fun is_Abs (Abs _) = true | |
| 285 | | is_Abs _ = false | |
| 286 | ||
| 287 | fun eta_redex (t $ Bound 0) = | |
| 288 | let | |
| 289 | fun free n (Bound i) = i = n | |
| 290 | | free n (t $ u) = free n t orelse free n u | |
| 291 | | free n (Abs(_,_,t)) = free (n+1) t | |
| 292 | | free n _ = false | |
| 293 | in | |
| 294 | not (free 0 t) | |
| 295 | end | |
| 296 | | eta_redex _ = false | |
| 297 | ||
| 298 | fun eta_contract sg assumes origt = | |
| 299 | let | |
| 300 | val (typet,Tinst) = freeze_thaw_term origt | |
| 301 | val (init,thaw) = freeze_thaw (reflexive (cterm_of sg typet)) | |
| 302 | val final = inst_tfrees sg Tinst o thaw | |
| 303 | val t = #1 (Logic.dest_equals (prop_of init)) | |
| 304 | val _ = | |
| 305 | let | |
| 306 | val lhs = #1 (Logic.dest_equals (prop_of (final init))) | |
| 307 | in | |
| 308 | if not (lhs aconv origt) | |
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15574diff
changeset | 309 | then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; | 
| 14516 | 310 | writeln (string_of_cterm (cterm_of sg origt)); | 
| 311 | writeln (string_of_cterm (cterm_of sg lhs)); | |
| 312 | writeln (string_of_cterm (cterm_of sg typet)); | |
| 313 | writeln (string_of_cterm (cterm_of sg t)); | |
| 314 | app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of sg T)))) Tinst; | |
| 315 | writeln "done") | |
| 316 | else () | |
| 317 | end | |
| 318 | in | |
| 319 | case t of | |
| 320 | 	    Const("all",_) $ (Abs(x,xT,Const("==",eqT) $ P $ Q)) =>
 | |
| 321 | ((if eta_redex P andalso eta_redex Q | |
| 322 | then | |
| 323 | let | |
| 324 | val cert = cterm_of sg | |
| 325 | val v = Free(variant (add_term_free_names(t,[])) "v",xT) | |
| 326 | val cv = cert v | |
| 327 | val ct = cert t | |
| 328 | val th = (assume ct) | |
| 329 | |> forall_elim cv | |
| 330 | |> abstract_rule x cv | |
| 331 | val ext_th = eta_conversion (cert (Abs(x,xT,P))) | |
| 332 | val th' = transitive (symmetric ext_th) th | |
| 333 | val cu = cert (prop_of th') | |
| 334 | val uth = combination (assume cu) (reflexive cv) | |
| 335 | val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v))) | |
| 336 | |> transitive uth | |
| 337 | |> forall_intr cv | |
| 338 | |> implies_intr cu | |
| 339 | val rew_th = equal_intr (th' |> implies_intr ct) uth' | |
| 340 | val res = final rew_th | |
| 341 | val lhs = (#1 (Logic.dest_equals (prop_of res))) | |
| 342 | in | |
| 15531 | 343 | SOME res | 
| 14516 | 344 | end | 
| 15531 | 345 | else NONE) | 
| 14516 | 346 | handle e => (writeln "eta_contract:";print_exn e)) | 
| 15531 | 347 | 	  | _ => (error ("Bad eta_contract argument" ^ (string_of_cterm (cterm_of sg t))); NONE)
 | 
| 14516 | 348 | end | 
| 349 | ||
| 350 | fun beta_fun sg assume t = | |
| 15531 | 351 | SOME (beta_conversion true (cterm_of sg t)) | 
| 14516 | 352 | |
| 353 | fun eta_expand sg assumes origt = | |
| 354 | let | |
| 355 | val (typet,Tinst) = freeze_thaw_term origt | |
| 356 | val (init,thaw) = freeze_thaw (reflexive (cterm_of sg typet)) | |
| 357 | val final = inst_tfrees sg Tinst o thaw | |
| 358 | val t = #1 (Logic.dest_equals (prop_of init)) | |
| 359 | val _ = | |
| 360 | let | |
| 361 | val lhs = #1 (Logic.dest_equals (prop_of (final init))) | |
| 362 | in | |
| 363 | if not (lhs aconv origt) | |
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15574diff
changeset | 364 | then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; | 
| 14516 | 365 | writeln (string_of_cterm (cterm_of sg origt)); | 
| 366 | writeln (string_of_cterm (cterm_of sg lhs)); | |
| 367 | writeln (string_of_cterm (cterm_of sg typet)); | |
| 368 | writeln (string_of_cterm (cterm_of sg t)); | |
| 369 | app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of sg T)))) Tinst; | |
| 370 | writeln "done") | |
| 371 | else () | |
| 372 | end | |
| 373 | in | |
| 374 | case t of | |
| 375 | 	    Const("==",T) $ P $ Q =>
 | |
| 376 | if is_Abs P orelse is_Abs Q | |
| 377 | then (case domain_type T of | |
| 378 | 		      Type("fun",[aT,bT]) =>
 | |
| 379 | let | |
| 380 | val cert = cterm_of sg | |
| 381 | val vname = variant (add_term_free_names(t,[])) "v" | |
| 382 | val v = Free(vname,aT) | |
| 383 | val cv = cert v | |
| 384 | val ct = cert t | |
| 385 | val th1 = (combination (assume ct) (reflexive cv)) | |
| 386 | |> forall_intr cv | |
| 387 | |> implies_intr ct | |
| 388 | val concl = cert (concl_of th1) | |
| 389 | val th2 = (assume concl) | |
| 390 | |> forall_elim cv | |
| 391 | |> abstract_rule vname cv | |
| 392 | val (lhs,rhs) = Logic.dest_equals (prop_of th2) | |
| 393 | val elhs = eta_conversion (cert lhs) | |
| 394 | val erhs = eta_conversion (cert rhs) | |
| 395 | val th2' = transitive | |
| 396 | (transitive (symmetric elhs) th2) | |
| 397 | erhs | |
| 398 | val res = equal_intr th1 (th2' |> implies_intr concl) | |
| 399 | val res' = final res | |
| 400 | in | |
| 15531 | 401 | SOME res' | 
| 14516 | 402 | end | 
| 15531 | 403 | | _ => NONE) | 
| 404 | else NONE | |
| 405 | 	  | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of sg t))); NONE)
 | |
| 14516 | 406 | end | 
| 407 | handle e => (writeln "eta_expand internal error";print_exn e) | |
| 408 | ||
| 14854 | 409 | fun mk_tfree s = TFree("'"^s,[])
 | 
| 14516 | 410 | val xT = mk_tfree "a" | 
| 411 | val yT = mk_tfree "b" | |
| 412 | val P  = Var(("P",0),xT-->yT-->propT)
 | |
| 413 | val Q  = Var(("Q",0),xT-->yT)
 | |
| 414 | val R  = Var(("R",0),xT-->yT)
 | |
| 415 | val S  = Var(("S",0),xT)
 | |
| 416 | in | |
| 417 | fun beta_simproc sg = Simplifier.simproc_i | |
| 418 | sg | |
| 419 | "Beta-contraction" | |
| 420 | 		      [Abs("x",xT,Q) $ S]
 | |
| 421 | beta_fun | |
| 422 | ||
| 423 | fun quant_simproc sg = Simplifier.simproc_i | |
| 424 | sg | |
| 425 | "Ordered rewriting of nested quantifiers" | |
| 426 | 			   [all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))]
 | |
| 427 | quant_rewrite | |
| 428 | fun eta_expand_simproc sg = Simplifier.simproc_i | |
| 429 | sg | |
| 430 | "Smart eta-expansion by equivalences" | |
| 431 | [Logic.mk_equals(Q,R)] | |
| 432 | eta_expand | |
| 433 | fun eta_contract_simproc sg = Simplifier.simproc_i | |
| 434 | sg | |
| 435 | "Smart handling of eta-contractions" | |
| 436 | 			 [all xT $ (Abs("x",xT,Logic.mk_equals(Q $ Bound 0,R $ Bound 0)))]
 | |
| 437 | eta_contract | |
| 438 | end | |
| 439 | ||
| 440 | (* Disambiguates the names of bound variables in a term, returning t | |
| 441 | == t' where all the names of bound variables in t' are unique *) | |
| 442 | ||
| 443 | fun disamb_bound sg t = | |
| 444 | let | |
| 445 | ||
| 446 | fun F (t $ u,idx) = | |
| 447 | let | |
| 448 | val (t',idx') = F (t,idx) | |
| 449 | val (u',idx'') = F (u,idx') | |
| 450 | in | |
| 451 | (t' $ u',idx'') | |
| 452 | end | |
| 453 | | F (Abs(x,xT,t),idx) = | |
| 454 | let | |
| 455 | val x' = "x" ^ (LargeInt.toString idx) (* amazing *) | |
| 456 | val (t',idx') = F (t,idx+1) | |
| 457 | in | |
| 458 | (Abs(x',xT,t'),idx') | |
| 459 | end | |
| 460 | | F arg = arg | |
| 461 | val (t',_) = F (t,0) | |
| 462 | val ct = cterm_of sg t | |
| 463 | val ct' = cterm_of sg t' | |
| 464 | val res = transitive (reflexive ct) (reflexive ct') | |
| 465 | 	val _ = message ("disamb_term: " ^ (string_of_thm res))
 | |
| 466 | in | |
| 467 | res | |
| 468 | end | |
| 469 | ||
| 470 | (* Transforms a term t to some normal form t', returning the theorem t | |
| 471 | == t'. This is originally a help function for make_equal, but might | |
| 472 | be handy in its own right, for example for indexing terms. *) | |
| 473 | ||
| 474 | fun norm_term thy t = | |
| 475 | let | |
| 476 | val sg = sign_of thy | |
| 477 | ||
| 478 | val norms = ShuffleData.get thy | |
| 479 | val ss = empty_ss setmksimps single | |
| 480 | addsimps (map (transfer_sg sg) norms) | |
| 481 | fun chain f th = | |
| 482 | let | |
| 483 | val rhs = snd (dest_equals (cprop_of th)) | |
| 484 | in | |
| 485 | transitive th (f rhs) | |
| 486 | end | |
| 487 | ||
| 488 | val th = | |
| 489 | t |> disamb_bound sg | |
| 490 | |> chain (Simplifier.full_rewrite | |
| 491 | (ss addsimprocs [quant_simproc sg,eta_expand_simproc sg,eta_contract_simproc sg])) | |
| 492 | |> chain eta_conversion | |
| 493 | |> strip_shyps | |
| 494 | 	val _ = message ("norm_term: " ^ (string_of_thm th))
 | |
| 495 | in | |
| 496 | th | |
| 497 | end | |
| 498 | handle e => (writeln "norm_term internal error"; print_sign_exn (sign_of thy) e) | |
| 499 | ||
| 500 | ||
| 501 | (* Closes a theorem with respect to free and schematic variables (does | |
| 502 | not touch type variables, though). *) | |
| 503 | ||
| 504 | fun close_thm th = | |
| 505 | let | |
| 506 | val sg = sign_of_thm th | |
| 507 | val c = prop_of th | |
| 14854 | 508 | val vars = add_term_frees (c,add_term_vars(c,[])) | 
| 14516 | 509 | in | 
| 14854 | 510 | Drule.forall_intr_list (map (cterm_of sg) vars) th | 
| 14516 | 511 | end | 
| 512 | handle e => (writeln "close_thm internal error"; print_exn e) | |
| 513 | ||
| 514 | (* Normalizes a theorem's conclusion using norm_term. *) | |
| 515 | ||
| 516 | fun norm_thm thy th = | |
| 517 | let | |
| 518 | val c = prop_of th | |
| 519 | in | |
| 520 | equal_elim (norm_term thy c) th | |
| 521 | end | |
| 522 | ||
| 523 | (* make_equal sg t u tries to construct the theorem t == u under the | |
| 15531 | 524 | signature sg. If it succeeds, SOME (t == u) is returned, otherwise | 
| 525 | NONE is returned. *) | |
| 14516 | 526 | |
| 527 | fun make_equal sg t u = | |
| 528 | let | |
| 529 | val t_is_t' = norm_term sg t | |
| 530 | val u_is_u' = norm_term sg u | |
| 531 | val th = transitive t_is_t' (symmetric u_is_u') | |
| 15531 | 532 | 	val _ = message ("make_equal: SOME " ^ (string_of_thm th))
 | 
| 14516 | 533 | in | 
| 15531 | 534 | SOME th | 
| 14516 | 535 | end | 
| 15531 | 536 | handle e as THM _ => (message "make_equal: NONE";NONE) | 
| 14516 | 537 | |
| 538 | fun match_consts ignore t (* th *) = | |
| 539 | let | |
| 540 | fun add_consts (Const (c, _), cs) = | |
| 541 | if c mem_string ignore | |
| 542 | then cs | |
| 543 | else c ins_string cs | |
| 544 | | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs)) | |
| 545 | | add_consts (Abs (_, _, t), cs) = add_consts (t, cs) | |
| 546 | | add_consts (_, cs) = cs | |
| 547 | val t_consts = add_consts(t,[]) | |
| 548 | in | |
| 549 | fn (name,th) => | |
| 550 | let | |
| 551 | val th_consts = add_consts(prop_of th,[]) | |
| 552 | in | |
| 553 | eq_set(t_consts,th_consts) | |
| 554 | end | |
| 555 | end | |
| 556 | ||
| 557 | val collect_ignored = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 558 | foldr (fn (thm,cs) => | 
| 14516 | 559 | let | 
| 560 | val (lhs,rhs) = Logic.dest_equals (prop_of thm) | |
| 561 | val ignore_lhs = term_consts lhs \\ term_consts rhs | |
| 562 | val ignore_rhs = term_consts rhs \\ term_consts lhs | |
| 563 | in | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 564 | foldr (op ins_string) cs (ignore_lhs @ ignore_rhs) | 
| 14516 | 565 | end) | 
| 566 | ||
| 567 | (* set_prop t thms tries to make a theorem with the proposition t from | |
| 568 | one of the theorems thms, by shuffling the propositions around. If it | |
| 15531 | 569 | succeeds, SOME theorem is returned, otherwise NONE. *) | 
| 14516 | 570 | |
| 571 | fun set_prop thy t = | |
| 572 | let | |
| 573 | val sg = sign_of thy | |
| 14854 | 574 | val vars = add_term_frees (t,add_term_vars (t,[])) | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 575 | val closed_t = foldr (fn (v,body) => let val vT = type_of v | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 576 | 					     in all vT $ (Abs("x",vT,abstract_over(v,body))) end) t vars
 | 
| 14516 | 577 | val rew_th = norm_term thy closed_t | 
| 578 | val rhs = snd (dest_equals (cprop_of rew_th)) | |
| 579 | ||
| 580 | val shuffles = ShuffleData.get thy | |
| 15531 | 581 | fun process [] = NONE | 
| 14516 | 582 | | process ((name,th)::thms) = | 
| 583 | let | |
| 584 | val norm_th = varifyT (norm_thm thy (close_thm (transfer_sg sg th))) | |
| 585 | val triv_th = trivial rhs | |
| 586 | 		val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
 | |
| 587 | val mod_th = case Seq.pull (bicompose true (false,norm_th,0) 1 triv_th) of | |
| 15531 | 588 | SOME(th,_) => SOME th | 
| 589 | | NONE => NONE | |
| 14516 | 590 | in | 
| 591 | case mod_th of | |
| 15531 | 592 | SOME mod_th => | 
| 14516 | 593 | let | 
| 594 | val closed_th = equal_elim (symmetric rew_th) mod_th | |
| 595 | in | |
| 596 | 			message ("Shuffler.set_prop succeeded by " ^ name);
 | |
| 15531 | 597 | SOME (name,forall_elim_list (map (cterm_of sg) vars) closed_th) | 
| 14516 | 598 | end | 
| 15531 | 599 | | NONE => process thms | 
| 14516 | 600 | end | 
| 601 | handle e as THM _ => process thms | |
| 602 | in | |
| 603 | fn thms => | |
| 604 | case process thms of | |
| 15531 | 605 | res as SOME (name,th) => if (prop_of th) aconv t | 
| 14516 | 606 | then res | 
| 607 | else error "Internal error in set_prop" | |
| 15531 | 608 | | NONE => NONE | 
| 14516 | 609 | end | 
| 610 | handle e => (writeln "set_prop internal error"; print_exn e) | |
| 611 | ||
| 612 | fun find_potential thy t = | |
| 613 | let | |
| 614 | val shuffles = ShuffleData.get thy | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 615 | val ignored = collect_ignored [] shuffles | 
| 14516 | 616 | val rel_consts = term_consts t \\ ignored | 
| 617 | val pot_thms = PureThy.thms_containing_consts thy rel_consts | |
| 618 | in | |
| 15570 | 619 | List.filter (match_consts ignored t) pot_thms | 
| 14516 | 620 | end | 
| 621 | ||
| 622 | fun gen_shuffle_tac thy search thms i st = | |
| 623 | let | |
| 624 | 	val _ = message ("Shuffling " ^ (string_of_thm st))
 | |
| 15570 | 625 | val t = List.nth(prems_of st,i-1) | 
| 14516 | 626 | val set = set_prop thy t | 
| 627 | fun process_tac thms st = | |
| 628 | case set thms of | |
| 15531 | 629 | SOME (_,th) => Seq.of_list (compose (th,i,st)) | 
| 630 | | NONE => Seq.empty | |
| 14516 | 631 | in | 
| 632 | (process_tac thms APPEND (if search | |
| 633 | then process_tac (find_potential thy t) | |
| 634 | else no_tac)) st | |
| 635 | end | |
| 636 | ||
| 637 | fun shuffle_tac thms i st = | |
| 638 | gen_shuffle_tac (the_context()) false thms i st | |
| 639 | ||
| 640 | fun search_tac thms i st = | |
| 641 | gen_shuffle_tac (the_context()) true thms i st | |
| 642 | ||
| 643 | fun shuffle_meth (thms:thm list) ctxt = | |
| 644 | let | |
| 645 | val thy = ProofContext.theory_of ctxt | |
| 646 | in | |
| 647 | Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy false (map (pair "") thms)) | |
| 648 | end | |
| 649 | ||
| 650 | fun search_meth ctxt = | |
| 651 | let | |
| 652 | val thy = ProofContext.theory_of ctxt | |
| 653 | val prems = ProofContext.prems_of ctxt | |
| 654 | in | |
| 655 | Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy true (map (pair "premise") prems)) | |
| 656 | end | |
| 657 | ||
| 658 | val print_shuffles = ShuffleData.print | |
| 659 | ||
| 660 | fun add_shuffle_rule thm thy = | |
| 661 | let | |
| 662 | val shuffles = ShuffleData.get thy | |
| 663 | in | |
| 664 | if exists (curry Thm.eq_thm thm) shuffles | |
| 665 | then (warning ((string_of_thm thm) ^ " already known to the shuffler"); | |
| 666 | thy) | |
| 667 | else ShuffleData.put (thm::shuffles) thy | |
| 668 | end | |
| 669 | ||
| 670 | fun shuffle_attr (thy,thm) = (add_shuffle_rule thm thy,thm) | |
| 671 | ||
| 672 | val setup = [Method.add_method ("shuffle_tac",Method.thms_ctxt_args shuffle_meth,"solve goal by shuffling terms around"),
 | |
| 673 | 	     Method.add_method ("search_tac",Method.ctxt_args search_meth,"search for suitable theorems"),
 | |
| 674 | ShuffleData.init, | |
| 675 | add_shuffle_rule weaken, | |
| 676 | add_shuffle_rule equiv_comm, | |
| 677 | add_shuffle_rule imp_comm, | |
| 678 | add_shuffle_rule Drule.norm_hhf_eq, | |
| 679 | add_shuffle_rule Drule.triv_forall_equality, | |
| 680 | 	     Attrib.add_attributes [("shuffle_rule",(Attrib.no_args shuffle_attr,K Attrib.undef_local_attribute),"tell the shuffler about the theorem")]]
 | |
| 681 | end |