| author | wenzelm | 
| Wed, 13 Apr 2005 18:46:22 +0200 | |
| changeset 15710 | b6b3df30cb0f | 
| parent 15574 | b1d1b5bfc464 | 
| child 16505 | c4b2e3cd84ab | 
| permissions | -rw-r--r-- | 
| 10769 | 1 | (* Title: TFL/tfl.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Konrad Slind, Cambridge University Computer Laboratory | |
| 4 | Copyright 1997 University of Cambridge | |
| 5 | ||
| 6 | First part of main module. | |
| 7 | *) | |
| 8 | ||
| 9 | signature PRIM = | |
| 10 | sig | |
| 11 | val trace: bool ref | |
| 14240 | 12 | val trace_thms: string -> thm list -> unit | 
| 13 | val trace_cterms: string -> cterm list -> unit | |
| 10769 | 14 | type pattern | 
| 15 |   val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
 | |
| 16 | val wfrec_definition0: theory -> string -> term -> term -> theory * thm | |
| 17 | val post_definition: thm list -> theory * (thm * pattern list) -> | |
| 18 |    {theory: theory,
 | |
| 19 | rules: thm, | |
| 20 | rows: int list, | |
| 21 | TCs: term list list, | |
| 22 | full_pats_TCs: (term * term list) list} | |
| 23 | val wfrec_eqns: theory -> xstring -> thm list -> term list -> | |
| 24 |    {WFR: term,
 | |
| 25 | SV: term list, | |
| 26 | proto_def: term, | |
| 27 | extracta: (thm * term list) list, | |
| 28 | pats: pattern list} | |
| 29 | val lazyR_def: theory -> xstring -> thm list -> term list -> | |
| 30 |    {theory: theory,
 | |
| 31 | rules: thm, | |
| 32 | R: term, | |
| 33 | SV: term list, | |
| 34 | full_pats_TCs: (term * term list) list, | |
| 35 | patterns : pattern list} | |
| 36 | val mk_induction: theory -> | |
| 37 |     {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
 | |
| 11632 | 38 |   val postprocess: bool -> {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm}
 | 
| 39 |     -> theory -> {rules: thm, induction: thm, TCs: term list list}
 | |
| 40 |     -> {rules: thm, induction: thm, nested_tcs: thm list}
 | |
| 10769 | 41 | end; | 
| 42 | ||
| 43 | structure Prim: PRIM = | |
| 44 | struct | |
| 45 | ||
| 46 | val trace = ref false; | |
| 47 | ||
| 48 | open BasisLibrary; | |
| 49 | ||
| 50 | structure R = Rules; | |
| 51 | structure S = USyntax; | |
| 52 | structure U = Utils; | |
| 53 | ||
| 54 | ||
| 55 | fun TFL_ERR func mesg = U.ERR {module = "Tfl", func = func, mesg = mesg};
 | |
| 56 | ||
| 57 | val concl = #2 o R.dest_thm; | |
| 58 | val hyp = #1 o R.dest_thm; | |
| 59 | ||
| 60 | val list_mk_type = U.end_itlist (curry (op -->)); | |
| 61 | ||
| 62 | fun enumerate xs = ListPair.zip(xs, 0 upto (length xs - 1)); | |
| 63 | ||
| 64 | fun front_last [] = raise TFL_ERR "front_last" "empty list" | |
| 65 | | front_last [x] = ([],x) | |
| 66 | | front_last (h::t) = | |
| 67 | let val (pref,x) = front_last t | |
| 68 | in | |
| 69 | (h::pref,x) | |
| 70 | end; | |
| 71 | ||
| 72 | ||
| 73 | (*--------------------------------------------------------------------------- | |
| 74 | * The next function is common to pattern-match translation and | |
| 75 | * proof of completeness of cases for the induction theorem. | |
| 76 | * | |
| 77 | * The curried function "gvvariant" returns a function to generate distinct | |
| 78 | * variables that are guaranteed not to be in names. The names of | |
| 79 | * the variables go u, v, ..., z, aa, ..., az, ... The returned | |
| 80 | * function contains embedded refs! | |
| 81 | *---------------------------------------------------------------------------*) | |
| 82 | fun gvvariant names = | |
| 83 | let val slist = ref names | |
| 84 | val vname = ref "u" | |
| 85 | fun new() = | |
| 86 | if !vname mem_string (!slist) | |
| 12902 | 87 | then (vname := Symbol.bump_string (!vname); new()) | 
| 10769 | 88 | else (slist := !vname :: !slist; !vname) | 
| 89 | in | |
| 90 | fn ty => Free(new(), ty) | |
| 91 | end; | |
| 92 | ||
| 93 | ||
| 94 | (*--------------------------------------------------------------------------- | |
| 95 | * Used in induction theorem production. This is the simple case of | |
| 96 | * partitioning up pattern rows by the leading constructor. | |
| 97 | *---------------------------------------------------------------------------*) | |
| 98 | fun ipartition gv (constructors,rows) = | |
| 99 | let fun pfail s = raise TFL_ERR "partition.part" s | |
| 100 |       fun part {constrs = [],   rows = [],   A} = rev A
 | |
| 101 |         | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
 | |
| 102 |         | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
 | |
| 103 |         | part {constrs = c::crst, rows,     A} =
 | |
| 104 | let val (Name,Ty) = dest_Const c | |
| 105 | val L = binder_types Ty | |
| 106 | val (in_group, not_in_group) = | |
| 107 | U.itlist (fn (row as (p::rst, rhs)) => | |
| 108 | fn (in_group,not_in_group) => | |
| 109 | let val (pc,args) = S.strip_comb p | |
| 110 | in if (#1(dest_Const pc) = Name) | |
| 111 | then ((args@rst, rhs)::in_group, not_in_group) | |
| 112 | else (in_group, row::not_in_group) | |
| 113 | end) rows ([],[]) | |
| 114 | val col_types = U.take type_of (length L, #1(hd in_group)) | |
| 115 | in | |
| 116 |           part{constrs = crst, rows = not_in_group,
 | |
| 117 |                A = {constructor = c,
 | |
| 118 | new_formals = map gv col_types, | |
| 119 | group = in_group}::A} | |
| 120 | end | |
| 121 |   in part{constrs = constructors, rows = rows, A = []}
 | |
| 122 | end; | |
| 123 | ||
| 124 | ||
| 125 | ||
| 126 | (*--------------------------------------------------------------------------- | |
| 127 | * Each pattern carries with it a tag (i,b) where | |
| 128 | * i is the clause it came from and | |
| 129 | * b=true indicates that clause was given by the user | |
| 130 | * (or is an instantiation of a user supplied pattern) | |
| 131 | * b=false --> i = ~1 | |
| 132 | *---------------------------------------------------------------------------*) | |
| 133 | ||
| 134 | type pattern = term * (int * bool) | |
| 135 | ||
| 136 | fun pattern_map f (tm,x) = (f tm, x); | |
| 137 | ||
| 138 | fun pattern_subst theta = pattern_map (subst_free theta); | |
| 139 | ||
| 140 | val pat_of = fst; | |
| 141 | fun row_of_pat x = fst (snd x); | |
| 142 | fun given x = snd (snd x); | |
| 143 | ||
| 144 | (*--------------------------------------------------------------------------- | |
| 145 | * Produce an instance of a constructor, plus genvars for its arguments. | |
| 146 | *---------------------------------------------------------------------------*) | |
| 147 | fun fresh_constr ty_match colty gv c = | |
| 148 | let val (_,Ty) = dest_Const c | |
| 149 | val L = binder_types Ty | |
| 150 | and ty = body_type Ty | |
| 151 | val ty_theta = ty_match ty colty | |
| 152 | val c' = S.inst ty_theta c | |
| 153 | val gvars = map (S.inst ty_theta o gv) L | |
| 154 | in (c', gvars) | |
| 155 | end; | |
| 156 | ||
| 157 | ||
| 158 | (*--------------------------------------------------------------------------- | |
| 159 | * Goes through a list of rows and picks out the ones beginning with a | |
| 160 | * pattern with constructor = Name. | |
| 161 | *---------------------------------------------------------------------------*) | |
| 162 | fun mk_group Name rows = | |
| 163 | U.itlist (fn (row as ((prfx, p::rst), rhs)) => | |
| 164 | fn (in_group,not_in_group) => | |
| 165 | let val (pc,args) = S.strip_comb p | |
| 166 | in if ((#1 (Term.dest_Const pc) = Name) handle TERM _ => false) | |
| 167 | then (((prfx,args@rst), rhs)::in_group, not_in_group) | |
| 168 | else (in_group, row::not_in_group) end) | |
| 169 | rows ([],[]); | |
| 170 | ||
| 171 | (*--------------------------------------------------------------------------- | |
| 172 | * Partition the rows. Not efficient: we should use hashing. | |
| 173 | *---------------------------------------------------------------------------*) | |
| 174 | fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows" | |
| 175 | | partition gv ty_match | |
| 176 | (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) = | |
| 177 | let val fresh = fresh_constr ty_match colty gv | |
| 178 |      fun part {constrs = [],      rows, A} = rev A
 | |
| 179 |        | part {constrs = c::crst, rows, A} =
 | |
| 180 | let val (c',gvars) = fresh c | |
| 181 | val (Name,Ty) = dest_Const c' | |
| 182 | val (in_group, not_in_group) = mk_group Name rows | |
| 183 | val in_group' = | |
| 184 | if (null in_group) (* Constructor not given *) | |
| 185 | then [((prfx, #2(fresh c)), (S.ARB res_ty, (~1,false)))] | |
| 186 | else in_group | |
| 187 | in | |
| 188 |          part{constrs = crst,
 | |
| 189 | rows = not_in_group, | |
| 190 |               A = {constructor = c',
 | |
| 191 | new_formals = gvars, | |
| 192 | group = in_group'}::A} | |
| 193 | end | |
| 194 | in part{constrs=constructors, rows=rows, A=[]}
 | |
| 195 | end; | |
| 196 | ||
| 197 | (*--------------------------------------------------------------------------- | |
| 198 | * Misc. routines used in mk_case | |
| 199 | *---------------------------------------------------------------------------*) | |
| 200 | ||
| 201 | fun mk_pat (c,l) = | |
| 202 | let val L = length (binder_types (type_of c)) | |
| 203 | fun build (prfx,tag,plist) = | |
| 15570 | 204 | let val args = Library.take (L,plist) | 
| 205 | and plist' = Library.drop(L,plist) | |
| 10769 | 206 | in (prfx,tag,list_comb(c,args)::plist') end | 
| 207 | in map build l end; | |
| 208 | ||
| 209 | fun v_to_prfx (prfx, v::pats) = (v::prfx,pats) | |
| 210 | | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx"; | |
| 211 | ||
| 212 | fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats) | |
| 213 | | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats"; | |
| 214 | ||
| 215 | ||
| 216 | (*---------------------------------------------------------------------------- | |
| 217 | * Translation of pattern terms into nested case expressions. | |
| 218 | * | |
| 219 | * This performs the translation and also builds the full set of patterns. | |
| 220 | * Thus it supports the construction of induction theorems even when an | |
| 221 | * incomplete set of patterns is given. | |
| 222 | *---------------------------------------------------------------------------*) | |
| 223 | ||
| 224 | fun mk_case ty_info ty_match usednames range_ty = | |
| 225 | let | |
| 226 | fun mk_case_fail s = raise TFL_ERR "mk_case" s | |
| 227 | val fresh_var = gvvariant usednames | |
| 228 | val divide = partition fresh_var ty_match | |
| 229 | fun expand constructors ty ((_,[]), _) = mk_case_fail"expand_var_row" | |
| 230 | | expand constructors ty (row as ((prfx, p::rst), rhs)) = | |
| 231 | if (is_Free p) | |
| 232 | then let val fresh = fresh_constr ty_match ty fresh_var | |
| 233 | fun expnd (c,gvs) = | |
| 234 | let val capp = list_comb(c,gvs) | |
| 235 | in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs) | |
| 236 | end | |
| 237 | in map expnd (map fresh constructors) end | |
| 238 | else [row] | |
| 239 |  fun mk{rows=[],...} = mk_case_fail"no rows"
 | |
| 240 |    | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
 | |
| 241 | ([(prfx,tag,[])], tm) | |
| 242 |    | mk{path=[], rows = _::_} = mk_case_fail"blunder"
 | |
| 243 |    | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
 | |
| 244 |         mk{path = path,
 | |
| 245 | rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst} | |
| 246 |    | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
 | |
| 247 | let val (pat_rectangle,rights) = ListPair.unzip rows | |
| 248 | val col0 = map(hd o #2) pat_rectangle | |
| 249 | in | |
| 250 | if (forall is_Free col0) | |
| 251 | then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e) | |
| 252 | (ListPair.zip (col0, rights)) | |
| 253 | val pat_rectangle' = map v_to_prfx pat_rectangle | |
| 254 |               val (pref_patl,tm) = mk{path = rstp,
 | |
| 255 | rows = ListPair.zip (pat_rectangle', | |
| 256 | rights')} | |
| 257 | in (map v_to_pats pref_patl, tm) | |
| 258 | end | |
| 259 | else | |
| 260 | let val pty as Type (ty_name,_) = type_of p | |
| 261 | in | |
| 262 | case (ty_info ty_name) | |
| 15531 | 263 |      of NONE => mk_case_fail("Not a known datatype: "^ty_name)
 | 
| 264 |       | SOME{case_const,constructors} =>
 | |
| 10769 | 265 | let | 
| 266 | val case_const_name = #1(dest_Const case_const) | |
| 267 | val nrows = List.concat (map (expand constructors pty) rows) | |
| 268 | val subproblems = divide(constructors, pty, range_ty, nrows) | |
| 269 | val groups = map #group subproblems | |
| 270 | and new_formals = map #new_formals subproblems | |
| 271 | and constructors' = map #constructor subproblems | |
| 272 |             val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
 | |
| 273 | (ListPair.zip (new_formals, groups)) | |
| 274 | val rec_calls = map mk news | |
| 275 | val (pat_rect,dtrees) = ListPair.unzip rec_calls | |
| 276 | val case_functions = map S.list_mk_abs | |
| 277 | (ListPair.zip (new_formals, dtrees)) | |
| 278 | val types = map type_of (case_functions@[u]) @ [range_ty] | |
| 279 | val case_const' = Const(case_const_name, list_mk_type types) | |
| 280 | val tree = list_comb(case_const', case_functions@[u]) | |
| 281 | val pat_rect1 = List.concat | |
| 282 | (ListPair.map mk_pat (constructors', pat_rect)) | |
| 283 | in (pat_rect1,tree) | |
| 284 | end | |
| 285 | end end | |
| 286 | in mk | |
| 287 | end; | |
| 288 | ||
| 289 | ||
| 290 | (* Repeated variable occurrences in a pattern are not allowed. *) | |
| 291 | fun FV_multiset tm = | |
| 292 | case (S.dest_term tm) | |
| 293 |      of S.VAR{Name,Ty} => [Free(Name,Ty)]
 | |
| 294 | | S.CONST _ => [] | |
| 295 |       | S.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
 | |
| 296 | | S.LAMB _ => raise TFL_ERR "FV_multiset" "lambda"; | |
| 297 | ||
| 298 | fun no_repeat_vars thy pat = | |
| 299 | let fun check [] = true | |
| 300 | | check (v::rst) = | |
| 301 | if mem_term (v,rst) then | |
| 302 | raise TFL_ERR "no_repeat_vars" | |
| 303 | (quote (#1 (dest_Free v)) ^ | |
| 304 | " occurs repeatedly in the pattern " ^ | |
| 305 | quote (string_of_cterm (Thry.typecheck thy pat))) | |
| 306 | else check rst | |
| 307 | in check (FV_multiset pat) | |
| 308 | end; | |
| 309 | ||
| 310 | fun dest_atom (Free p) = p | |
| 311 | | dest_atom (Const p) = p | |
| 312 | | dest_atom _ = raise TFL_ERR "dest_atom" "function name not an identifier"; | |
| 313 | ||
| 314 | fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q); | |
| 315 | ||
| 316 | local fun mk_functional_err s = raise TFL_ERR "mk_functional" s | |
| 317 | fun single [_$_] = | |
| 318 | mk_functional_err "recdef does not allow currying" | |
| 319 | | single [f] = f | |
| 320 | | single fs = | |
| 321 | (*multiple function names?*) | |
| 322 | if length (gen_distinct same_name fs) < length fs | |
| 323 | then mk_functional_err | |
| 324 | "The function being declared appears with multiple types" | |
| 325 | else mk_functional_err | |
| 326 | (Int.toString (length fs) ^ | |
| 327 | " distinct function names being declared") | |
| 328 | in | |
| 329 | fun mk_functional thy clauses = | |
| 330 | let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses | |
| 331 | handle TERM _ => raise TFL_ERR "mk_functional" | |
| 332 | "recursion equations must use the = relation") | |
| 333 | val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L) | |
| 334 | val atom = single (gen_distinct (op aconv) funcs) | |
| 335 | val (fname,ftype) = dest_atom atom | |
| 336 | val dummy = map (no_repeat_vars thy) pats | |
| 337 | val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats, | |
| 338 | map (fn (t,i) => (t,(i,true))) (enumerate R)) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 339 | val names = foldr add_term_names [] R | 
| 10769 | 340 | val atype = type_of(hd pats) | 
| 341 | and aname = variant names "a" | |
| 342 | val a = Free(aname,atype) | |
| 343 | val ty_info = Thry.match_info thy | |
| 344 | val ty_match = Thry.match_type thy | |
| 345 | val range_ty = type_of (hd R) | |
| 346 | val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty | |
| 347 |                                     {path=[a], rows=rows}
 | |
| 348 | val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts | |
| 349 | handle Match => mk_functional_err "error in pattern-match translation" | |
| 350 | val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1 | |
| 351 | val finals = map row_of_pat patts2 | |
| 352 | val originals = map (row_of_pat o #2) rows | |
| 353 | val dummy = case (originals\\finals) | |
| 354 | of [] => () | |
| 355 | | L => mk_functional_err | |
| 356 |  ("The following clauses are redundant (covered by preceding clauses): " ^
 | |
| 357 | commas (map (fn i => Int.toString (i + 1)) L)) | |
| 358 |  in {functional = Abs(Sign.base_name fname, ftype,
 | |
| 359 | abstract_over (atom, | |
| 360 | absfree(aname,atype, case_tm))), | |
| 361 | pats = patts2} | |
| 362 | end end; | |
| 363 | ||
| 364 | ||
| 365 | (*---------------------------------------------------------------------------- | |
| 366 | * | |
| 367 | * PRINCIPLES OF DEFINITION | |
| 368 | * | |
| 369 | *---------------------------------------------------------------------------*) | |
| 370 | ||
| 371 | ||
| 372 | (*For Isabelle, the lhs of a definition must be a constant.*) | |
| 373 | fun mk_const_def sign (Name, Ty, rhs) = | |
| 15531 | 374 | Sign.infer_types (Sign.pp sign) sign (K NONE) (K NONE) [] false | 
| 10769 | 375 |                ([Const("==",dummyT) $ Const(Name,Ty) $ rhs], propT)
 | 
| 376 | |> #1; | |
| 377 | ||
| 378 | (*Make all TVars available for instantiation by adding a ? to the front*) | |
| 379 | fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts) | |
| 380 |   | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
 | |
| 381 |   | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
 | |
| 382 | ||
| 383 | local val f_eq_wfrec_R_M = | |
| 384 | #ant(S.dest_imp(#2(S.strip_forall (concl Thms.WFREC_COROLLARY)))) | |
| 385 |       val {lhs=f, rhs} = S.dest_eq f_eq_wfrec_R_M
 | |
| 386 | val (fname,_) = dest_Free f | |
| 387 | val (wfrec,_) = S.strip_comb rhs | |
| 388 | in | |
| 389 | fun wfrec_definition0 thy fid R (functional as Abs(Name, Ty, _)) = | |
| 390 | let val def_name = if Name<>fid then | |
| 391 | raise TFL_ERR "wfrec_definition0" | |
| 392 |                                       ("Expected a definition of " ^
 | |
| 393 | quote fid ^ " but found one of " ^ | |
| 394 | quote Name) | |
| 395 | else Name ^ "_def" | |
| 396 | val wfrec_R_M = map_term_types poly_tvars | |
| 397 | (wfrec $ map_term_types poly_tvars R) | |
| 398 | $ functional | |
| 399 | val def_term = mk_const_def (Theory.sign_of thy) (Name, Ty, wfrec_R_M) | |
| 400 | val (thy', [def]) = PureThy.add_defs_i false [Thm.no_attributes (def_name, def_term)] thy | |
| 401 | in (thy', def) end; | |
| 402 | end; | |
| 403 | ||
| 404 | ||
| 405 | ||
| 406 | (*--------------------------------------------------------------------------- | |
| 407 | * This structure keeps track of congruence rules that aren't derived | |
| 408 | * from a datatype definition. | |
| 409 | *---------------------------------------------------------------------------*) | |
| 410 | fun extraction_thms thy = | |
| 411 |  let val {case_rewrites,case_congs} = Thry.extract_info thy
 | |
| 412 | in (case_rewrites, case_congs) | |
| 413 | end; | |
| 414 | ||
| 415 | ||
| 416 | (*--------------------------------------------------------------------------- | |
| 417 | * Pair patterns with termination conditions. The full list of patterns for | |
| 418 | * a definition is merged with the TCs arising from the user-given clauses. | |
| 419 | * There can be fewer clauses than the full list, if the user omitted some | |
| 420 | * cases. This routine is used to prepare input for mk_induction. | |
| 421 | *---------------------------------------------------------------------------*) | |
| 422 | fun merge full_pats TCs = | |
| 423 | let fun insert (p,TCs) = | |
| 424 | let fun insrt ((x as (h,[]))::rst) = | |
| 425 | if (p aconv h) then (p,TCs)::rst else x::insrt rst | |
| 426 | | insrt (x::rst) = x::insrt rst | |
| 427 | | insrt[] = raise TFL_ERR "merge.insert" "pattern not found" | |
| 428 | in insrt end | |
| 429 | fun pass ([],ptcl_final) = ptcl_final | |
| 430 | | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl) | |
| 431 | in | |
| 432 | pass (TCs, map (fn p => (p,[])) full_pats) | |
| 433 | end; | |
| 434 | ||
| 435 | ||
| 15570 | 436 | fun givens pats = map pat_of (List.filter given pats); | 
| 10769 | 437 | |
| 438 | fun post_definition meta_tflCongs (theory, (def, pats)) = | |
| 439 | let val tych = Thry.typecheck theory | |
| 440 | val f = #lhs(S.dest_eq(concl def)) | |
| 441 | val corollary = R.MATCH_MP Thms.WFREC_COROLLARY def | |
| 15570 | 442 | val pats' = List.filter given pats | 
| 10769 | 443 | val given_pats = map pat_of pats' | 
| 444 | val rows = map row_of_pat pats' | |
| 445 | val WFR = #ant(S.dest_imp(concl corollary)) | |
| 446 | val R = #Rand(S.dest_comb WFR) | |
| 447 | val corollary' = R.UNDISCH corollary (* put WF R on assums *) | |
| 448 | val corollaries = map (fn pat => R.SPEC (tych pat) corollary') | |
| 449 | given_pats | |
| 450 | val (case_rewrites,context_congs) = extraction_thms theory | |
| 14219 | 451 | (*case_ss causes minimal simplification: bodies of case expressions are | 
| 452 | not simplified. Otherwise large examples (Red-Black trees) are too | |
| 453 | slow.*) | |
| 14217 
9f5679e97eac
Fixed inefficiency in post_definition by adding weak case congruence
 berghofe parents: 
12902diff
changeset | 454 | val case_ss = HOL_basic_ss addcongs | 
| 
9f5679e97eac
Fixed inefficiency in post_definition by adding weak case congruence
 berghofe parents: 
12902diff
changeset | 455 | DatatypePackage.weak_case_congs_of theory addsimps case_rewrites | 
| 
9f5679e97eac
Fixed inefficiency in post_definition by adding weak case congruence
 berghofe parents: 
12902diff
changeset | 456 | val corollaries' = map (Simplifier.simplify case_ss) corollaries | 
| 10769 | 457 | val extract = R.CONTEXT_REWRITE_RULE | 
| 458 | (f, [R], cut_apply, meta_tflCongs@context_congs) | |
| 459 | val (rules, TCs) = ListPair.unzip (map extract corollaries') | |
| 460 | val rules0 = map (rewrite_rule [Thms.CUT_DEF]) rules | |
| 461 | val mk_cond_rule = R.FILTER_DISCH_ALL(not o curry (op aconv) WFR) | |
| 462 | val rules1 = R.LIST_CONJ(map mk_cond_rule rules0) | |
| 463 | in | |
| 464 |  {theory = theory,
 | |
| 465 | rules = rules1, | |
| 466 | rows = rows, | |
| 467 | full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)), | |
| 468 | TCs = TCs} | |
| 469 | end; | |
| 470 | ||
| 471 | ||
| 472 | (*--------------------------------------------------------------------------- | |
| 473 | * Perform the extraction without making the definition. Definition and | |
| 474 | * extraction commute for the non-nested case. (Deferred recdefs) | |
| 475 | * | |
| 476 | * The purpose of wfrec_eqns is merely to instantiate the recursion theorem | |
| 477 | * and extract termination conditions: no definition is made. | |
| 478 | *---------------------------------------------------------------------------*) | |
| 479 | ||
| 480 | fun wfrec_eqns thy fid tflCongs eqns = | |
| 481 |  let val {lhs,rhs} = S.dest_eq (hd eqns)
 | |
| 482 | val (f,args) = S.strip_comb lhs | |
| 483 | val (fname,fty) = dest_atom f | |
| 484 | val (SV,a) = front_last args (* SV = schematic variables *) | |
| 485 | val g = list_comb(f,SV) | |
| 486 | val h = Free(fname,type_of g) | |
| 487 | val eqns1 = map (subst_free[(g,h)]) eqns | |
| 488 |      val {functional as Abs(Name, Ty, _),  pats} = mk_functional thy eqns1
 | |
| 489 | val given_pats = givens pats | |
| 490 | (* val f = Free(Name,Ty) *) | |
| 491 |      val Type("fun", [f_dty, f_rty]) = Ty
 | |
| 492 | val dummy = if Name<>fid then | |
| 493 | raise TFL_ERR "wfrec_eqns" | |
| 494 |                                       ("Expected a definition of " ^
 | |
| 495 | quote fid ^ " but found one of " ^ | |
| 496 | quote Name) | |
| 497 | else () | |
| 498 | val (case_rewrites,context_congs) = extraction_thms thy | |
| 499 | val tych = Thry.typecheck thy | |
| 500 | val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY | |
| 501 |      val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
 | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 502 | val R = Free (variant (foldr add_term_names [] eqns) Rname, | 
| 10769 | 503 | Rtype) | 
| 504 | val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0 | |
| 505 | val ([proto_def, WFR],_) = S.strip_imp(concl WFREC_THM) | |
| 506 | val dummy = | |
| 507 | if !trace then | |
| 508 |                writeln ("ORIGINAL PROTO_DEF: " ^
 | |
| 509 | Sign.string_of_term (Theory.sign_of thy) proto_def) | |
| 510 | else () | |
| 511 | val R1 = S.rand WFR | |
| 512 | val corollary' = R.UNDISCH(R.UNDISCH WFREC_THM) | |
| 513 | val corollaries = map (fn pat => R.SPEC (tych pat) corollary') given_pats | |
| 514 | val corollaries' = map (rewrite_rule case_rewrites) corollaries | |
| 515 | fun extract X = R.CONTEXT_REWRITE_RULE | |
| 516 | (f, R1::SV, cut_apply, tflCongs@context_congs) X | |
| 517 |  in {proto_def = proto_def,
 | |
| 518 | SV=SV, | |
| 519 | WFR=WFR, | |
| 520 | pats=pats, | |
| 521 | extracta = map extract corollaries'} | |
| 522 | end; | |
| 523 | ||
| 524 | ||
| 525 | (*--------------------------------------------------------------------------- | |
| 526 | * Define the constant after extracting the termination conditions. The | |
| 527 | * wellfounded relation used in the definition is computed by using the | |
| 528 | * choice operator on the extracted conditions (plus the condition that | |
| 529 | * such a relation must be wellfounded). | |
| 530 | *---------------------------------------------------------------------------*) | |
| 531 | ||
| 532 | fun lazyR_def thy fid tflCongs eqns = | |
| 533 |  let val {proto_def,WFR,pats,extracta,SV} =
 | |
| 534 | wfrec_eqns thy fid tflCongs eqns | |
| 535 | val R1 = S.rand WFR | |
| 536 | val f = #lhs(S.dest_eq proto_def) | |
| 537 | val (extractants,TCl) = ListPair.unzip extracta | |
| 538 | val dummy = if !trace | |
| 539 | then (writeln "Extractants = "; | |
| 540 | prths extractants; | |
| 541 | ()) | |
| 542 | else () | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 543 | val TCs = foldr (gen_union (op aconv)) [] TCl | 
| 10769 | 544 | val full_rqt = WFR::TCs | 
| 545 |      val R' = S.mk_select{Bvar=R1, Body=S.list_mk_conj full_rqt}
 | |
| 546 | val R'abs = S.rand R' | |
| 547 | val proto_def' = subst_free[(R1,R')] proto_def | |
| 548 |      val dummy = if !trace then writeln ("proto_def' = " ^
 | |
| 549 | Sign.string_of_term | |
| 550 | (Theory.sign_of thy) proto_def') | |
| 551 | else () | |
| 552 |      val {lhs,rhs} = S.dest_eq proto_def'
 | |
| 553 | val (c,args) = S.strip_comb lhs | |
| 554 | val (Name,Ty) = dest_atom c | |
| 555 | val defn = mk_const_def (Theory.sign_of thy) | |
| 556 | (Name, Ty, S.list_mk_abs (args,rhs)) | |
| 557 | val (theory, [def0]) = | |
| 558 | thy | |
| 559 | |> PureThy.add_defs_i false | |
| 560 | [Thm.no_attributes (fid ^ "_def", defn)] | |
| 561 | val def = freezeT def0; | |
| 562 |      val dummy = if !trace then writeln ("DEF = " ^ string_of_thm def)
 | |
| 563 | else () | |
| 564 | (* val fconst = #lhs(S.dest_eq(concl def)) *) | |
| 565 | val tych = Thry.typecheck theory | |
| 566 | val full_rqt_prop = map (Dcterm.mk_prop o tych) full_rqt | |
| 567 | (*lcp: a lot of object-logic inference to remove*) | |
| 568 | val baz = R.DISCH_ALL | |
| 569 | (U.itlist R.DISCH full_rqt_prop | |
| 570 | (R.LIST_CONJ extractants)) | |
| 571 |      val dum = if !trace then writeln ("baz = " ^ string_of_thm baz)
 | |
| 572 | else () | |
| 573 | val f_free = Free (fid, fastype_of f) (*'cos f is a Const*) | |
| 574 | val SV' = map tych SV; | |
| 575 | val SVrefls = map reflexive SV' | |
| 576 | val def0 = (U.rev_itlist (fn x => fn th => R.rbeta(combination th x)) | |
| 577 | SVrefls def) | |
| 578 | RS meta_eq_to_obj_eq | |
| 579 | val def' = R.MP (R.SPEC (tych R') (R.GEN (tych R1) baz)) def0 | |
| 580 | val body_th = R.LIST_CONJ (map R.ASSUME full_rqt_prop) | |
| 11455 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 581 | val SELECT_AX = (*in this way we hope to avoid a STATIC dependence upon | 
| 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 582 | theory Hilbert_Choice*) | 
| 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 583 | thm "Hilbert_Choice.tfl_some" | 
| 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 584 | handle ERROR => error | 
| 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 585 | "defer_recdef requires theory Main or at least Hilbert_Choice as parent" | 
| 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 paulson parents: 
10769diff
changeset | 586 | val bar = R.MP (R.ISPECL[tych R'abs, tych R1] SELECT_AX) body_th | 
| 10769 | 587 |  in {theory = theory, R=R1, SV=SV,
 | 
| 588 | rules = U.rev_itlist (U.C R.MP) (R.CONJUNCTS bar) def', | |
| 589 | full_pats_TCs = merge (map pat_of pats) (ListPair.zip (givens pats, TCl)), | |
| 590 | patterns = pats} | |
| 591 | end; | |
| 592 | ||
| 593 | ||
| 594 | ||
| 595 | (*---------------------------------------------------------------------------- | |
| 596 | * | |
| 597 | * INDUCTION THEOREM | |
| 598 | * | |
| 599 | *---------------------------------------------------------------------------*) | |
| 600 | ||
| 601 | ||
| 602 | (*------------------------ Miscellaneous function -------------------------- | |
| 603 | * | |
| 604 | * [x_1,...,x_n] ?v_1...v_n. M[v_1,...,v_n] | |
| 605 | * ----------------------------------------------------------- | |
| 606 | * ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]), | |
| 607 | * ... | |
| 608 | * (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] ) | |
| 609 | * | |
| 610 | * This function is totally ad hoc. Used in the production of the induction | |
| 611 | * theorem. The nchotomy theorem can have clauses that look like | |
| 612 | * | |
| 613 | * ?v1..vn. z = C vn..v1 | |
| 614 | * | |
| 615 | * in which the order of quantification is not the order of occurrence of the | |
| 616 | * quantified variables as arguments to C. Since we have no control over this | |
| 617 | * aspect of the nchotomy theorem, we make the correspondence explicit by | |
| 618 | * pairing the incoming new variable with the term it gets beta-reduced into. | |
| 619 | *---------------------------------------------------------------------------*) | |
| 620 | ||
| 621 | fun alpha_ex_unroll (xlist, tm) = | |
| 622 | let val (qvars,body) = S.strip_exists tm | |
| 623 | val vlist = #2(S.strip_comb (S.rhs body)) | |
| 624 | val plist = ListPair.zip (vlist, xlist) | |
| 15570 | 625 | val args = map (fn qv => valOf (gen_assoc (op aconv) (plist, qv))) qvars | 
| 15531 | 626 | handle Option => sys_error | 
| 10769 | 627 | "TFL fault [alpha_ex_unroll]: no correspondence" | 
| 628 | fun build ex [] = [] | |
| 629 | | build (_$rex) (v::rst) = | |
| 630 | let val ex1 = betapply(rex, v) | |
| 631 | in ex1 :: build ex1 rst | |
| 632 | end | |
| 633 | val (nex::exl) = rev (tm::build tm args) | |
| 634 | in | |
| 635 | (nex, ListPair.zip (args, rev exl)) | |
| 636 | end; | |
| 637 | ||
| 638 | ||
| 639 | ||
| 640 | (*---------------------------------------------------------------------------- | |
| 641 | * | |
| 642 | * PROVING COMPLETENESS OF PATTERNS | |
| 643 | * | |
| 644 | *---------------------------------------------------------------------------*) | |
| 645 | ||
| 646 | fun mk_case ty_info usednames thy = | |
| 647 | let | |
| 648 | val divide = ipartition (gvvariant usednames) | |
| 649 | val tych = Thry.typecheck thy | |
| 650 | fun tych_binding(x,y) = (tych x, tych y) | |
| 651 | fun fail s = raise TFL_ERR "mk_case" s | |
| 652 |  fun mk{rows=[],...} = fail"no rows"
 | |
| 653 |    | mk{path=[], rows = [([], (thm, bindings))]} =
 | |
| 654 | R.IT_EXISTS (map tych_binding bindings) thm | |
| 655 |    | mk{path = u::rstp, rows as (p::_, _)::_} =
 | |
| 656 | let val (pat_rectangle,rights) = ListPair.unzip rows | |
| 657 | val col0 = map hd pat_rectangle | |
| 658 | val pat_rectangle' = map tl pat_rectangle | |
| 659 | in | |
| 660 | if (forall is_Free col0) (* column 0 is all variables *) | |
| 661 | then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)])) | |
| 662 | (ListPair.zip (rights, col0)) | |
| 663 |           in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
 | |
| 664 | end | |
| 665 | else (* column 0 is all constructors *) | |
| 666 | let val Type (ty_name,_) = type_of p | |
| 667 | in | |
| 668 | case (ty_info ty_name) | |
| 15531 | 669 |      of NONE => fail("Not a known datatype: "^ty_name)
 | 
| 670 |       | SOME{constructors,nchotomy} =>
 | |
| 10769 | 671 | let val thm' = R.ISPEC (tych u) nchotomy | 
| 672 | val disjuncts = S.strip_disj (concl thm') | |
| 673 | val subproblems = divide(constructors, rows) | |
| 674 | val groups = map #group subproblems | |
| 675 | and new_formals = map #new_formals subproblems | |
| 676 | val existentials = ListPair.map alpha_ex_unroll | |
| 677 | (new_formals, disjuncts) | |
| 678 | val constraints = map #1 existentials | |
| 679 | val vexl = map #2 existentials | |
| 680 | fun expnd tm (pats,(th,b)) = (pats,(R.SUBS[R.ASSUME(tych tm)]th,b)) | |
| 681 |             val news = map (fn (nf,rows,c) => {path = nf@rstp,
 | |
| 682 | rows = map (expnd c) rows}) | |
| 683 | (U.zip3 new_formals groups constraints) | |
| 684 | val recursive_thms = map mk news | |
| 15570 | 685 | val build_exists = Library.foldr | 
| 10769 | 686 | (fn((x,t), th) => | 
| 687 | R.CHOOSE (tych x, R.ASSUME (tych t)) th) | |
| 688 | val thms' = ListPair.map build_exists (vexl, recursive_thms) | |
| 689 | val same_concls = R.EVEN_ORS thms' | |
| 690 | in R.DISJ_CASESL thm' same_concls | |
| 691 | end | |
| 692 | end end | |
| 693 | in mk | |
| 694 | end; | |
| 695 | ||
| 696 | ||
| 697 | fun complete_cases thy = | |
| 698 | let val tych = Thry.typecheck thy | |
| 699 | val ty_info = Thry.induct_info thy | |
| 700 | in fn pats => | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 701 | let val names = foldr add_term_names [] pats | 
| 10769 | 702 | val T = type_of (hd pats) | 
| 703 | val aname = Term.variant names "a" | |
| 704 | val vname = Term.variant (aname::names) "v" | |
| 705 | val a = Free (aname, T) | |
| 706 | val v = Free (vname, T) | |
| 707 | val a_eq_v = HOLogic.mk_eq(a,v) | |
| 708 |      val ex_th0 = R.EXISTS (tych (S.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
 | |
| 709 | (R.REFL (tych a)) | |
| 710 | val th0 = R.ASSUME (tych a_eq_v) | |
| 711 | val rows = map (fn x => ([x], (th0,[]))) pats | |
| 712 | in | |
| 713 | R.GEN (tych a) | |
| 714 | (R.RIGHT_ASSOC | |
| 715 | (R.CHOOSE(tych v, ex_th0) | |
| 716 | (mk_case ty_info (vname::aname::names) | |
| 717 |                  thy {path=[v], rows=rows})))
 | |
| 718 | end end; | |
| 719 | ||
| 720 | ||
| 721 | (*--------------------------------------------------------------------------- | |
| 722 | * Constructing induction hypotheses: one for each recursive call. | |
| 723 | * | |
| 724 | * Note. R will never occur as a variable in the ind_clause, because | |
| 725 | * to do so, it would have to be from a nested definition, and we don't | |
| 726 | * allow nested defns to have R variable. | |
| 727 | * | |
| 728 | * Note. When the context is empty, there can be no local variables. | |
| 729 | *---------------------------------------------------------------------------*) | |
| 730 | (* | |
| 731 | local infix 5 ==> | |
| 732 |       fun (tm1 ==> tm2) = S.mk_imp{ant = tm1, conseq = tm2}
 | |
| 733 | in | |
| 734 | fun build_ih f P (pat,TCs) = | |
| 735 | let val globals = S.free_vars_lr pat | |
| 15570 | 736 | fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) | 
| 10769 | 737 | fun dest_TC tm = | 
| 738 | let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm)) | |
| 739 | val (R,y,_) = S.dest_relation R_y_pat | |
| 740 | val P_y = if (nested tm) then R_y_pat ==> P$y else P$y | |
| 741 | in case cntxt | |
| 742 | of [] => (P_y, (tm,[])) | |
| 743 | | _ => let | |
| 744 | val imp = S.list_mk_conj cntxt ==> P_y | |
| 745 | val lvs = gen_rems (op aconv) (S.free_vars_lr imp, globals) | |
| 746 | val locals = #2(U.pluck (curry (op aconv) P) lvs) handle U.ERR _ => lvs | |
| 747 | in (S.list_mk_forall(locals,imp), (tm,locals)) end | |
| 748 | end | |
| 749 | in case TCs | |
| 750 | of [] => (S.list_mk_forall(globals, P$pat), []) | |
| 751 | | _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs) | |
| 752 | val ind_clause = S.list_mk_conj ihs ==> P$pat | |
| 753 | in (S.list_mk_forall(globals,ind_clause), TCs_locals) | |
| 754 | end | |
| 755 | end | |
| 756 | end; | |
| 757 | *) | |
| 758 | ||
| 759 | local infix 5 ==> | |
| 760 |       fun (tm1 ==> tm2) = S.mk_imp{ant = tm1, conseq = tm2}
 | |
| 761 | in | |
| 762 | fun build_ih f (P,SV) (pat,TCs) = | |
| 763 | let val pat_vars = S.free_vars_lr pat | |
| 764 | val globals = pat_vars@SV | |
| 15570 | 765 | fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) | 
| 10769 | 766 | fun dest_TC tm = | 
| 767 | let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm)) | |
| 768 | val (R,y,_) = S.dest_relation R_y_pat | |
| 769 | val P_y = if (nested tm) then R_y_pat ==> P$y else P$y | |
| 770 | in case cntxt | |
| 771 | of [] => (P_y, (tm,[])) | |
| 772 | | _ => let | |
| 773 | val imp = S.list_mk_conj cntxt ==> P_y | |
| 774 | val lvs = gen_rems (op aconv) (S.free_vars_lr imp, globals) | |
| 775 | val locals = #2(U.pluck (curry (op aconv) P) lvs) handle U.ERR _ => lvs | |
| 776 | in (S.list_mk_forall(locals,imp), (tm,locals)) end | |
| 777 | end | |
| 778 | in case TCs | |
| 779 | of [] => (S.list_mk_forall(pat_vars, P$pat), []) | |
| 780 | | _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs) | |
| 781 | val ind_clause = S.list_mk_conj ihs ==> P$pat | |
| 782 | in (S.list_mk_forall(pat_vars,ind_clause), TCs_locals) | |
| 783 | end | |
| 784 | end | |
| 785 | end; | |
| 786 | ||
| 787 | (*--------------------------------------------------------------------------- | |
| 788 | * This function makes good on the promise made in "build_ih". | |
| 789 | * | |
| 790 | * Input is tm = "(!y. R y pat ==> P y) ==> P pat", | |
| 791 | * TCs = TC_1[pat] ... TC_n[pat] | |
| 792 | * thm = ih1 /\ ... /\ ih_n |- ih[pat] | |
| 793 | *---------------------------------------------------------------------------*) | |
| 794 | fun prove_case f thy (tm,TCs_locals,thm) = | |
| 795 | let val tych = Thry.typecheck thy | |
| 796 | val antc = tych(#ant(S.dest_imp tm)) | |
| 797 | val thm' = R.SPEC_ALL thm | |
| 15570 | 798 | fun nested tm = isSome (S.find_term (curry (op aconv) f) tm) | 
| 10769 | 799 | fun get_cntxt TC = tych(#ant(S.dest_imp(#2(S.strip_forall(concl TC))))) | 
| 800 | fun mk_ih ((TC,locals),th2,nested) = | |
| 801 | R.GENL (map tych locals) | |
| 802 | (if nested then R.DISCH (get_cntxt TC) th2 handle U.ERR _ => th2 | |
| 803 | else if S.is_imp (concl TC) then R.IMP_TRANS TC th2 | |
| 804 | else R.MP th2 TC) | |
| 805 | in | |
| 806 | R.DISCH antc | |
| 807 | (if S.is_imp(concl thm') (* recursive calls in this clause *) | |
| 808 | then let val th1 = R.ASSUME antc | |
| 809 | val TCs = map #1 TCs_locals | |
| 810 | val ylist = map (#2 o S.dest_relation o #2 o S.strip_imp o | |
| 811 | #2 o S.strip_forall) TCs | |
| 812 | val TClist = map (fn(TC,lvs) => (R.SPEC_ALL(R.ASSUME(tych TC)),lvs)) | |
| 813 | TCs_locals | |
| 814 | val th2list = map (U.C R.SPEC th1 o tych) ylist | |
| 815 | val nlist = map nested TCs | |
| 816 | val triples = U.zip3 TClist th2list nlist | |
| 817 | val Pylist = map mk_ih triples | |
| 818 | in R.MP thm' (R.LIST_CONJ Pylist) end | |
| 819 | else thm') | |
| 820 | end; | |
| 821 | ||
| 822 | ||
| 823 | (*--------------------------------------------------------------------------- | |
| 824 | * | |
| 825 | * x = (v1,...,vn) |- M[x] | |
| 826 | * --------------------------------------------- | |
| 827 | * ?v1 ... vn. x = (v1,...,vn) |- M[x] | |
| 828 | * | |
| 829 | *---------------------------------------------------------------------------*) | |
| 830 | fun LEFT_ABS_VSTRUCT tych thm = | |
| 831 | let fun CHOOSER v (tm,thm) = | |
| 832 |         let val ex_tm = S.mk_exists{Bvar=v,Body=tm}
 | |
| 833 | in (ex_tm, R.CHOOSE(tych v, R.ASSUME (tych ex_tm)) thm) | |
| 834 | end | |
| 15570 | 835 | val [veq] = List.filter (can S.dest_eq) (#1 (R.dest_thm thm)) | 
| 10769 | 836 |       val {lhs,rhs} = S.dest_eq veq
 | 
| 837 | val L = S.free_vars_lr rhs | |
| 838 | in #2 (U.itlist CHOOSER L (veq,thm)) end; | |
| 839 | ||
| 840 | ||
| 841 | (*---------------------------------------------------------------------------- | |
| 842 | * Input : f, R, and [(pat1,TCs1),..., (patn,TCsn)] | |
| 843 | * | |
| 844 | * Instantiates WF_INDUCTION_THM, getting Sinduct and then tries to prove | |
| 845 | * recursion induction (Rinduct) by proving the antecedent of Sinduct from | |
| 846 | * the antecedent of Rinduct. | |
| 847 | *---------------------------------------------------------------------------*) | |
| 848 | fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
 | |
| 849 | let val tych = Thry.typecheck thy | |
| 850 | val Sinduction = R.UNDISCH (R.ISPEC (tych R) Thms.WF_INDUCTION_THM) | |
| 851 | val (pats,TCsl) = ListPair.unzip pat_TCs_list | |
| 852 | val case_thm = complete_cases thy pats | |
| 853 | val domain = (type_of o hd) pats | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 854 | val Pname = Term.variant (foldr (Library.foldr add_term_names) | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 855 | [] (pats::TCsl)) "P" | 
| 10769 | 856 | val P = Free(Pname, domain --> HOLogic.boolT) | 
| 857 | val Sinduct = R.SPEC (tych P) Sinduction | |
| 858 | val Sinduct_assumf = S.rand ((#ant o S.dest_imp o concl) Sinduct) | |
| 859 | val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list | |
| 860 | val (Rassums,TCl') = ListPair.unzip Rassums_TCl' | |
| 861 | val Rinduct_assum = R.ASSUME (tych (S.list_mk_conj Rassums)) | |
| 862 | val cases = map (fn pat => betapply (Sinduct_assumf, pat)) pats | |
| 863 | val tasks = U.zip3 cases TCl' (R.CONJUNCTS Rinduct_assum) | |
| 864 | val proved_cases = map (prove_case fconst thy) tasks | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 865 | val v = Free (variant (foldr add_term_names [] (map concl proved_cases)) | 
| 10769 | 866 | "v", | 
| 867 | domain) | |
| 868 | val vtyped = tych v | |
| 869 | val substs = map (R.SYM o R.ASSUME o tych o (curry HOLogic.mk_eq v)) pats | |
| 870 | val proved_cases1 = ListPair.map (fn (th,th') => R.SUBS[th]th') | |
| 871 | (substs, proved_cases) | |
| 872 | val abs_cases = map (LEFT_ABS_VSTRUCT tych) proved_cases1 | |
| 873 | val dant = R.GEN vtyped (R.DISJ_CASESL (R.ISPEC vtyped case_thm) abs_cases) | |
| 874 | val dc = R.MP Sinduct dant | |
| 875 | val Parg_ty = type_of(#Bvar(S.dest_forall(concl dc))) | |
| 876 | val vars = map (gvvariant[Pname]) (S.strip_prod_type Parg_ty) | |
| 877 | val dc' = U.itlist (R.GEN o tych) vars | |
| 878 | (R.SPEC (tych(S.mk_vstruct Parg_ty vars)) dc) | |
| 879 | in | |
| 880 | R.GEN (tych P) (R.DISCH (tych(concl Rinduct_assum)) dc') | |
| 881 | end | |
| 882 | handle U.ERR _ => raise TFL_ERR "mk_induction" "failed derivation"; | |
| 883 | ||
| 884 | ||
| 885 | ||
| 886 | ||
| 887 | (*--------------------------------------------------------------------------- | |
| 888 | * | |
| 889 | * POST PROCESSING | |
| 890 | * | |
| 891 | *---------------------------------------------------------------------------*) | |
| 892 | ||
| 893 | ||
| 894 | fun simplify_induction thy hth ind = | |
| 895 | let val tych = Thry.typecheck thy | |
| 896 | val (asl,_) = R.dest_thm ind | |
| 897 | val (_,tc_eq_tc') = R.dest_thm hth | |
| 898 | val tc = S.lhs tc_eq_tc' | |
| 899 | fun loop [] = ind | |
| 900 | | loop (asm::rst) = | |
| 901 | if (can (Thry.match_term thy asm) tc) | |
| 902 | then R.UNDISCH | |
| 903 | (R.MATCH_MP | |
| 904 | (R.MATCH_MP Thms.simp_thm (R.DISCH (tych asm) ind)) | |
| 905 | hth) | |
| 906 | else loop rst | |
| 907 | in loop asl | |
| 908 | end; | |
| 909 | ||
| 910 | ||
| 911 | (*--------------------------------------------------------------------------- | |
| 912 | * The termination condition is an antecedent to the rule, and an | |
| 913 | * assumption to the theorem. | |
| 914 | *---------------------------------------------------------------------------*) | |
| 915 | fun elim_tc tcthm (rule,induction) = | |
| 916 | (R.MP rule tcthm, R.PROVE_HYP tcthm induction) | |
| 917 | ||
| 918 | ||
| 14240 | 919 | fun trace_thms s L = | 
| 920 | if !trace then writeln (cat_lines (s :: map string_of_thm L)) | |
| 921 | else (); | |
| 922 | ||
| 923 | fun trace_cterms s L = | |
| 924 | if !trace then writeln (cat_lines (s :: map string_of_cterm L)) | |
| 925 | else ();; | |
| 926 | ||
| 927 | ||
| 11632 | 928 | fun postprocess strict {wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
 | 
| 10769 | 929 | let val tych = Thry.typecheck theory | 
| 11632 | 930 | val prove = R.prove strict; | 
| 10769 | 931 | |
| 932 | (*--------------------------------------------------------------------- | |
| 933 | * Attempt to eliminate WF condition. It's the only assumption of rules | |
| 934 | *---------------------------------------------------------------------*) | |
| 935 | val (rules1,induction1) = | |
| 11632 | 936 | let val thm = prove(tych(HOLogic.mk_Trueprop | 
| 10769 | 937 | (hd(#1(R.dest_thm rules)))), | 
| 938 | wf_tac) | |
| 939 | in (R.PROVE_HYP thm rules, R.PROVE_HYP thm induction) | |
| 940 | end handle U.ERR _ => (rules,induction); | |
| 941 | ||
| 942 | (*---------------------------------------------------------------------- | |
| 943 | * The termination condition (tc) is simplified to |- tc = tc' (there | |
| 944 | * might not be a change!) and then 3 attempts are made: | |
| 945 | * | |
| 946 | * 1. if |- tc = T, then eliminate it with eqT; otherwise, | |
| 947 | * 2. apply the terminator to tc'. If |- tc' = T then eliminate; else | |
| 948 | * 3. replace tc by tc' in both the rules and the induction theorem. | |
| 949 | *---------------------------------------------------------------------*) | |
| 950 | ||
| 951 | fun simplify_tc tc (r,ind) = | |
| 952 | let val tc1 = tych tc | |
| 14240 | 953 | val _ = trace_cterms "TC before simplification: " [tc1] | 
| 10769 | 954 | val tc_eq = simplifier tc1 | 
| 14240 | 955 | val _ = trace_thms "result: " [tc_eq] | 
| 10769 | 956 | in | 
| 957 | elim_tc (R.MATCH_MP Thms.eqT tc_eq) (r,ind) | |
| 958 | handle U.ERR _ => | |
| 959 | (elim_tc (R.MATCH_MP(R.MATCH_MP Thms.rev_eq_mp tc_eq) | |
| 11632 | 960 | (prove(tych(HOLogic.mk_Trueprop(S.rhs(concl tc_eq))), | 
| 10769 | 961 | terminator))) | 
| 962 | (r,ind) | |
| 963 | handle U.ERR _ => | |
| 964 | (R.UNDISCH(R.MATCH_MP (R.MATCH_MP Thms.simp_thm r) tc_eq), | |
| 965 | simplify_induction theory tc_eq ind)) | |
| 966 | end | |
| 967 | ||
| 968 | (*---------------------------------------------------------------------- | |
| 969 | * Nested termination conditions are harder to get at, since they are | |
| 970 | * left embedded in the body of the function (and in induction | |
| 971 | * theorem hypotheses). Our "solution" is to simplify them, and try to | |
| 972 | * prove termination, but leave the application of the resulting theorem | |
| 973 | * to a higher level. So things go much as in "simplify_tc": the | |
| 974 | * termination condition (tc) is simplified to |- tc = tc' (there might | |
| 975 | * not be a change) and then 2 attempts are made: | |
| 976 | * | |
| 977 | * 1. if |- tc = T, then return |- tc; otherwise, | |
| 978 | * 2. apply the terminator to tc'. If |- tc' = T then return |- tc; else | |
| 979 | * 3. return |- tc = tc' | |
| 980 | *---------------------------------------------------------------------*) | |
| 981 | fun simplify_nested_tc tc = | |
| 982 | let val tc_eq = simplifier (tych (#2 (S.strip_forall tc))) | |
| 983 | in | |
| 984 | R.GEN_ALL | |
| 985 | (R.MATCH_MP Thms.eqT tc_eq | |
| 986 | handle U.ERR _ => | |
| 987 | (R.MATCH_MP(R.MATCH_MP Thms.rev_eq_mp tc_eq) | |
| 11632 | 988 | (prove(tych(HOLogic.mk_Trueprop (S.rhs(concl tc_eq))), | 
| 10769 | 989 | terminator)) | 
| 990 | handle U.ERR _ => tc_eq)) | |
| 991 | end | |
| 992 | ||
| 993 | (*------------------------------------------------------------------- | |
| 994 | * Attempt to simplify the termination conditions in each rule and | |
| 995 | * in the induction theorem. | |
| 996 | *-------------------------------------------------------------------*) | |
| 997 | fun strip_imp tm = if S.is_neg tm then ([],tm) else S.strip_imp tm | |
| 998 | fun loop ([],extras,R,ind) = (rev R, ind, extras) | |
| 999 | | loop ((r,ftcs)::rst, nthms, R, ind) = | |
| 1000 | let val tcs = #1(strip_imp (concl r)) | |
| 1001 | val extra_tcs = gen_rems (op aconv) (ftcs, tcs) | |
| 1002 | val extra_tc_thms = map simplify_nested_tc extra_tcs | |
| 1003 | val (r1,ind1) = U.rev_itlist simplify_tc tcs (r,ind) | |
| 1004 | val r2 = R.FILTER_DISCH_ALL(not o S.is_WFR) r1 | |
| 1005 | in loop(rst, nthms@extra_tc_thms, r2::R, ind1) | |
| 1006 | end | |
| 1007 | val rules_tcs = ListPair.zip (R.CONJUNCTS rules1, TCs) | |
| 1008 | val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1) | |
| 1009 | in | |
| 1010 |   {induction = ind2, rules = R.LIST_CONJ rules2, nested_tcs = extras}
 | |
| 1011 | end; | |
| 1012 | ||
| 1013 | ||
| 1014 | end; |