| author | wenzelm | 
| Fri, 21 Apr 2023 15:30:59 +0200 | |
| changeset 77901 | 5728d5ebce34 | 
| parent 77879 | dd222e2af01a | 
| child 79336 | 032a31db4c6f | 
| permissions | -rw-r--r-- | 
| 62058 | 1 | (* Title: HOL/Library/old_recdef.ML | 
| 60520 | 2 | Author: Konrad Slind, Cambridge University Computer Laboratory | 
| 3 | Author: Lucas Dixon, University of Edinburgh | |
| 4 | ||
| 5 | Old TFL/recdef package. | |
| 6 | *) | |
| 7 | ||
| 8 | signature CASE_SPLIT = | |
| 9 | sig | |
| 10 | (* try to recursively split conjectured thm to given list of thms *) | |
| 11 | val splitto : Proof.context -> thm list -> thm -> thm | |
| 12 | end; | |
| 13 | ||
| 14 | signature UTILS = | |
| 15 | sig | |
| 16 |   exception ERR of {module: string, func: string, mesg: string}
 | |
| 17 |   val end_itlist: ('a -> 'a -> 'a) -> 'a list -> 'a
 | |
| 18 |   val itlist2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
 | |
| 19 |   val pluck: ('a -> bool) -> 'a list -> 'a * 'a list
 | |
| 20 |   val zip3: 'a list -> 'b list -> 'c list -> ('a*'b*'c) list
 | |
| 21 |   val take: ('a -> 'b) -> int * 'a list -> 'b list
 | |
| 22 | end; | |
| 23 | ||
| 24 | signature USYNTAX = | |
| 25 | sig | |
| 26 |   datatype lambda = VAR   of {Name : string, Ty : typ}
 | |
| 27 |                   | CONST of {Name : string, Ty : typ}
 | |
| 28 |                   | COMB  of {Rator: term, Rand : term}
 | |
| 29 |                   | LAMB  of {Bvar : term, Body : term}
 | |
| 30 | ||
| 31 | val alpha : typ | |
| 32 | ||
| 33 | (* Types *) | |
| 34 | val type_vars : typ -> typ list | |
| 35 | val type_varsl : typ list -> typ list | |
| 36 | val mk_vartype : string -> typ | |
| 37 | val is_vartype : typ -> bool | |
| 38 | val strip_prod_type : typ -> typ list | |
| 39 | ||
| 40 | (* Terms *) | |
| 41 | val free_vars_lr : term -> term list | |
| 42 | val type_vars_in_term : term -> typ list | |
| 43 | val dest_term : term -> lambda | |
| 44 | ||
| 45 | (* Prelogic *) | |
| 46 | val inst : (typ*typ) list -> term -> term | |
| 47 | ||
| 48 | (* Construction routines *) | |
| 49 |   val mk_abs    :{Bvar  : term, Body : term} -> term
 | |
| 50 | ||
| 51 |   val mk_imp    :{ant : term, conseq :  term} -> term
 | |
| 52 |   val mk_select :{Bvar : term, Body : term} -> term
 | |
| 53 |   val mk_forall :{Bvar : term, Body : term} -> term
 | |
| 54 |   val mk_exists :{Bvar : term, Body : term} -> term
 | |
| 55 |   val mk_conj   :{conj1 : term, conj2 : term} -> term
 | |
| 56 |   val mk_disj   :{disj1 : term, disj2 : term} -> term
 | |
| 57 |   val mk_pabs   :{varstruct : term, body : term} -> term
 | |
| 58 | ||
| 59 | (* Destruction routines *) | |
| 60 |   val dest_const: term -> {Name : string, Ty : typ}
 | |
| 61 |   val dest_comb : term -> {Rator : term, Rand : term}
 | |
| 62 |   val dest_abs  : string list -> term -> {Bvar : term, Body : term} * string list
 | |
| 63 |   val dest_eq     : term -> {lhs : term, rhs : term}
 | |
| 64 |   val dest_imp    : term -> {ant : term, conseq : term}
 | |
| 65 |   val dest_forall : term -> {Bvar : term, Body : term}
 | |
| 66 |   val dest_exists : term -> {Bvar : term, Body : term}
 | |
| 67 | val dest_neg : term -> term | |
| 68 |   val dest_conj   : term -> {conj1 : term, conj2 : term}
 | |
| 69 |   val dest_disj   : term -> {disj1 : term, disj2 : term}
 | |
| 70 |   val dest_pair   : term -> {fst : term, snd : term}
 | |
| 71 |   val dest_pabs   : string list -> term -> {varstruct : term, body : term, used : string list}
 | |
| 72 | ||
| 73 | val lhs : term -> term | |
| 74 | val rhs : term -> term | |
| 75 | val rand : term -> term | |
| 76 | ||
| 77 | (* Query routines *) | |
| 78 | val is_imp : term -> bool | |
| 79 | val is_forall : term -> bool | |
| 80 | val is_exists : term -> bool | |
| 81 | val is_neg : term -> bool | |
| 82 | val is_conj : term -> bool | |
| 83 | val is_disj : term -> bool | |
| 84 | val is_pair : term -> bool | |
| 85 | val is_pabs : term -> bool | |
| 86 | ||
| 87 | (* Construction of a term from a list of Preterms *) | |
| 88 | val list_mk_abs : (term list * term) -> term | |
| 89 | val list_mk_imp : (term list * term) -> term | |
| 90 | val list_mk_forall : (term list * term) -> term | |
| 91 | val list_mk_conj : term list -> term | |
| 92 | ||
| 93 | (* Destructing a term to a list of Preterms *) | |
| 94 | val strip_comb : term -> (term * term list) | |
| 95 | val strip_abs : term -> (term list * term) | |
| 96 | val strip_imp : term -> (term list * term) | |
| 97 | val strip_forall : term -> (term list * term) | |
| 98 | val strip_exists : term -> (term list * term) | |
| 99 | val strip_disj : term -> term list | |
| 100 | ||
| 101 | (* Miscellaneous *) | |
| 102 | val mk_vstruct : typ -> term list -> term | |
| 103 | val gen_all : term -> term | |
| 104 | val find_term : (term -> bool) -> term -> term option | |
| 105 | val dest_relation : term -> term * term * term | |
| 106 | val is_WFR : term -> bool | |
| 107 | val ARB : typ -> term | |
| 108 | end; | |
| 109 | ||
| 110 | signature DCTERM = | |
| 111 | sig | |
| 112 | val dest_comb: cterm -> cterm * cterm | |
| 74518 | 113 | val dest_abs: cterm -> cterm * cterm | 
| 60520 | 114 | val capply: cterm -> cterm -> cterm | 
| 115 | val cabs: cterm -> cterm -> cterm | |
| 116 | val mk_conj: cterm * cterm -> cterm | |
| 117 | val mk_disj: cterm * cterm -> cterm | |
| 118 | val mk_exists: cterm * cterm -> cterm | |
| 119 | val dest_conj: cterm -> cterm * cterm | |
| 120 |   val dest_const: cterm -> {Name: string, Ty: typ}
 | |
| 121 | val dest_disj: cterm -> cterm * cterm | |
| 122 | val dest_eq: cterm -> cterm * cterm | |
| 123 | val dest_exists: cterm -> cterm * cterm | |
| 124 | val dest_forall: cterm -> cterm * cterm | |
| 125 | val dest_imp: cterm -> cterm * cterm | |
| 126 | val dest_neg: cterm -> cterm | |
| 127 | val dest_pair: cterm -> cterm * cterm | |
| 128 |   val dest_var: cterm -> {Name:string, Ty:typ}
 | |
| 129 | val is_conj: cterm -> bool | |
| 130 | val is_disj: cterm -> bool | |
| 131 | val is_eq: cterm -> bool | |
| 132 | val is_exists: cterm -> bool | |
| 133 | val is_forall: cterm -> bool | |
| 134 | val is_imp: cterm -> bool | |
| 135 | val is_neg: cterm -> bool | |
| 136 | val is_pair: cterm -> bool | |
| 137 | val list_mk_disj: cterm list -> cterm | |
| 138 | val strip_abs: cterm -> cterm list * cterm | |
| 139 | val strip_comb: cterm -> cterm * cterm list | |
| 140 | val strip_disj: cterm -> cterm list | |
| 141 | val strip_exists: cterm -> cterm list * cterm | |
| 142 | val strip_forall: cterm -> cterm list * cterm | |
| 143 | val strip_imp: cterm -> cterm list * cterm | |
| 144 | val drop_prop: cterm -> cterm | |
| 145 | val mk_prop: cterm -> cterm | |
| 146 | end; | |
| 147 | ||
| 148 | signature RULES = | |
| 149 | sig | |
| 150 | val dest_thm: thm -> term list * term | |
| 151 | ||
| 152 | (* Inference rules *) | |
| 153 | val REFL: cterm -> thm | |
| 154 | val ASSUME: cterm -> thm | |
| 155 | val MP: thm -> thm -> thm | |
| 156 | val MATCH_MP: thm -> thm -> thm | |
| 157 | val CONJUNCT1: thm -> thm | |
| 158 | val CONJUNCT2: thm -> thm | |
| 159 | val CONJUNCTS: thm -> thm list | |
| 160 | val DISCH: cterm -> thm -> thm | |
| 161 | val UNDISCH: thm -> thm | |
| 162 | val SPEC: cterm -> thm -> thm | |
| 163 | val ISPEC: cterm -> thm -> thm | |
| 164 | val ISPECL: cterm list -> thm -> thm | |
| 165 | val GEN: Proof.context -> cterm -> thm -> thm | |
| 166 | val GENL: Proof.context -> cterm list -> thm -> thm | |
| 167 | val LIST_CONJ: thm list -> thm | |
| 168 | ||
| 169 | val SYM: thm -> thm | |
| 170 | val DISCH_ALL: thm -> thm | |
| 171 | val FILTER_DISCH_ALL: (term -> bool) -> thm -> thm | |
| 172 | val SPEC_ALL: thm -> thm | |
| 173 | val GEN_ALL: Proof.context -> thm -> thm | |
| 174 | val IMP_TRANS: thm -> thm -> thm | |
| 175 | val PROVE_HYP: thm -> thm -> thm | |
| 176 | ||
| 177 | val CHOOSE: Proof.context -> cterm * thm -> thm -> thm | |
| 60781 | 178 | val EXISTS: Proof.context -> cterm * cterm -> thm -> thm | 
| 60520 | 179 | val IT_EXISTS: Proof.context -> (cterm * cterm) list -> thm -> thm | 
| 180 | ||
| 181 | val EVEN_ORS: thm list -> thm list | |
| 182 | val DISJ_CASESL: thm -> thm list -> thm | |
| 183 | ||
| 184 | val list_beta_conv: cterm -> cterm list -> thm | |
| 185 | val SUBS: Proof.context -> thm list -> thm -> thm | |
| 186 | val simpl_conv: Proof.context -> thm list -> cterm -> thm | |
| 187 | ||
| 188 | val rbeta: thm -> thm | |
| 189 | val tracing: bool Unsynchronized.ref | |
| 190 | val CONTEXT_REWRITE_RULE: Proof.context -> | |
| 191 | term * term list * thm * thm list -> thm -> thm * term list | |
| 192 | val RIGHT_ASSOC: Proof.context -> thm -> thm | |
| 193 | ||
| 74532 | 194 | val prove: Proof.context -> bool -> term -> (Proof.context -> tactic) -> thm | 
| 60520 | 195 | end; | 
| 196 | ||
| 197 | signature THRY = | |
| 198 | sig | |
| 199 | val match_term: theory -> term -> term -> (term * term) list * (typ * typ) list | |
| 200 | val match_type: theory -> typ -> typ -> (typ * typ) list | |
| 201 | val typecheck: theory -> term -> cterm | |
| 202 | (*datatype facts of various flavours*) | |
| 203 |   val match_info: theory -> string -> {constructors: term list, case_const: term} option
 | |
| 204 |   val induct_info: theory -> string -> {constructors: term list, nchotomy: thm} option
 | |
| 205 |   val extract_info: theory -> {case_congs: thm list, case_rewrites: thm list}
 | |
| 206 | end; | |
| 207 | ||
| 208 | signature PRIM = | |
| 209 | sig | |
| 210 | val trace: bool Unsynchronized.ref | |
| 211 | val trace_thms: Proof.context -> string -> thm list -> unit | |
| 212 | val trace_cterm: Proof.context -> string -> cterm -> unit | |
| 213 | type pattern | |
| 214 |   val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
 | |
| 215 | val wfrec_definition0: string -> term -> term -> theory -> thm * theory | |
| 216 | val post_definition: Proof.context -> thm list -> thm * pattern list -> | |
| 217 |    {rules: thm,
 | |
| 218 | rows: int list, | |
| 219 | TCs: term list list, | |
| 220 | full_pats_TCs: (term * term list) list} | |
| 60699 | 221 | val mk_induction: Proof.context -> | 
| 60520 | 222 |     {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
 | 
| 223 | val postprocess: Proof.context -> bool -> | |
| 74532 | 224 |     {wf_tac: Proof.context -> tactic,
 | 
| 225 | terminator: Proof.context -> tactic, | |
| 226 | simplifier: Proof.context -> cterm -> thm} -> | |
| 60520 | 227 |     {rules: thm, induction: thm, TCs: term list list} ->
 | 
| 228 |     {rules: thm, induction: thm, nested_tcs: thm list}
 | |
| 229 | end; | |
| 230 | ||
| 231 | signature TFL = | |
| 232 | sig | |
| 233 | val define_i: bool -> thm list -> thm list -> xstring -> term -> term list -> Proof.context -> | |
| 234 |     {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | |
| 235 | val define: bool -> thm list -> thm list -> xstring -> string -> string list -> Proof.context -> | |
| 236 |     {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | |
| 237 | end; | |
| 238 | ||
| 239 | signature OLD_RECDEF = | |
| 240 | sig | |
| 241 | val get_recdef: theory -> string | |
| 242 |     -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
 | |
| 243 |   val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
 | |
| 244 | val simp_add: attribute | |
| 245 | val simp_del: attribute | |
| 246 | val cong_add: attribute | |
| 247 | val cong_del: attribute | |
| 248 | val wf_add: attribute | |
| 249 | val wf_del: attribute | |
| 250 | val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list -> | |
| 251 | Token.src option -> theory -> theory | |
| 252 |       * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | |
| 253 | val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list -> | |
| 254 |     theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | |
| 255 | end; | |
| 256 | ||
| 257 | structure Old_Recdef: OLD_RECDEF = | |
| 258 | struct | |
| 259 | ||
| 260 | (*** extra case splitting for TFL ***) | |
| 261 | ||
| 262 | structure CaseSplit: CASE_SPLIT = | |
| 263 | struct | |
| 264 | ||
| 265 | (* make a casethm from an induction thm *) | |
| 60752 | 266 | fun cases_thm_of_induct_thm ctxt = | 
| 267 | Seq.hd o (ALLGOALS (fn i => REPEAT (eresolve_tac ctxt [Drule.thin_rl] i))); | |
| 60520 | 268 | |
| 269 | (* get the case_thm (my version) from a type *) | |
| 60752 | 270 | fun case_thm_of_ty ctxt ty = | 
| 60520 | 271 | let | 
| 60752 | 272 | val thy = Proof_Context.theory_of ctxt | 
| 60520 | 273 | val ty_str = case ty of | 
| 274 | Type(ty_str, _) => ty_str | |
| 275 |                    | TFree(s,_)  => error ("Free type: " ^ s)
 | |
| 60521 | 276 |                    | TVar((s,_),_) => error ("Free variable: " ^ s)
 | 
| 60520 | 277 |       val {induct, ...} = BNF_LFP_Compat.the_info thy [BNF_LFP_Compat.Keep_Nesting] ty_str
 | 
| 278 | in | |
| 60752 | 279 | cases_thm_of_induct_thm ctxt induct | 
| 60520 | 280 | end; | 
| 281 | ||
| 282 | ||
| 283 | (* for use when there are no prems to the subgoal *) | |
| 284 | (* does a case split on the given variable *) | |
| 285 | fun mk_casesplit_goal_thm ctxt (vstr,ty) gt = | |
| 286 | let | |
| 287 | val thy = Proof_Context.theory_of ctxt; | |
| 288 | ||
| 289 | val x = Free(vstr,ty); | |
| 290 | val abst = Abs(vstr, ty, Term.abstract_over (x, gt)); | |
| 291 | ||
| 60752 | 292 | val case_thm = case_thm_of_ty ctxt ty; | 
| 60520 | 293 | |
| 294 | val abs_ct = Thm.cterm_of ctxt abst; | |
| 295 | val free_ct = Thm.cterm_of ctxt x; | |
| 296 | ||
| 297 | val (Pv, Dv, type_insts) = | |
| 298 | case (Thm.concl_of case_thm) of | |
| 60521 | 299 | (_ $ (Pv $ (Dv as Var(_, Dty)))) => | 
| 60520 | 300 | (Pv, Dv, | 
| 301 | Sign.typ_match thy (Dty, ty) Vartab.empty) | |
| 302 | | _ => error "not a valid case thm"; | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 303 | val type_cinsts = map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T)) | 
| 60520 | 304 | (Vartab.dest type_insts); | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 305 | val Pv = dest_Var (Envir.subst_term_types type_insts Pv); | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 306 | val Dv = dest_Var (Envir.subst_term_types type_insts Dv); | 
| 60520 | 307 | in | 
| 308 | Conv.fconv_rule Drule.beta_eta_conversion | |
| 309 | (case_thm | |
| 74282 | 310 | |> Thm.instantiate (TVars.make type_cinsts, Vars.empty) | 
| 77879 | 311 | |> Thm.instantiate (TVars.empty, Vars.make2 (Pv, abs_ct) (Dv, free_ct))) | 
| 60520 | 312 | end; | 
| 313 | ||
| 314 | ||
| 315 | (* the find_XXX_split functions are simply doing a lightwieght (I | |
| 316 | think) term matching equivalent to find where to do the next split *) | |
| 317 | ||
| 318 | (* assuming two twems are identical except for a free in one at a | |
| 319 | subterm, or constant in another, ie assume that one term is a plit of | |
| 320 | another, then gives back the free variable that has been split. *) | |
| 321 | exception find_split_exp of string | |
| 322 | fun find_term_split (Free v, _ $ _) = SOME v | |
| 323 | | find_term_split (Free v, Const _) = SOME v | |
| 324 | | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *) | |
| 60521 | 325 | | find_term_split (Free _, Var _) = NONE (* keep searching *) | 
| 60520 | 326 | | find_term_split (a $ b, a2 $ b2) = | 
| 327 | (case find_term_split (a, a2) of | |
| 328 | NONE => find_term_split (b,b2) | |
| 329 | | vopt => vopt) | |
| 60521 | 330 | | find_term_split (Abs(_,_,t1), Abs(_,_,t2)) = | 
| 60520 | 331 | find_term_split (t1, t2) | 
| 60521 | 332 | | find_term_split (Const (x,_), Const(x2,_)) = | 
| 60520 | 333 | if x = x2 then NONE else (* keep searching *) | 
| 334 | raise find_split_exp (* stop now *) | |
| 335 | "Terms are not identical upto a free varaible! (Consts)" | |
| 336 | | find_term_split (Bound i, Bound j) = | |
| 337 | if i = j then NONE else (* keep searching *) | |
| 338 | raise find_split_exp (* stop now *) | |
| 339 | "Terms are not identical upto a free varaible! (Bound)" | |
| 340 | | find_term_split _ = | |
| 341 | raise find_split_exp (* stop now *) | |
| 342 | "Terms are not identical upto a free varaible! (Other)"; | |
| 343 | ||
| 344 | (* assume that "splitth" is a case split form of subgoal i of "genth", | |
| 345 | then look for a free variable to split, breaking the subgoal closer to | |
| 346 | splitth. *) | |
| 347 | fun find_thm_split splitth i genth = | |
| 348 | find_term_split (Logic.get_goal (Thm.prop_of genth) i, | |
| 349 | Thm.concl_of splitth) handle find_split_exp _ => NONE; | |
| 350 | ||
| 351 | (* as above but searches "splitths" for a theorem that suggest a case split *) | |
| 352 | fun find_thms_split splitths i genth = | |
| 353 | Library.get_first (fn sth => find_thm_split sth i genth) splitths; | |
| 354 | ||
| 355 | ||
| 356 | (* split the subgoal i of "genth" until we get to a member of | |
| 357 | splitths. Assumes that genth will be a general form of splitths, that | |
| 358 | can be case-split, as needed. Otherwise fails. Note: We assume that | |
| 359 | all of "splitths" are split to the same level, and thus it doesn't | |
| 360 | matter which one we choose to look for the next split. Simply add | |
| 361 | search on splitthms and split variable, to change this. *) | |
| 362 | (* Note: possible efficiency measure: when a case theorem is no longer | |
| 363 | useful, drop it? *) | |
| 364 | (* Note: This should not be a separate tactic but integrated into the | |
| 365 | case split done during recdef's case analysis, this would avoid us | |
| 366 | having to (re)search for variables to split. *) | |
| 367 | fun splitto ctxt splitths genth = | |
| 368 | let | |
| 369 | val _ = not (null splitths) orelse error "splitto: no given splitths"; | |
| 370 | ||
| 371 | (* check if we are a member of splitths - FIXME: quicker and | |
| 372 | more flexible with discrim net. *) | |
| 373 | fun solve_by_splitth th split = | |
| 374 | Thm.biresolution (SOME ctxt) false [(false,split)] 1 th; | |
| 375 | ||
| 376 | fun split th = | |
| 377 | (case find_thms_split splitths 1 th of | |
| 378 | NONE => | |
| 379 | (writeln (cat_lines | |
| 61268 | 380 | (["th:", Thm.string_of_thm ctxt th, "split ths:"] @ | 
| 381 | map (Thm.string_of_thm ctxt) splitths @ ["\n--"])); | |
| 60520 | 382 | error "splitto: cannot find variable to split on") | 
| 383 | | SOME v => | |
| 384 | let | |
| 385 | val gt = HOLogic.dest_Trueprop (#1 (Logic.dest_implies (Thm.prop_of th))); | |
| 386 | val split_thm = mk_casesplit_goal_thm ctxt v gt; | |
| 387 | val (subthms, expf) = IsaND.fixed_subgoal_thms ctxt split_thm; | |
| 388 | in | |
| 389 | expf (map recsplitf subthms) | |
| 390 | end) | |
| 391 | ||
| 392 | and recsplitf th = | |
| 393 | (* note: multiple unifiers! we only take the first element, | |
| 394 | probably fine -- there is probably only one anyway. *) | |
| 395 | (case get_first (Seq.pull o solve_by_splitth th) splitths of | |
| 396 | NONE => split th | |
| 397 | | SOME (solved_th, _) => solved_th); | |
| 398 | in | |
| 399 | recsplitf genth | |
| 400 | end; | |
| 401 | ||
| 402 | end; | |
| 403 | ||
| 404 | ||
| 60521 | 405 | |
| 60520 | 406 | (*** basic utilities ***) | 
| 407 | ||
| 408 | structure Utils: UTILS = | |
| 409 | struct | |
| 410 | ||
| 411 | (*standard exception for TFL*) | |
| 412 | exception ERR of {module: string, func: string, mesg: string};
 | |
| 413 | ||
| 414 | fun UTILS_ERR func mesg = ERR {module = "Utils", func = func, mesg = mesg};
 | |
| 415 | ||
| 416 | ||
| 60521 | 417 | fun end_itlist _ [] = raise (UTILS_ERR "end_itlist" "list too short") | 
| 418 | | end_itlist _ [x] = x | |
| 60520 | 419 | | end_itlist f (x :: xs) = f x (end_itlist f xs); | 
| 420 | ||
| 421 | fun itlist2 f L1 L2 base_value = | |
| 422 | let fun it ([],[]) = base_value | |
| 423 | | it ((a::rst1),(b::rst2)) = f a b (it (rst1,rst2)) | |
| 424 | | it _ = raise UTILS_ERR "itlist2" "different length lists" | |
| 425 | in it (L1,L2) | |
| 426 | end; | |
| 427 | ||
| 428 | fun pluck p = | |
| 429 | let fun remv ([],_) = raise UTILS_ERR "pluck" "item not found" | |
| 430 | | remv (h::t, A) = if p h then (h, rev A @ t) else remv (t,h::A) | |
| 431 | in fn L => remv(L,[]) | |
| 432 | end; | |
| 433 | ||
| 434 | fun take f = | |
| 60521 | 435 | let fun grab(0, _) = [] | 
| 60520 | 436 | | grab(n, x::rst) = f x::grab(n-1,rst) | 
| 437 | in grab | |
| 438 | end; | |
| 439 | ||
| 440 | fun zip3 [][][] = [] | |
| 441 | | zip3 (x::l1) (y::l2) (z::l3) = (x,y,z)::zip3 l1 l2 l3 | |
| 442 | | zip3 _ _ _ = raise UTILS_ERR "zip3" "different lengths"; | |
| 443 | ||
| 444 | ||
| 445 | end; | |
| 446 | ||
| 447 | ||
| 60521 | 448 | |
| 60520 | 449 | (*** emulation of HOL's abstract syntax functions ***) | 
| 450 | ||
| 451 | structure USyntax: USYNTAX = | |
| 452 | struct | |
| 453 | ||
| 454 | infix 4 ##; | |
| 455 | ||
| 456 | fun USYN_ERR func mesg = Utils.ERR {module = "USyntax", func = func, mesg = mesg};
 | |
| 457 | ||
| 458 | ||
| 459 | (*--------------------------------------------------------------------------- | |
| 460 | * | |
| 461 | * Types | |
| 462 | * | |
| 463 | *---------------------------------------------------------------------------*) | |
| 464 | val mk_prim_vartype = TVar; | |
| 69593 | 465 | fun mk_vartype s = mk_prim_vartype ((s, 0), \<^sort>\<open>type\<close>); | 
| 60520 | 466 | |
| 467 | (* But internally, it's useful *) | |
| 468 | fun dest_vtype (TVar x) = x | |
| 469 | | dest_vtype _ = raise USYN_ERR "dest_vtype" "not a flexible type variable"; | |
| 470 | ||
| 471 | val is_vartype = can dest_vtype; | |
| 472 | ||
| 473 | val type_vars = map mk_prim_vartype o Misc_Legacy.typ_tvars | |
| 474 | fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []); | |
| 475 | ||
| 476 | val alpha = mk_vartype "'a" | |
| 477 | ||
| 478 | val strip_prod_type = HOLogic.flatten_tupleT; | |
| 479 | ||
| 480 | ||
| 481 | ||
| 482 | (*--------------------------------------------------------------------------- | |
| 483 | * | |
| 484 | * Terms | |
| 485 | * | |
| 486 | *---------------------------------------------------------------------------*) | |
| 487 | ||
| 488 | (* Free variables, in order of occurrence, from left to right in the | |
| 489 | * syntax tree. *) | |
| 490 | fun free_vars_lr tm = | |
| 491 | let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end | |
| 492 | fun add (t, frees) = case t of | |
| 493 | Free _ => if (memb t frees) then frees else t::frees | |
| 494 | | Abs (_,_,body) => add(body,frees) | |
| 495 | | f$t => add(t, add(f, frees)) | |
| 496 | | _ => frees | |
| 497 | in rev(add(tm,[])) | |
| 498 | end; | |
| 499 | ||
| 500 | ||
| 501 | ||
| 502 | val type_vars_in_term = map mk_prim_vartype o Misc_Legacy.term_tvars; | |
| 503 | ||
| 504 | ||
| 505 | ||
| 506 | (* Prelogic *) | |
| 507 | fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty) | |
| 508 | fun inst theta = subst_vars (map dest_tybinding theta,[]) | |
| 509 | ||
| 510 | ||
| 511 | (* Construction routines *) | |
| 512 | ||
| 513 | fun mk_abs{Bvar as Var((s,_),ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | |
| 514 |   | mk_abs{Bvar as Free(s,ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | |
| 515 | | mk_abs _ = raise USYN_ERR "mk_abs" "Bvar is not a variable"; | |
| 516 | ||
| 517 | ||
| 518 | fun mk_imp{ant,conseq} =
 | |
| 69593 | 519 | let val c = Const(\<^const_name>\<open>HOL.implies\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) | 
| 60520 | 520 | in list_comb(c,[ant,conseq]) | 
| 521 | end; | |
| 522 | ||
| 523 | fun mk_select (r as {Bvar,Body}) =
 | |
| 524 | let val ty = type_of Bvar | |
| 69593 | 525 | val c = Const(\<^const_name>\<open>Eps\<close>,(ty --> HOLogic.boolT) --> ty) | 
| 60520 | 526 | in list_comb(c,[mk_abs r]) | 
| 527 | end; | |
| 528 | ||
| 529 | fun mk_forall (r as {Bvar,Body}) =
 | |
| 530 | let val ty = type_of Bvar | |
| 69593 | 531 | val c = Const(\<^const_name>\<open>All\<close>,(ty --> HOLogic.boolT) --> HOLogic.boolT) | 
| 60520 | 532 | in list_comb(c,[mk_abs r]) | 
| 533 | end; | |
| 534 | ||
| 535 | fun mk_exists (r as {Bvar,Body}) =
 | |
| 536 | let val ty = type_of Bvar | |
| 69593 | 537 | val c = Const(\<^const_name>\<open>Ex\<close>,(ty --> HOLogic.boolT) --> HOLogic.boolT) | 
| 60520 | 538 | in list_comb(c,[mk_abs r]) | 
| 539 | end; | |
| 540 | ||
| 541 | ||
| 542 | fun mk_conj{conj1,conj2} =
 | |
| 69593 | 543 | let val c = Const(\<^const_name>\<open>HOL.conj\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) | 
| 60520 | 544 | in list_comb(c,[conj1,conj2]) | 
| 545 | end; | |
| 546 | ||
| 547 | fun mk_disj{disj1,disj2} =
 | |
| 69593 | 548 | let val c = Const(\<^const_name>\<open>HOL.disj\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) | 
| 60520 | 549 | in list_comb(c,[disj1,disj2]) | 
| 550 | end; | |
| 551 | ||
| 552 | fun prod_ty ty1 ty2 = HOLogic.mk_prodT (ty1,ty2); | |
| 553 | ||
| 554 | local | |
| 555 | fun mk_uncurry (xt, yt, zt) = | |
| 69593 | 556 | Const(\<^const_name>\<open>case_prod\<close>, (xt --> yt --> zt) --> prod_ty xt yt --> zt) | 
| 557 | fun dest_pair(Const(\<^const_name>\<open>Pair\<close>,_) $ M $ N) = {fst=M, snd=N}
 | |
| 60520 | 558 | | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair" | 
| 559 | fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false | |
| 560 | in | |
| 561 | fun mk_pabs{varstruct,body} =
 | |
| 562 | let fun mpa (varstruct, body) = | |
| 563 | if is_var varstruct | |
| 564 |        then mk_abs {Bvar = varstruct, Body = body}
 | |
| 565 |        else let val {fst, snd} = dest_pair varstruct
 | |
| 566 | in mk_uncurry (type_of fst, type_of snd, type_of body) $ | |
| 567 | mpa (fst, mpa (snd, body)) | |
| 568 | end | |
| 569 | in mpa (varstruct, body) end | |
| 570 | handle TYPE _ => raise USYN_ERR "mk_pabs" ""; | |
| 571 | end; | |
| 572 | ||
| 573 | (* Destruction routines *) | |
| 574 | ||
| 575 | datatype lambda = VAR   of {Name : string, Ty : typ}
 | |
| 576 |                 | CONST of {Name : string, Ty : typ}
 | |
| 577 |                 | COMB  of {Rator: term, Rand : term}
 | |
| 578 |                 | LAMB  of {Bvar : term, Body : term};
 | |
| 579 | ||
| 580 | ||
| 60521 | 581 | fun dest_term(Var((s,_),ty)) = VAR{Name = s, Ty = ty}
 | 
| 60520 | 582 |   | dest_term(Free(s,ty))    = VAR{Name = s, Ty = ty}
 | 
| 583 |   | dest_term(Const(s,ty))   = CONST{Name = s, Ty = ty}
 | |
| 584 |   | dest_term(M$N)           = COMB{Rator=M,Rand=N}
 | |
| 585 | | dest_term(Abs(s,ty,M)) = let val v = Free(s,ty) | |
| 586 |                                in LAMB{Bvar = v, Body = Term.betapply (M,v)}
 | |
| 587 | end | |
| 588 | | dest_term(Bound _) = raise USYN_ERR "dest_term" "Bound"; | |
| 589 | ||
| 590 | fun dest_const(Const(s,ty)) = {Name = s, Ty = ty}
 | |
| 591 | | dest_const _ = raise USYN_ERR "dest_const" "not a constant"; | |
| 592 | ||
| 593 | fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2}
 | |
| 594 | | dest_comb _ = raise USYN_ERR "dest_comb" "not a comb"; | |
| 595 | ||
| 60524 | 596 | fun dest_abs used (a as Abs(s, ty, _)) = | 
| 60520 | 597 | let | 
| 598 | val s' = singleton (Name.variant_list used) s; | |
| 599 | val v = Free(s', ty); | |
| 600 |      in ({Bvar = v, Body = Term.betapply (a,v)}, s'::used)
 | |
| 601 | end | |
| 602 | | dest_abs _ _ = raise USYN_ERR "dest_abs" "not an abstraction"; | |
| 603 | ||
| 69593 | 604 | fun dest_eq(Const(\<^const_name>\<open>HOL.eq\<close>,_) $ M $ N) = {lhs=M, rhs=N}
 | 
| 60520 | 605 | | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality"; | 
| 606 | ||
| 69593 | 607 | fun dest_imp(Const(\<^const_name>\<open>HOL.implies\<close>,_) $ M $ N) = {ant=M, conseq=N}
 | 
| 60520 | 608 | | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication"; | 
| 609 | ||
| 69593 | 610 | fun dest_forall(Const(\<^const_name>\<open>All\<close>,_) $ (a as Abs _)) = fst (dest_abs [] a) | 
| 60520 | 611 | | dest_forall _ = raise USYN_ERR "dest_forall" "not a forall"; | 
| 612 | ||
| 69593 | 613 | fun dest_exists(Const(\<^const_name>\<open>Ex\<close>,_) $ (a as Abs _)) = fst (dest_abs [] a) | 
| 60520 | 614 | | dest_exists _ = raise USYN_ERR "dest_exists" "not an existential"; | 
| 615 | ||
| 69593 | 616 | fun dest_neg(Const(\<^const_name>\<open>Not\<close>,_) $ M) = M | 
| 60520 | 617 | | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation"; | 
| 618 | ||
| 69593 | 619 | fun dest_conj(Const(\<^const_name>\<open>HOL.conj\<close>,_) $ M $ N) = {conj1=M, conj2=N}
 | 
| 60520 | 620 | | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction"; | 
| 621 | ||
| 69593 | 622 | fun dest_disj(Const(\<^const_name>\<open>HOL.disj\<close>,_) $ M $ N) = {disj1=M, disj2=N}
 | 
| 60520 | 623 | | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction"; | 
| 624 | ||
| 625 | fun mk_pair{fst,snd} =
 | |
| 626 | let val ty1 = type_of fst | |
| 627 | val ty2 = type_of snd | |
| 69593 | 628 | val c = Const(\<^const_name>\<open>Pair\<close>,ty1 --> ty2 --> prod_ty ty1 ty2) | 
| 60520 | 629 | in list_comb(c,[fst,snd]) | 
| 630 | end; | |
| 631 | ||
| 69593 | 632 | fun dest_pair(Const(\<^const_name>\<open>Pair\<close>,_) $ M $ N) = {fst=M, snd=N}
 | 
| 60520 | 633 | | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"; | 
| 634 | ||
| 635 | ||
| 69593 | 636 | local fun ucheck t = (if #Name (dest_const t) = \<^const_name>\<open>case_prod\<close> then t | 
| 60520 | 637 | else raise Match) | 
| 638 | in | |
| 639 | fun dest_pabs used tm = | |
| 640 |    let val ({Bvar,Body}, used') = dest_abs used tm
 | |
| 641 |    in {varstruct = Bvar, body = Body, used = used'}
 | |
| 642 | end handle Utils.ERR _ => | |
| 643 |           let val {Rator,Rand} = dest_comb tm
 | |
| 644 | val _ = ucheck Rator | |
| 645 |               val {varstruct = lv, body, used = used'} = dest_pabs used Rand
 | |
| 646 |               val {varstruct = rv, body, used = used''} = dest_pabs used' body
 | |
| 647 |           in {varstruct = mk_pair {fst = lv, snd = rv}, body = body, used = used''}
 | |
| 648 | end | |
| 649 | end; | |
| 650 | ||
| 651 | ||
| 652 | val lhs = #lhs o dest_eq | |
| 653 | val rhs = #rhs o dest_eq | |
| 654 | val rand = #Rand o dest_comb | |
| 655 | ||
| 656 | ||
| 657 | (* Query routines *) | |
| 658 | val is_imp = can dest_imp | |
| 659 | val is_forall = can dest_forall | |
| 660 | val is_exists = can dest_exists | |
| 661 | val is_neg = can dest_neg | |
| 662 | val is_conj = can dest_conj | |
| 663 | val is_disj = can dest_disj | |
| 664 | val is_pair = can dest_pair | |
| 665 | val is_pabs = can (dest_pabs []) | |
| 666 | ||
| 667 | ||
| 668 | (* Construction of a cterm from a list of Terms *) | |
| 669 | ||
| 670 | fun list_mk_abs(L,tm) = fold_rev (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm;
 | |
| 671 | ||
| 672 | (* These others are almost never used *) | |
| 673 | fun list_mk_imp(A,c) = fold_rev (fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c;
 | |
| 674 | fun list_mk_forall(V,t) = fold_rev (fn v => fn b => mk_forall{Bvar=v, Body=b})V t;
 | |
| 675 | val list_mk_conj = Utils.end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm})
 | |
| 676 | ||
| 677 | ||
| 678 | (* Need to reverse? *) | |
| 679 | fun gen_all tm = list_mk_forall(Misc_Legacy.term_frees tm, tm); | |
| 680 | ||
| 681 | (* Destructing a cterm to a list of Terms *) | |
| 682 | fun strip_comb tm = | |
| 683 | let fun dest(M$N, A) = dest(M, N::A) | |
| 684 | | dest x = x | |
| 685 | in dest(tm,[]) | |
| 686 | end; | |
| 687 | ||
| 688 | fun strip_abs(tm as Abs _) = | |
| 689 |        let val ({Bvar,Body}, _) = dest_abs [] tm
 | |
| 690 | val (bvs, core) = strip_abs Body | |
| 691 | in (Bvar::bvs, core) | |
| 692 | end | |
| 693 | | strip_abs M = ([],M); | |
| 694 | ||
| 695 | ||
| 696 | fun strip_imp fm = | |
| 697 | if (is_imp fm) | |
| 698 |    then let val {ant,conseq} = dest_imp fm
 | |
| 699 | val (was,wb) = strip_imp conseq | |
| 700 | in ((ant::was), wb) | |
| 701 | end | |
| 702 | else ([],fm); | |
| 703 | ||
| 704 | fun strip_forall fm = | |
| 705 | if (is_forall fm) | |
| 706 |    then let val {Bvar,Body} = dest_forall fm
 | |
| 707 | val (bvs,core) = strip_forall Body | |
| 708 | in ((Bvar::bvs), core) | |
| 709 | end | |
| 710 | else ([],fm); | |
| 711 | ||
| 712 | ||
| 713 | fun strip_exists fm = | |
| 714 | if (is_exists fm) | |
| 715 |    then let val {Bvar, Body} = dest_exists fm
 | |
| 716 | val (bvs,core) = strip_exists Body | |
| 717 | in (Bvar::bvs, core) | |
| 718 | end | |
| 719 | else ([],fm); | |
| 720 | ||
| 721 | fun strip_disj w = | |
| 722 | if (is_disj w) | |
| 723 |    then let val {disj1,disj2} = dest_disj w
 | |
| 724 | in (strip_disj disj1@strip_disj disj2) | |
| 725 | end | |
| 726 | else [w]; | |
| 727 | ||
| 728 | ||
| 729 | (* Miscellaneous *) | |
| 730 | ||
| 731 | fun mk_vstruct ty V = | |
| 69593 | 732 | let fun follow_prod_type (Type(\<^type_name>\<open>Product_Type.prod\<close>,[ty1,ty2])) vs = | 
| 60520 | 733 | let val (ltm,vs1) = follow_prod_type ty1 vs | 
| 734 | val (rtm,vs2) = follow_prod_type ty2 vs1 | |
| 735 |               in (mk_pair{fst=ltm, snd=rtm}, vs2) end
 | |
| 736 | | follow_prod_type _ (v::vs) = (v,vs) | |
| 737 | in #1 (follow_prod_type ty V) end; | |
| 738 | ||
| 739 | ||
| 740 | (* Search a term for a sub-term satisfying the predicate p. *) | |
| 741 | fun find_term p = | |
| 742 | let fun find tm = | |
| 743 | if (p tm) then SOME tm | |
| 744 | else case tm of | |
| 745 | Abs(_,_,body) => find body | |
| 746 | | (t$u) => (case find t of NONE => find u | some => some) | |
| 747 | | _ => NONE | |
| 748 | in find | |
| 749 | end; | |
| 750 | ||
| 751 | fun dest_relation tm = | |
| 752 | if (type_of tm = HOLogic.boolT) | |
| 69593 | 753 | then let val (Const(\<^const_name>\<open>Set.member\<close>,_) $ (Const(\<^const_name>\<open>Pair\<close>,_)$y$x) $ R) = tm | 
| 60520 | 754 | in (R,y,x) | 
| 755 | end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure" | |
| 756 | else raise USYN_ERR "dest_relation" "not a boolean term"; | |
| 757 | ||
| 69593 | 758 | fun is_WFR (Const(\<^const_name>\<open>Wellfounded.wf\<close>,_)$_) = true | 
| 60520 | 759 | | is_WFR _ = false; | 
| 760 | ||
| 761 | fun ARB ty = mk_select{Bvar=Free("v",ty),
 | |
| 69593 | 762 | Body=Const(\<^const_name>\<open>True\<close>,HOLogic.boolT)}; | 
| 60520 | 763 | |
| 764 | end; | |
| 765 | ||
| 766 | ||
| 60521 | 767 | |
| 60520 | 768 | (*** derived cterm destructors ***) | 
| 769 | ||
| 770 | structure Dcterm: DCTERM = | |
| 771 | struct | |
| 772 | ||
| 773 | fun ERR func mesg = Utils.ERR {module = "Dcterm", func = func, mesg = mesg};
 | |
| 774 | ||
| 775 | ||
| 776 | fun dest_comb t = Thm.dest_comb t | |
| 777 | handle CTERM (msg, _) => raise ERR "dest_comb" msg; | |
| 778 | ||
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 779 | fun dest_abs t = Thm.dest_abs_global t | 
| 60520 | 780 | handle CTERM (msg, _) => raise ERR "dest_abs" msg; | 
| 781 | ||
| 782 | fun capply t u = Thm.apply t u | |
| 783 | handle CTERM (msg, _) => raise ERR "capply" msg; | |
| 784 | ||
| 785 | fun cabs a t = Thm.lambda a t | |
| 786 | handle CTERM (msg, _) => raise ERR "cabs" msg; | |
| 787 | ||
| 788 | ||
| 789 | (*--------------------------------------------------------------------------- | |
| 790 | * Some simple constructor functions. | |
| 791 | *---------------------------------------------------------------------------*) | |
| 792 | ||
| 69593 | 793 | val mk_hol_const = Thm.cterm_of \<^theory_context>\<open>HOL\<close> o Const; | 
| 60520 | 794 | |
| 795 | fun mk_exists (r as (Bvar, Body)) = | |
| 796 | let val ty = Thm.typ_of_cterm Bvar | |
| 69593 | 797 | val c = mk_hol_const(\<^const_name>\<open>Ex\<close>, (ty --> HOLogic.boolT) --> HOLogic.boolT) | 
| 60520 | 798 | in capply c (uncurry cabs r) end; | 
| 799 | ||
| 800 | ||
| 69593 | 801 | local val c = mk_hol_const(\<^const_name>\<open>HOL.conj\<close>, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) | 
| 60520 | 802 | in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2 | 
| 803 | end; | |
| 804 | ||
| 69593 | 805 | local val c = mk_hol_const(\<^const_name>\<open>HOL.disj\<close>, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) | 
| 60520 | 806 | in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2 | 
| 807 | end; | |
| 808 | ||
| 809 | ||
| 810 | (*--------------------------------------------------------------------------- | |
| 811 | * The primitives. | |
| 812 | *---------------------------------------------------------------------------*) | |
| 813 | fun dest_const ctm = | |
| 814 | (case Thm.term_of ctm | |
| 815 |       of Const(s,ty) => {Name = s, Ty = ty}
 | |
| 816 | | _ => raise ERR "dest_const" "not a constant"); | |
| 817 | ||
| 818 | fun dest_var ctm = | |
| 819 | (case Thm.term_of ctm | |
| 60521 | 820 |       of Var((s,_),ty) => {Name=s, Ty=ty}
 | 
| 60520 | 821 |        | Free(s,ty)    => {Name=s, Ty=ty}
 | 
| 822 | | _ => raise ERR "dest_var" "not a variable"); | |
| 823 | ||
| 824 | ||
| 825 | (*--------------------------------------------------------------------------- | |
| 826 | * Derived destructor operations. | |
| 827 | *---------------------------------------------------------------------------*) | |
| 828 | ||
| 829 | fun dest_monop expected tm = | |
| 830 | let | |
| 831 |    fun err () = raise ERR "dest_monop" ("Not a(n) " ^ quote expected);
 | |
| 832 | val (c, N) = dest_comb tm handle Utils.ERR _ => err (); | |
| 833 | val name = #Name (dest_const c handle Utils.ERR _ => err ()); | |
| 834 | in if name = expected then N else err () end; | |
| 835 | ||
| 836 | fun dest_binop expected tm = | |
| 837 | let | |
| 838 |    fun err () = raise ERR "dest_binop" ("Not a(n) " ^ quote expected);
 | |
| 839 | val (M, N) = dest_comb tm handle Utils.ERR _ => err () | |
| 840 | in (dest_monop expected M, N) handle Utils.ERR _ => err () end; | |
| 841 | ||
| 842 | fun dest_binder expected tm = | |
| 74518 | 843 | dest_abs (dest_monop expected tm) | 
| 60520 | 844 |   handle Utils.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
 | 
| 845 | ||
| 846 | ||
| 69593 | 847 | val dest_neg = dest_monop \<^const_name>\<open>Not\<close> | 
| 848 | val dest_pair = dest_binop \<^const_name>\<open>Pair\<close> | |
| 849 | val dest_eq = dest_binop \<^const_name>\<open>HOL.eq\<close> | |
| 850 | val dest_imp = dest_binop \<^const_name>\<open>HOL.implies\<close> | |
| 851 | val dest_conj = dest_binop \<^const_name>\<open>HOL.conj\<close> | |
| 852 | val dest_disj = dest_binop \<^const_name>\<open>HOL.disj\<close> | |
| 853 | val dest_exists = dest_binder \<^const_name>\<open>Ex\<close> | |
| 854 | val dest_forall = dest_binder \<^const_name>\<open>All\<close> | |
| 60520 | 855 | |
| 856 | (* Query routines *) | |
| 857 | ||
| 858 | val is_eq = can dest_eq | |
| 859 | val is_imp = can dest_imp | |
| 860 | val is_forall = can dest_forall | |
| 861 | val is_exists = can dest_exists | |
| 862 | val is_neg = can dest_neg | |
| 863 | val is_conj = can dest_conj | |
| 864 | val is_disj = can dest_disj | |
| 865 | val is_pair = can dest_pair | |
| 866 | ||
| 867 | ||
| 868 | (*--------------------------------------------------------------------------- | |
| 869 | * Iterated creation. | |
| 870 | *---------------------------------------------------------------------------*) | |
| 871 | val list_mk_disj = Utils.end_itlist (fn d1 => fn tm => mk_disj (d1, tm)); | |
| 872 | ||
| 873 | (*--------------------------------------------------------------------------- | |
| 874 | * Iterated destruction. (To the "right" in a term.) | |
| 875 | *---------------------------------------------------------------------------*) | |
| 876 | fun strip break tm = | |
| 877 | let fun dest (p as (ctm,accum)) = | |
| 878 | let val (M,N) = break ctm | |
| 879 | in dest (N, M::accum) | |
| 880 | end handle Utils.ERR _ => p | |
| 881 | in dest (tm,[]) | |
| 882 | end; | |
| 883 | ||
| 884 | fun rev2swap (x,l) = (rev l, x); | |
| 885 | ||
| 886 | val strip_comb = strip (Library.swap o dest_comb) (* Goes to the "left" *) | |
| 887 | val strip_imp = rev2swap o strip dest_imp | |
| 74518 | 888 | val strip_abs = rev2swap o strip dest_abs | 
| 60520 | 889 | val strip_forall = rev2swap o strip dest_forall | 
| 890 | val strip_exists = rev2swap o strip dest_exists | |
| 891 | ||
| 892 | val strip_disj = rev o (op::) o strip dest_disj | |
| 893 | ||
| 894 | ||
| 895 | (*--------------------------------------------------------------------------- | |
| 896 | * Going into and out of prop | |
| 897 | *---------------------------------------------------------------------------*) | |
| 898 | ||
| 899 | fun is_Trueprop ct = | |
| 900 | (case Thm.term_of ct of | |
| 69593 | 901 | Const (\<^const_name>\<open>Trueprop\<close>, _) $ _ => true | 
| 60520 | 902 | | _ => false); | 
| 903 | ||
| 69593 | 904 | fun mk_prop ct = if is_Trueprop ct then ct else Thm.apply \<^cterm>\<open>Trueprop\<close> ct; | 
| 60520 | 905 | fun drop_prop ct = if is_Trueprop ct then Thm.dest_arg ct else ct; | 
| 906 | ||
| 907 | end; | |
| 908 | ||
| 909 | ||
| 60521 | 910 | |
| 60520 | 911 | (*** emulation of HOL inference rules for TFL ***) | 
| 912 | ||
| 913 | structure Rules: RULES = | |
| 914 | struct | |
| 915 | ||
| 916 | fun RULES_ERR func mesg = Utils.ERR {module = "Rules", func = func, mesg = mesg};
 | |
| 917 | ||
| 918 | ||
| 60949 | 919 | fun cconcl thm = Dcterm.drop_prop (Thm.cprop_of thm); | 
| 920 | fun chyps thm = map Dcterm.drop_prop (Thm.chyps_of thm); | |
| 60520 | 921 | |
| 922 | fun dest_thm thm = | |
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 923 | (map HOLogic.dest_Trueprop (Thm.hyps_of thm), HOLogic.dest_Trueprop (Thm.prop_of thm)) | 
| 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 924 | handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop"; | 
| 60520 | 925 | |
| 926 | ||
| 927 | (* Inference rules *) | |
| 928 | ||
| 929 | (*--------------------------------------------------------------------------- | |
| 930 | * Equality (one step) | |
| 931 | *---------------------------------------------------------------------------*) | |
| 932 | ||
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 933 | fun REFL tm = HOLogic.mk_obj_eq (Thm.reflexive tm) | 
| 60520 | 934 | handle THM (msg, _, _) => raise RULES_ERR "REFL" msg; | 
| 935 | ||
| 936 | fun SYM thm = thm RS sym | |
| 937 | handle THM (msg, _, _) => raise RULES_ERR "SYM" msg; | |
| 938 | ||
| 939 | fun ALPHA thm ctm1 = | |
| 940 | let | |
| 941 | val ctm2 = Thm.cprop_of thm; | |
| 942 | val ctm2_eq = Thm.reflexive ctm2; | |
| 943 | val ctm1_eq = Thm.reflexive ctm1; | |
| 944 | in Thm.equal_elim (Thm.transitive ctm2_eq ctm1_eq) thm end | |
| 945 | handle THM (msg, _, _) => raise RULES_ERR "ALPHA" msg; | |
| 946 | ||
| 947 | fun rbeta th = | |
| 948 | (case Dcterm.strip_comb (cconcl th) of | |
| 60521 | 949 | (_, [_, r]) => Thm.transitive th (Thm.beta_conversion false r) | 
| 60520 | 950 | | _ => raise RULES_ERR "rbeta" ""); | 
| 951 | ||
| 952 | ||
| 953 | (*---------------------------------------------------------------------------- | |
| 954 | * Implication and the assumption list | |
| 955 | * | |
| 956 | * Assumptions get stuck on the meta-language assumption list. Implications | |
| 957 | * are in the object language, so discharging an assumption "A" from theorem | |
| 958 | * "B" results in something that looks like "A --> B". | |
| 959 | *---------------------------------------------------------------------------*) | |
| 960 | ||
| 961 | fun ASSUME ctm = Thm.assume (Dcterm.mk_prop ctm); | |
| 962 | ||
| 963 | ||
| 964 | (*--------------------------------------------------------------------------- | |
| 965 | * Implication in TFL is -->. Meta-language implication (==>) is only used | |
| 966 | * in the implementation of some of the inference rules below. | |
| 967 | *---------------------------------------------------------------------------*) | |
| 968 | fun MP th1 th2 = th2 RS (th1 RS mp) | |
| 969 | handle THM (msg, _, _) => raise RULES_ERR "MP" msg; | |
| 970 | ||
| 971 | (*forces the first argument to be a proposition if necessary*) | |
| 972 | fun DISCH tm thm = Thm.implies_intr (Dcterm.mk_prop tm) thm COMP impI | |
| 973 | handle THM (msg, _, _) => raise RULES_ERR "DISCH" msg; | |
| 974 | ||
| 60949 | 975 | fun DISCH_ALL thm = fold_rev DISCH (Thm.chyps_of thm) thm; | 
| 60520 | 976 | |
| 977 | ||
| 978 | fun FILTER_DISCH_ALL P thm = | |
| 979 | let fun check tm = P (Thm.term_of tm) | |
| 980 | in fold_rev (fn tm => fn th => if check tm then DISCH tm th else th) (chyps thm) thm | |
| 981 | end; | |
| 982 | ||
| 983 | fun UNDISCH thm = | |
| 984 | let val tm = Dcterm.mk_prop (#1 (Dcterm.dest_imp (cconcl thm))) | |
| 985 | in Thm.implies_elim (thm RS mp) (ASSUME tm) end | |
| 986 | handle Utils.ERR _ => raise RULES_ERR "UNDISCH" "" | |
| 987 | | THM _ => raise RULES_ERR "UNDISCH" ""; | |
| 988 | ||
| 989 | fun PROVE_HYP ath bth = MP (DISCH (cconcl ath) bth) ath; | |
| 990 | ||
| 60522 | 991 | fun IMP_TRANS th1 th2 = th2 RS (th1 RS @{thm tfl_imp_trans})
 | 
| 60520 | 992 | handle THM (msg, _, _) => raise RULES_ERR "IMP_TRANS" msg; | 
| 993 | ||
| 994 | ||
| 995 | (*---------------------------------------------------------------------------- | |
| 996 | * Conjunction | |
| 997 | *---------------------------------------------------------------------------*) | |
| 998 | ||
| 999 | fun CONJUNCT1 thm = thm RS conjunct1 | |
| 1000 | handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT1" msg; | |
| 1001 | ||
| 1002 | fun CONJUNCT2 thm = thm RS conjunct2 | |
| 1003 | handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT2" msg; | |
| 1004 | ||
| 1005 | fun CONJUNCTS th = CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th) handle Utils.ERR _ => [th]; | |
| 1006 | ||
| 1007 | fun LIST_CONJ [] = raise RULES_ERR "LIST_CONJ" "empty list" | |
| 1008 | | LIST_CONJ [th] = th | |
| 1009 | | LIST_CONJ (th :: rst) = MP (MP (conjI COMP (impI RS impI)) th) (LIST_CONJ rst) | |
| 1010 | handle THM (msg, _, _) => raise RULES_ERR "LIST_CONJ" msg; | |
| 1011 | ||
| 1012 | ||
| 1013 | (*---------------------------------------------------------------------------- | |
| 1014 | * Disjunction | |
| 1015 | *---------------------------------------------------------------------------*) | |
| 1016 | local | |
| 1017 | val prop = Thm.prop_of disjI1 | |
| 60524 | 1018 | val [_,Q] = Misc_Legacy.term_vars prop | 
| 69593 | 1019 | val disj1 = Thm.forall_intr (Thm.cterm_of \<^context> Q) disjI1 | 
| 60520 | 1020 | in | 
| 1021 | fun DISJ1 thm tm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj1) | |
| 1022 | handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg; | |
| 1023 | end; | |
| 1024 | ||
| 1025 | local | |
| 1026 | val prop = Thm.prop_of disjI2 | |
| 60521 | 1027 | val [P,_] = Misc_Legacy.term_vars prop | 
| 69593 | 1028 | val disj2 = Thm.forall_intr (Thm.cterm_of \<^context> P) disjI2 | 
| 60520 | 1029 | in | 
| 1030 | fun DISJ2 tm thm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj2) | |
| 1031 | handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg; | |
| 1032 | end; | |
| 1033 | ||
| 1034 | ||
| 1035 | (*---------------------------------------------------------------------------- | |
| 1036 | * | |
| 1037 | * A1 |- M1, ..., An |- Mn | |
| 1038 | * --------------------------------------------------- | |
| 1039 | * [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn] | |
| 1040 | * | |
| 1041 | *---------------------------------------------------------------------------*) | |
| 1042 | ||
| 1043 | ||
| 1044 | fun EVEN_ORS thms = | |
| 1045 | let fun blue ldisjs [] _ = [] | |
| 1046 | | blue ldisjs (th::rst) rdisjs = | |
| 1047 | let val tail = tl rdisjs | |
| 1048 | val rdisj_tl = Dcterm.list_mk_disj tail | |
| 1049 | in fold_rev DISJ2 ldisjs (DISJ1 th rdisj_tl) | |
| 1050 | :: blue (ldisjs @ [cconcl th]) rst tail | |
| 1051 | end handle Utils.ERR _ => [fold_rev DISJ2 ldisjs th] | |
| 1052 | in blue [] thms (map cconcl thms) end; | |
| 1053 | ||
| 1054 | ||
| 1055 | (*---------------------------------------------------------------------------- | |
| 1056 | * | |
| 1057 | * A |- P \/ Q B,P |- R C,Q |- R | |
| 1058 | * --------------------------------------------------- | |
| 1059 | * A U B U C |- R | |
| 1060 | * | |
| 1061 | *---------------------------------------------------------------------------*) | |
| 1062 | ||
| 1063 | fun DISJ_CASES th1 th2 th3 = | |
| 1064 | let | |
| 1065 | val c = Dcterm.drop_prop (cconcl th1); | |
| 1066 | val (disj1, disj2) = Dcterm.dest_disj c; | |
| 1067 | val th2' = DISCH disj1 th2; | |
| 1068 | val th3' = DISCH disj2 th3; | |
| 1069 | in | |
| 60522 | 1070 |     th3' RS (th2' RS (th1 RS @{thm tfl_disjE}))
 | 
| 60520 | 1071 | handle THM (msg, _, _) => raise RULES_ERR "DISJ_CASES" msg | 
| 1072 | end; | |
| 1073 | ||
| 1074 | ||
| 1075 | (*----------------------------------------------------------------------------- | |
| 1076 | * | |
| 1077 | * |- A1 \/ ... \/ An [A1 |- M, ..., An |- M] | |
| 1078 | * --------------------------------------------------- | |
| 1079 | * |- M | |
| 1080 | * | |
| 1081 | * Note. The list of theorems may be all jumbled up, so we have to | |
| 1082 | * first organize it to align with the first argument (the disjunctive | |
| 1083 | * theorem). | |
| 1084 | *---------------------------------------------------------------------------*) | |
| 1085 | ||
| 1086 | fun organize eq = (* a bit slow - analogous to insertion sort *) | |
| 1087 | let fun extract a alist = | |
| 1088 | let fun ex (_,[]) = raise RULES_ERR "organize" "not a permutation.1" | |
| 1089 | | ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t) | |
| 1090 | in ex ([],alist) | |
| 1091 | end | |
| 1092 | fun place [] [] = [] | |
| 1093 | | place (a::rst) alist = | |
| 1094 | let val (item,next) = extract a alist | |
| 1095 | in item::place rst next | |
| 1096 | end | |
| 1097 | | place _ _ = raise RULES_ERR "organize" "not a permutation.2" | |
| 1098 | in place | |
| 1099 | end; | |
| 1100 | ||
| 1101 | fun DISJ_CASESL disjth thl = | |
| 1102 | let val c = cconcl disjth | |
| 1103 | fun eq th atm = | |
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 1104 | exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th) | 
| 60520 | 1105 | val tml = Dcterm.strip_disj c | 
| 60521 | 1106 | fun DL _ [] = raise RULES_ERR "DISJ_CASESL" "no cases" | 
| 60520 | 1107 | | DL th [th1] = PROVE_HYP th th1 | 
| 1108 | | DL th [th1,th2] = DISJ_CASES th th1 th2 | |
| 1109 | | DL th (th1::rst) = | |
| 1110 | let val tm = #2 (Dcterm.dest_disj (Dcterm.drop_prop(cconcl th))) | |
| 1111 | in DISJ_CASES th th1 (DL (ASSUME tm) rst) end | |
| 1112 | in DL disjth (organize eq tml thl) | |
| 1113 | end; | |
| 1114 | ||
| 1115 | ||
| 1116 | (*---------------------------------------------------------------------------- | |
| 1117 | * Universals | |
| 1118 | *---------------------------------------------------------------------------*) | |
| 1119 | local (* this is fragile *) | |
| 1120 | val prop = Thm.prop_of spec | |
| 1121 | val x = hd (tl (Misc_Legacy.term_vars prop)) | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 1122 | val TV = dest_TVar (type_of x) | 
| 69593 | 1123 | val gspec = Thm.forall_intr (Thm.cterm_of \<^context> x) spec | 
| 60520 | 1124 | in | 
| 1125 | fun SPEC tm thm = | |
| 74282 | 1126 | let val gspec' = | 
| 77879 | 1127 | Drule.instantiate_normalize (TVars.make1 (TV, Thm.ctyp_of_cterm tm), Vars.empty) gspec | 
| 60520 | 1128 | in thm RS (Thm.forall_elim tm gspec') end | 
| 1129 | end; | |
| 1130 | ||
| 1131 | fun SPEC_ALL thm = fold SPEC (#1 (Dcterm.strip_forall(cconcl thm))) thm; | |
| 1132 | ||
| 1133 | val ISPEC = SPEC | |
| 1134 | val ISPECL = fold ISPEC; | |
| 1135 | ||
| 1136 | (* Not optimized! Too complicated. *) | |
| 1137 | local | |
| 1138 | val prop = Thm.prop_of allI | |
| 1139 | val [P] = Misc_Legacy.add_term_vars (prop, []) | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 1140 | fun cty_theta ctxt = map (fn (i, (S, ty)) => ((i, S), Thm.ctyp_of ctxt ty)) | 
| 60520 | 1141 | fun ctm_theta ctxt = | 
| 1142 | map (fn (i, (_, tm2)) => | |
| 1143 | let val ctm2 = Thm.cterm_of ctxt tm2 | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60524diff
changeset | 1144 | in ((i, Thm.typ_of_cterm ctm2), ctm2) end) | 
| 60520 | 1145 | fun certify ctxt (ty_theta,tm_theta) = | 
| 74282 | 1146 | (TVars.make (cty_theta ctxt (Vartab.dest ty_theta)), | 
| 1147 | Vars.make (ctm_theta ctxt (Vartab.dest tm_theta))) | |
| 60520 | 1148 | in | 
| 1149 | fun GEN ctxt v th = | |
| 1150 | let val gth = Thm.forall_intr v th | |
| 1151 | val thy = Proof_Context.theory_of ctxt | |
| 69593 | 1152 | val Const(\<^const_name>\<open>Pure.all\<close>,_)$Abs(x,ty,rst) = Thm.prop_of gth | 
| 60520 | 1153 | val P' = Abs(x,ty, HOLogic.dest_Trueprop rst) (* get rid of trueprop *) | 
| 1154 | val theta = Pattern.match thy (P,P') (Vartab.empty, Vartab.empty); | |
| 1155 | val allI2 = Drule.instantiate_normalize (certify ctxt theta) allI | |
| 1156 | val thm = Thm.implies_elim allI2 gth | |
| 1157 | val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm | |
| 1158 | val prop' = tp $ (A $ Abs(x,ty,M)) | |
| 1159 | in ALPHA thm (Thm.cterm_of ctxt prop') end | |
| 1160 | end; | |
| 1161 | ||
| 1162 | fun GENL ctxt = fold_rev (GEN ctxt); | |
| 1163 | ||
| 1164 | fun GEN_ALL ctxt thm = | |
| 1165 | let | |
| 1166 | val prop = Thm.prop_of thm | |
| 1167 | val vlist = map (Thm.cterm_of ctxt) (Misc_Legacy.add_term_vars (prop, [])) | |
| 1168 | in GENL ctxt vlist thm end; | |
| 1169 | ||
| 1170 | ||
| 1171 | fun MATCH_MP th1 th2 = | |
| 1172 | if (Dcterm.is_forall (Dcterm.drop_prop(cconcl th1))) | |
| 1173 | then MATCH_MP (th1 RS spec) th2 | |
| 1174 | else MP th1 th2; | |
| 1175 | ||
| 1176 | ||
| 1177 | (*---------------------------------------------------------------------------- | |
| 1178 | * Existentials | |
| 1179 | *---------------------------------------------------------------------------*) | |
| 1180 | ||
| 1181 | ||
| 1182 | ||
| 1183 | (*--------------------------------------------------------------------------- | |
| 1184 | * Existential elimination | |
| 1185 | * | |
| 1186 | * A1 |- ?x.t[x] , A2, "t[v]" |- t' | |
| 1187 | * ------------------------------------ (variable v occurs nowhere) | |
| 1188 | * A1 u A2 |- t' | |
| 1189 | * | |
| 1190 | *---------------------------------------------------------------------------*) | |
| 1191 | ||
| 1192 | fun CHOOSE ctxt (fvar, exth) fact = | |
| 1193 | let | |
| 1194 | val lam = #2 (Dcterm.dest_comb (Dcterm.drop_prop (cconcl exth))) | |
| 1195 | val redex = Dcterm.capply lam fvar | |
| 1196 | val t$u = Thm.term_of redex | |
| 1197 | val residue = Thm.cterm_of ctxt (Term.betapply (t, u)) | |
| 1198 | in | |
| 60522 | 1199 |     GEN ctxt fvar (DISCH residue fact) RS (exth RS @{thm tfl_exE})
 | 
| 60520 | 1200 | handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg | 
| 1201 | end; | |
| 1202 | ||
| 60781 | 1203 | fun EXISTS ctxt (template,witness) thm = | 
| 60520 | 1204 | let val abstr = #2 (Dcterm.dest_comb template) in | 
| 60781 | 1205 |     thm RS (infer_instantiate ctxt [(("P", 0), abstr), (("x", 0), witness)] exI)
 | 
| 60520 | 1206 | handle THM (msg, _, _) => raise RULES_ERR "EXISTS" msg | 
| 60781 | 1207 | end; | 
| 60520 | 1208 | |
| 1209 | (*---------------------------------------------------------------------------- | |
| 1210 | * | |
| 1211 | * A |- M[x_1,...,x_n] | |
| 1212 | * ---------------------------- [(x |-> y)_1,...,(x |-> y)_n] | |
| 1213 | * A |- ?y_1...y_n. M | |
| 1214 | * | |
| 1215 | *---------------------------------------------------------------------------*) | |
| 1216 | (* Could be improved, but needs "subst_free" for certified terms *) | |
| 1217 | ||
| 1218 | fun IT_EXISTS ctxt blist th = | |
| 1219 | let | |
| 1220 | val blist' = map (apply2 Thm.term_of) blist | |
| 1221 |     fun ex v M = Thm.cterm_of ctxt (USyntax.mk_exists{Bvar=v,Body = M})
 | |
| 1222 | in | |
| 1223 | fold_rev (fn (b as (r1,r2)) => fn thm => | |
| 60781 | 1224 | EXISTS ctxt (ex r2 (subst_free [b] | 
| 60520 | 1225 | (HOLogic.dest_Trueprop(Thm.prop_of thm))), Thm.cterm_of ctxt r1) | 
| 1226 | thm) | |
| 1227 | blist' th | |
| 1228 | end; | |
| 1229 | ||
| 1230 | (*---------------------------------------------------------------------------- | |
| 1231 | * Rewriting | |
| 1232 | *---------------------------------------------------------------------------*) | |
| 1233 | ||
| 1234 | fun SUBS ctxt thl = | |
| 1235 | rewrite_rule ctxt (map (fn th => th RS eq_reflection handle THM _ => th) thl); | |
| 1236 | ||
| 74532 | 1237 | fun rew_conv ctxt ctm = | 
| 1238 | Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) | |
| 1239 | (Variable.declare_term (Thm.term_of ctm) ctxt) ctm; | |
| 60520 | 1240 | |
| 1241 | fun simpl_conv ctxt thl ctm = | |
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1242 | HOLogic.mk_obj_eq (rew_conv (ctxt addsimps thl) ctm); | 
| 60520 | 1243 | |
| 1244 | ||
| 60522 | 1245 | fun RIGHT_ASSOC ctxt = rewrite_rule ctxt @{thms tfl_disj_assoc};
 | 
| 60520 | 1246 | |
| 1247 | ||
| 1248 | ||
| 1249 | (*--------------------------------------------------------------------------- | |
| 1250 | * TERMINATION CONDITION EXTRACTION | |
| 1251 | *---------------------------------------------------------------------------*) | |
| 1252 | ||
| 1253 | ||
| 1254 | (* Object language quantifier, i.e., "!" *) | |
| 1255 | fun Forall v M = USyntax.mk_forall{Bvar=v, Body=M};
 | |
| 1256 | ||
| 1257 | ||
| 1258 | (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *) | |
| 1259 | fun is_cong thm = | |
| 1260 | case (Thm.prop_of thm) of | |
| 69593 | 1261 | (Const(\<^const_name>\<open>Pure.imp\<close>,_)$(Const(\<^const_name>\<open>Trueprop\<close>,_)$ _) $ | 
| 1262 | (Const(\<^const_name>\<open>Pure.eq\<close>,_) $ (Const (\<^const_name>\<open>Wfrec.cut\<close>,_) $ _ $ _ $ _ $ _) $ _)) => | |
| 60520 | 1263 | false | 
| 1264 | | _ => true; | |
| 1265 | ||
| 1266 | ||
| 69593 | 1267 | fun dest_equal(Const (\<^const_name>\<open>Pure.eq\<close>,_) $ | 
| 1268 | (Const (\<^const_name>\<open>Trueprop\<close>,_) $ lhs) | |
| 1269 |                $ (Const (\<^const_name>\<open>Trueprop\<close>,_) $ rhs)) = {lhs=lhs, rhs=rhs}
 | |
| 1270 |   | dest_equal(Const (\<^const_name>\<open>Pure.eq\<close>,_) $ lhs $ rhs) = {lhs=lhs, rhs=rhs}
 | |
| 60520 | 1271 | | dest_equal tm = USyntax.dest_eq tm; | 
| 1272 | ||
| 1273 | fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm)); | |
| 1274 | ||
| 69593 | 1275 | fun dest_all used (Const(\<^const_name>\<open>Pure.all\<close>,_) $ (a as Abs _)) = USyntax.dest_abs used a | 
| 60520 | 1276 | | dest_all _ _ = raise RULES_ERR "dest_all" "not a !!"; | 
| 1277 | ||
| 1278 | val is_all = can (dest_all []); | |
| 1279 | ||
| 1280 | fun strip_all used fm = | |
| 1281 | if (is_all fm) | |
| 1282 |    then let val ({Bvar, Body}, used') = dest_all used fm
 | |
| 1283 | val (bvs, core, used'') = strip_all used' Body | |
| 1284 | in ((Bvar::bvs), core, used'') | |
| 1285 | end | |
| 1286 | else ([], fm, used); | |
| 1287 | ||
| 69593 | 1288 | fun list_break_all(Const(\<^const_name>\<open>Pure.all\<close>,_) $ Abs (s,ty,body)) = | 
| 60520 | 1289 | let val (L,core) = list_break_all body | 
| 1290 | in ((s,ty)::L, core) | |
| 1291 | end | |
| 1292 | | list_break_all tm = ([],tm); | |
| 1293 | ||
| 1294 | (*--------------------------------------------------------------------------- | |
| 1295 | * Rename a term of the form | |
| 1296 | * | |
| 1297 | * !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn | |
| 1298 | * ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn. | |
| 1299 | * to one of | |
| 1300 | * | |
| 1301 | * !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn | |
| 1302 | * ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn. | |
| 1303 | * | |
| 1304 | * This prevents name problems in extraction, and helps the result to read | |
| 1305 | * better. There is a problem with varstructs, since they can introduce more | |
| 1306 | * than n variables, and some extra reasoning needs to be done. | |
| 1307 | *---------------------------------------------------------------------------*) | |
| 1308 | ||
| 1309 | fun get ([],_,L) = rev L | |
| 1310 | | get (ant::rst,n,L) = | |
| 1311 | case (list_break_all ant) | |
| 1312 | of ([],_) => get (rst, n+1,L) | |
| 60521 | 1313 | | (_,body) => | 
| 60520 | 1314 | let val eq = Logic.strip_imp_concl body | 
| 60521 | 1315 | val (f,_) = USyntax.strip_comb (get_lhs eq) | 
| 60520 | 1316 | val (vstrl,_) = USyntax.strip_abs f | 
| 1317 | val names = | |
| 1318 | Name.variant_list (Misc_Legacy.add_term_names(body, [])) (map (#1 o dest_Free) vstrl) | |
| 1319 | in get (rst, n+1, (names,n)::L) end | |
| 1320 | handle TERM _ => get (rst, n+1, L) | |
| 1321 | | Utils.ERR _ => get (rst, n+1, L); | |
| 1322 | ||
| 1323 | (* Note: Thm.rename_params_rule counts from 1, not 0 *) | |
| 1324 | fun rename thm = | |
| 1325 | let | |
| 1326 | val ants = Logic.strip_imp_prems (Thm.prop_of thm) | |
| 1327 | val news = get (ants,1,[]) | |
| 1328 | in fold Thm.rename_params_rule news thm end; | |
| 1329 | ||
| 1330 | ||
| 1331 | (*--------------------------------------------------------------------------- | |
| 1332 | * Beta-conversion to the rhs of an equation (taken from hol90/drule.sml) | |
| 1333 | *---------------------------------------------------------------------------*) | |
| 1334 | ||
| 1335 | fun list_beta_conv tm = | |
| 1336 | let fun rbeta th = Thm.transitive th (Thm.beta_conversion false (#2(Dcterm.dest_eq(cconcl th)))) | |
| 1337 | fun iter [] = Thm.reflexive tm | |
| 1338 | | iter (v::rst) = rbeta (Thm.combination(iter rst) (Thm.reflexive v)) | |
| 1339 | in iter end; | |
| 1340 | ||
| 1341 | ||
| 1342 | (*--------------------------------------------------------------------------- | |
| 1343 | * Trace information for the rewriter | |
| 1344 | *---------------------------------------------------------------------------*) | |
| 1345 | val tracing = Unsynchronized.ref false; | |
| 1346 | ||
| 1347 | fun say s = if !tracing then writeln s else (); | |
| 1348 | ||
| 1349 | fun print_thms ctxt s L = | |
| 61268 | 1350 | say (cat_lines (s :: map (Thm.string_of_thm ctxt) L)); | 
| 60520 | 1351 | |
| 1352 | fun print_term ctxt s t = | |
| 1353 | say (cat_lines [s, Syntax.string_of_term ctxt t]); | |
| 1354 | ||
| 1355 | ||
| 1356 | (*--------------------------------------------------------------------------- | |
| 1357 | * General abstraction handlers, should probably go in USyntax. | |
| 1358 | *---------------------------------------------------------------------------*) | |
| 1359 | fun mk_aabs (vstr, body) = | |
| 1360 |   USyntax.mk_abs {Bvar = vstr, Body = body}
 | |
| 1361 |   handle Utils.ERR _ => USyntax.mk_pabs {varstruct = vstr, body = body};
 | |
| 1362 | ||
| 1363 | fun list_mk_aabs (vstrl,tm) = | |
| 1364 | fold_rev (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm; | |
| 1365 | ||
| 1366 | fun dest_aabs used tm = | |
| 1367 |    let val ({Bvar,Body}, used') = USyntax.dest_abs used tm
 | |
| 1368 | in (Bvar, Body, used') end | |
| 1369 | handle Utils.ERR _ => | |
| 1370 |      let val {varstruct, body, used} = USyntax.dest_pabs used tm
 | |
| 1371 | in (varstruct, body, used) end; | |
| 1372 | ||
| 1373 | fun strip_aabs used tm = | |
| 1374 | let val (vstr, body, used') = dest_aabs used tm | |
| 1375 | val (bvs, core, used'') = strip_aabs used' body | |
| 1376 | in (vstr::bvs, core, used'') end | |
| 1377 | handle Utils.ERR _ => ([], tm, used); | |
| 1378 | ||
| 1379 | fun dest_combn tm 0 = (tm,[]) | |
| 1380 | | dest_combn tm n = | |
| 1381 |      let val {Rator,Rand} = USyntax.dest_comb tm
 | |
| 1382 | val (f,rands) = dest_combn Rator (n-1) | |
| 1383 | in (f,Rand::rands) | |
| 1384 | end; | |
| 1385 | ||
| 1386 | ||
| 1387 | ||
| 1388 | ||
| 1389 | local fun dest_pair M = let val {fst,snd} = USyntax.dest_pair M in (fst,snd) end
 | |
| 1390 | fun mk_fst tm = | |
| 69593 | 1391 | let val ty as Type(\<^type_name>\<open>Product_Type.prod\<close>, [fty,sty]) = type_of tm | 
| 1392 | in Const (\<^const_name>\<open>Product_Type.fst\<close>, ty --> fty) $ tm end | |
| 60520 | 1393 | fun mk_snd tm = | 
| 69593 | 1394 | let val ty as Type(\<^type_name>\<open>Product_Type.prod\<close>, [fty,sty]) = type_of tm | 
| 1395 | in Const (\<^const_name>\<open>Product_Type.snd\<close>, ty --> sty) $ tm end | |
| 60520 | 1396 | in | 
| 1397 | fun XFILL tych x vstruct = | |
| 1398 | let fun traverse p xocc L = | |
| 1399 | if (is_Free p) | |
| 1400 | then tych xocc::L | |
| 1401 | else let val (p1,p2) = dest_pair p | |
| 1402 | in traverse p1 (mk_fst xocc) (traverse p2 (mk_snd xocc) L) | |
| 1403 | end | |
| 1404 | in | |
| 1405 | traverse vstruct x [] | |
| 1406 | end end; | |
| 1407 | ||
| 1408 | (*--------------------------------------------------------------------------- | |
| 1409 | * Replace a free tuple (vstr) by a universally quantified variable (a). | |
| 1410 | * Note that the notion of "freeness" for a tuple is different than for a | |
| 1411 | * variable: if variables in the tuple also occur in any other place than | |
| 1412 | * an occurrences of the tuple, they aren't "free" (which is thus probably | |
| 1413 | * the wrong word to use). | |
| 1414 | *---------------------------------------------------------------------------*) | |
| 1415 | ||
| 1416 | fun VSTRUCT_ELIM ctxt tych a vstr th = | |
| 1417 | let val L = USyntax.free_vars_lr vstr | |
| 1418 | val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr))) | |
| 1419 | val thm1 = Thm.implies_intr bind1 (SUBS ctxt [SYM(Thm.assume bind1)] th) | |
| 1420 | val thm2 = forall_intr_list (map tych L) thm1 | |
| 1421 | val thm3 = forall_elim_list (XFILL tych a vstr) thm2 | |
| 1422 | in refl RS | |
| 1423 |      rewrite_rule ctxt [Thm.symmetric (@{thm surjective_pairing} RS eq_reflection)] thm3
 | |
| 1424 | end; | |
| 1425 | ||
| 1426 | fun PGEN ctxt tych a vstr th = | |
| 1427 | let val a1 = tych a | |
| 60781 | 1428 | in Thm.forall_intr a1 (VSTRUCT_ELIM ctxt tych a vstr th) end; | 
| 60520 | 1429 | |
| 1430 | ||
| 1431 | (*--------------------------------------------------------------------------- | |
| 1432 | * Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into | |
| 1433 | * | |
| 1434 | * (([x,y],N),vstr) | |
| 1435 | *---------------------------------------------------------------------------*) | |
| 1436 | fun dest_pbeta_redex used M n = | |
| 1437 | let val (f,args) = dest_combn M n | |
| 60521 | 1438 | val _ = dest_aabs used f | 
| 60520 | 1439 | in (strip_aabs used f,args) | 
| 1440 | end; | |
| 1441 | ||
| 1442 | fun pbeta_redex M n = can (fn t => dest_pbeta_redex [] t n) M; | |
| 1443 | ||
| 1444 | fun dest_impl tm = | |
| 1445 | let val ants = Logic.strip_imp_prems tm | |
| 1446 | val eq = Logic.strip_imp_concl tm | |
| 1447 | in (ants,get_lhs eq) | |
| 1448 | end; | |
| 1449 | ||
| 1450 | fun restricted t = is_some (USyntax.find_term | |
| 69593 | 1451 | (fn (Const(\<^const_name>\<open>Wfrec.cut\<close>,_)) =>true | _ => false) | 
| 60520 | 1452 | t) | 
| 1453 | ||
| 1454 | fun CONTEXT_REWRITE_RULE main_ctxt (func, G, cut_lemma, congs) th = | |
| 1455 | let val globals = func::G | |
| 1456 | val ctxt0 = empty_simpset main_ctxt | |
| 1457 |      val pbeta_reduce = simpl_conv ctxt0 [@{thm split_conv} RS eq_reflection];
 | |
| 1458 | val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref | |
| 1459 | val cut_lemma' = cut_lemma RS eq_reflection | |
| 1460 | fun prover used ctxt thm = | |
| 1461 | let fun cong_prover ctxt thm = | |
| 60521 | 1462 | let val _ = say "cong_prover:" | 
| 60520 | 1463 | val cntxt = Simplifier.prems_of ctxt | 
| 60521 | 1464 | val _ = print_thms ctxt "cntxt:" cntxt | 
| 1465 | val _ = say "cong rule:" | |
| 61268 | 1466 | val _ = say (Thm.string_of_thm ctxt thm) | 
| 60520 | 1467 | (* Unquantified eliminate *) | 
| 1468 | fun uq_eliminate (thm,imp) = | |
| 1469 | let val tych = Thm.cterm_of ctxt | |
| 1470 | val _ = print_term ctxt "To eliminate:" imp | |
| 1471 | val ants = map tych (Logic.strip_imp_prems imp) | |
| 1472 | val eq = Logic.strip_imp_concl imp | |
| 1473 | val lhs = tych(get_lhs eq) | |
| 1474 | val ctxt' = Simplifier.add_prems (map ASSUME ants) ctxt | |
| 1475 | val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ctxt' lhs | |
| 1476 | handle Utils.ERR _ => Thm.reflexive lhs | |
| 1477 | val _ = print_thms ctxt' "proven:" [lhs_eq_lhs1] | |
| 1478 | val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1 | |
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1479 | val lhs_eeq_lhs2 = HOLogic.mk_obj_eq lhs_eq_lhs2 | 
| 60520 | 1480 | in | 
| 1481 | lhs_eeq_lhs2 COMP thm | |
| 1482 | end | |
| 1483 | fun pq_eliminate (thm, vlist, imp_body, lhs_eq) = | |
| 1484 | let val ((vstrl, _, used'), args) = dest_pbeta_redex used lhs_eq (length vlist) | |
| 60521 | 1485 | val _ = forall (op aconv) (ListPair.zip (vlist, args)) | 
| 60520 | 1486 | orelse error "assertion failed in CONTEXT_REWRITE_RULE" | 
| 1487 | val imp_body1 = subst_free (ListPair.zip (args, vstrl)) | |
| 1488 | imp_body | |
| 1489 | val tych = Thm.cterm_of ctxt | |
| 1490 | val ants1 = map tych (Logic.strip_imp_prems imp_body1) | |
| 1491 | val eq1 = Logic.strip_imp_concl imp_body1 | |
| 1492 | val Q = get_lhs eq1 | |
| 1493 | val QeqQ1 = pbeta_reduce (tych Q) | |
| 1494 | val Q1 = #2(Dcterm.dest_eq(cconcl QeqQ1)) | |
| 1495 | val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt | |
| 1496 | val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' Q1 | |
| 1497 | handle Utils.ERR _ => Thm.reflexive Q1 | |
| 1498 | val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2)) | |
| 1499 | val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl)) | |
| 1500 | val Q2eeqQ3 = Thm.symmetric(pbeta_reduce Q3 RS eq_reflection) | |
| 1501 | val thA = Thm.transitive(QeqQ1 RS eq_reflection) Q1eeqQ2 | |
| 1502 | val QeeqQ3 = Thm.transitive thA Q2eeqQ3 handle THM _ => | |
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1503 | (HOLogic.mk_obj_eq Q2eeqQ3 | 
| 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1504 | RS (HOLogic.mk_obj_eq thA RS trans)) | 
| 60520 | 1505 | RS eq_reflection | 
| 1506 | val impth = implies_intr_list ants1 QeeqQ3 | |
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1507 | val impth1 = HOLogic.mk_obj_eq impth | 
| 60520 | 1508 | (* Need to abstract *) | 
| 1509 | val ant_th = Utils.itlist2 (PGEN ctxt' tych) args vstrl impth1 | |
| 1510 | in ant_th COMP thm | |
| 1511 | end | |
| 1512 | fun q_eliminate (thm, imp) = | |
| 1513 | let val (vlist, imp_body, used') = strip_all used imp | |
| 1514 | val (ants,Q) = dest_impl imp_body | |
| 1515 | in if (pbeta_redex Q) (length vlist) | |
| 1516 | then pq_eliminate (thm, vlist, imp_body, Q) | |
| 1517 | else | |
| 1518 | let val tych = Thm.cterm_of ctxt | |
| 1519 | val ants1 = map tych ants | |
| 1520 | val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt | |
| 1521 | val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm | |
| 1522 | (false,true,false) (prover used') ctxt' (tych Q) | |
| 1523 | handle Utils.ERR _ => Thm.reflexive (tych Q) | |
| 1524 | val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1 | |
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 1525 | val lhs_eq_lhs2 = HOLogic.mk_obj_eq lhs_eeq_lhs2 | 
| 60520 | 1526 | val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2 | 
| 1527 | in | |
| 1528 | ant_th COMP thm | |
| 1529 | end end | |
| 1530 | ||
| 1531 | fun eliminate thm = | |
| 1532 | case Thm.prop_of thm of | |
| 69593 | 1533 | Const(\<^const_name>\<open>Pure.imp\<close>,_) $ imp $ _ => | 
| 60520 | 1534 | eliminate | 
| 1535 | (if not(is_all imp) | |
| 1536 | then uq_eliminate (thm, imp) | |
| 1537 | else q_eliminate (thm, imp)) | |
| 1538 | (* Assume that the leading constant is ==, *) | |
| 1539 | | _ => thm (* if it is not a ==> *) | |
| 1540 | in SOME(eliminate (rename thm)) end | |
| 1541 | handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) | |
| 1542 | ||
| 1543 | fun restrict_prover ctxt thm = | |
| 1544 | let val _ = say "restrict_prover:" | |
| 1545 | val cntxt = rev (Simplifier.prems_of ctxt) | |
| 1546 | val _ = print_thms ctxt "cntxt:" cntxt | |
| 69593 | 1547 | val Const(\<^const_name>\<open>Pure.imp\<close>,_) $ (Const(\<^const_name>\<open>Trueprop\<close>,_) $ A) $ _ = | 
| 60520 | 1548 | Thm.prop_of thm | 
| 1549 | fun genl tm = let val vlist = subtract (op aconv) globals | |
| 1550 | (Misc_Legacy.add_term_frees(tm,[])) | |
| 1551 | in fold_rev Forall vlist tm | |
| 1552 | end | |
| 1553 | (*-------------------------------------------------------------- | |
| 1554 | * This actually isn't quite right, since it will think that | |
| 1555 | * not-fully applied occs. of "f" in the context mean that the | |
| 1556 | * current call is nested. The real solution is to pass in a | |
| 1557 | * term "f v1..vn" which is a pattern that any full application | |
| 1558 | * of "f" will match. | |
| 1559 | *-------------------------------------------------------------*) | |
| 1560 | val func_name = #1(dest_Const func) | |
| 1561 | fun is_func (Const (name,_)) = (name = func_name) | |
| 1562 | | is_func _ = false | |
| 1563 | val rcontext = rev cntxt | |
| 1564 | val cncl = HOLogic.dest_Trueprop o Thm.prop_of | |
| 1565 | val antl = case rcontext of [] => [] | |
| 1566 | | _ => [USyntax.list_mk_conj(map cncl rcontext)] | |
| 1567 | val TC = genl(USyntax.list_mk_imp(antl, A)) | |
| 1568 | val _ = print_term ctxt "func:" func | |
| 1569 | val _ = print_term ctxt "TC:" (HOLogic.mk_Trueprop TC) | |
| 1570 | val _ = tc_list := (TC :: !tc_list) | |
| 1571 | val nestedp = is_some (USyntax.find_term is_func TC) | |
| 1572 | val _ = if nestedp then say "nested" else say "not_nested" | |
| 1573 | val th' = if nestedp then raise RULES_ERR "solver" "nested function" | |
| 1574 | else let val cTC = Thm.cterm_of ctxt (HOLogic.mk_Trueprop TC) | |
| 1575 | in case rcontext of | |
| 1576 | [] => SPEC_ALL(ASSUME cTC) | |
| 1577 | | _ => MP (SPEC_ALL (ASSUME cTC)) | |
| 1578 | (LIST_CONJ rcontext) | |
| 1579 | end | |
| 1580 | val th'' = th' RS thm | |
| 1581 | in SOME (th'') | |
| 1582 | end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) | |
| 1583 | in | |
| 1584 | (if (is_cong thm) then cong_prover else restrict_prover) ctxt thm | |
| 1585 | end | |
| 1586 | val ctm = Thm.cprop_of th | |
| 1587 | val names = Misc_Legacy.add_term_names (Thm.term_of ctm, []) | |
| 1588 | val th1 = | |
| 1589 | Raw_Simplifier.rewrite_cterm (false, true, false) | |
| 1590 | (prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm | |
| 1591 | val th2 = Thm.equal_elim th1 th | |
| 1592 | in | |
| 1593 | (th2, filter_out restricted (!tc_list)) | |
| 1594 | end; | |
| 1595 | ||
| 1596 | ||
| 74532 | 1597 | fun prove ctxt strict t tac = | 
| 60520 | 1598 | let | 
| 70308 | 1599 | val ctxt' = Proof_Context.augment t ctxt; | 
| 60520 | 1600 | in | 
| 1601 | if strict | |
| 74532 | 1602 | then Goal.prove ctxt' [] [] t (tac o #context) | 
| 1603 | else Goal.prove ctxt' [] [] t (tac o #context) | |
| 60520 | 1604 | handle ERROR msg => (warning msg; raise RULES_ERR "prove" msg) | 
| 1605 | end; | |
| 1606 | ||
| 1607 | end; | |
| 1608 | ||
| 1609 | ||
| 60521 | 1610 | |
| 60520 | 1611 | (*** theory operations ***) | 
| 1612 | ||
| 1613 | structure Thry: THRY = | |
| 1614 | struct | |
| 1615 | ||
| 1616 | ||
| 1617 | fun THRY_ERR func mesg = Utils.ERR {module = "Thry", func = func, mesg = mesg};
 | |
| 1618 | ||
| 1619 | ||
| 1620 | (*--------------------------------------------------------------------------- | |
| 1621 | * Matching | |
| 1622 | *---------------------------------------------------------------------------*) | |
| 1623 | ||
| 1624 | local | |
| 1625 | ||
| 1626 | fun tybind (ixn, (S, T)) = (TVar (ixn, S), T); | |
| 1627 | ||
| 1628 | in | |
| 1629 | ||
| 1630 | fun match_term thry pat ob = | |
| 1631 | let | |
| 1632 | val (ty_theta, tm_theta) = Pattern.match thry (pat,ob) (Vartab.empty, Vartab.empty); | |
| 1633 | fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.subst_type ty_theta T), t) | |
| 1634 | in (map tmbind (Vartab.dest tm_theta), map tybind (Vartab.dest ty_theta)) | |
| 1635 | end; | |
| 1636 | ||
| 1637 | fun match_type thry pat ob = | |
| 1638 | map tybind (Vartab.dest (Sign.typ_match thry (pat, ob) Vartab.empty)); | |
| 1639 | ||
| 1640 | end; | |
| 1641 | ||
| 1642 | ||
| 1643 | (*--------------------------------------------------------------------------- | |
| 1644 | * Typing | |
| 1645 | *---------------------------------------------------------------------------*) | |
| 1646 | ||
| 1647 | fun typecheck thy t = | |
| 1648 | Thm.global_cterm_of thy t | |
| 1649 | handle TYPE (msg, _, _) => raise THRY_ERR "typecheck" msg | |
| 1650 | | TERM (msg, _) => raise THRY_ERR "typecheck" msg; | |
| 1651 | ||
| 1652 | ||
| 1653 | (*--------------------------------------------------------------------------- | |
| 1654 | * Get information about datatypes | |
| 1655 | *---------------------------------------------------------------------------*) | |
| 1656 | ||
| 1657 | fun match_info thy dtco = | |
| 1658 | case (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco, | |
| 1659 | BNF_LFP_Compat.get_constrs thy dtco) of | |
| 1660 |       (SOME {case_name, ... }, SOME constructors) =>
 | |
| 1661 |         SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors}
 | |
| 1662 | | _ => NONE; | |
| 1663 | ||
| 1664 | fun induct_info thy dtco = case BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco of | |
| 1665 | NONE => NONE | |
| 1666 |       | SOME {nchotomy, ...} =>
 | |
| 1667 |           SOME {nchotomy = nchotomy,
 | |
| 1668 | constructors = (map Const o the o BNF_LFP_Compat.get_constrs thy) dtco}; | |
| 1669 | ||
| 1670 | fun extract_info thy = | |
| 1671 | let val infos = map snd (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting])) | |
| 1672 |  in {case_congs = map (mk_meta_eq o #case_cong) infos,
 | |
| 1673 | case_rewrites = maps (map mk_meta_eq o #case_rewrites) infos} | |
| 1674 | end; | |
| 1675 | ||
| 1676 | ||
| 1677 | end; | |
| 1678 | ||
| 1679 | ||
| 60521 | 1680 | |
| 60520 | 1681 | (*** first part of main module ***) | 
| 1682 | ||
| 1683 | structure Prim: PRIM = | |
| 1684 | struct | |
| 1685 | ||
| 1686 | val trace = Unsynchronized.ref false; | |
| 1687 | ||
| 1688 | ||
| 1689 | fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg};
 | |
| 1690 | ||
| 1691 | val concl = #2 o Rules.dest_thm; | |
| 1692 | ||
| 1693 | val list_mk_type = Utils.end_itlist (curry (op -->)); | |
| 1694 | ||
| 1695 | ||
| 1696 | (*--------------------------------------------------------------------------- | |
| 1697 | * The next function is common to pattern-match translation and | |
| 1698 | * proof of completeness of cases for the induction theorem. | |
| 1699 | * | |
| 1700 | * The curried function "gvvariant" returns a function to generate distinct | |
| 1701 | * variables that are guaranteed not to be in names. The names of | |
| 1702 | * the variables go u, v, ..., z, aa, ..., az, ... The returned | |
| 1703 | * function contains embedded refs! | |
| 1704 | *---------------------------------------------------------------------------*) | |
| 1705 | fun gvvariant names = | |
| 1706 | let val slist = Unsynchronized.ref names | |
| 1707 | val vname = Unsynchronized.ref "u" | |
| 1708 | fun new() = | |
| 1709 | if member (op =) (!slist) (!vname) | |
| 1710 | then (vname := Symbol.bump_string (!vname); new()) | |
| 1711 | else (slist := !vname :: !slist; !vname) | |
| 1712 | in | |
| 1713 | fn ty => Free(new(), ty) | |
| 1714 | end; | |
| 1715 | ||
| 1716 | ||
| 1717 | (*--------------------------------------------------------------------------- | |
| 1718 | * Used in induction theorem production. This is the simple case of | |
| 1719 | * partitioning up pattern rows by the leading constructor. | |
| 1720 | *---------------------------------------------------------------------------*) | |
| 1721 | fun ipartition gv (constructors,rows) = | |
| 1722 | let fun pfail s = raise TFL_ERR "partition.part" s | |
| 1723 |       fun part {constrs = [],   rows = [],   A} = rev A
 | |
| 1724 |         | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
 | |
| 1725 |         | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
 | |
| 1726 |         | part {constrs = c::crst, rows,     A} =
 | |
| 1727 | let val (c, T) = dest_Const c | |
| 1728 | val L = binder_types T | |
| 1729 | val (in_group, not_in_group) = | |
| 1730 | fold_rev (fn (row as (p::rst, rhs)) => | |
| 1731 | fn (in_group,not_in_group) => | |
| 1732 | let val (pc,args) = USyntax.strip_comb p | |
| 1733 | in if (#1(dest_Const pc) = c) | |
| 1734 | then ((args@rst, rhs)::in_group, not_in_group) | |
| 1735 | else (in_group, row::not_in_group) | |
| 1736 | end) rows ([],[]) | |
| 1737 | val col_types = Utils.take type_of (length L, #1(hd in_group)) | |
| 1738 | in | |
| 1739 |           part{constrs = crst, rows = not_in_group,
 | |
| 1740 |                A = {constructor = c,
 | |
| 1741 | new_formals = map gv col_types, | |
| 1742 | group = in_group}::A} | |
| 1743 | end | |
| 1744 |   in part{constrs = constructors, rows = rows, A = []}
 | |
| 1745 | end; | |
| 1746 | ||
| 1747 | ||
| 1748 | ||
| 1749 | (*--------------------------------------------------------------------------- | |
| 1750 | * Each pattern carries with it a tag (i,b) where | |
| 1751 | * i is the clause it came from and | |
| 1752 | * b=true indicates that clause was given by the user | |
| 1753 | * (or is an instantiation of a user supplied pattern) | |
| 1754 | * b=false --> i = ~1 | |
| 1755 | *---------------------------------------------------------------------------*) | |
| 1756 | ||
| 1757 | type pattern = term * (int * bool) | |
| 1758 | ||
| 1759 | fun pattern_map f (tm,x) = (f tm, x); | |
| 1760 | ||
| 1761 | fun pattern_subst theta = pattern_map (subst_free theta); | |
| 1762 | ||
| 1763 | val pat_of = fst; | |
| 1764 | fun row_of_pat x = fst (snd x); | |
| 1765 | fun given x = snd (snd x); | |
| 1766 | ||
| 1767 | (*--------------------------------------------------------------------------- | |
| 1768 | * Produce an instance of a constructor, plus genvars for its arguments. | |
| 1769 | *---------------------------------------------------------------------------*) | |
| 1770 | fun fresh_constr ty_match colty gv c = | |
| 1771 | let val (_,Ty) = dest_Const c | |
| 1772 | val L = binder_types Ty | |
| 1773 | and ty = body_type Ty | |
| 1774 | val ty_theta = ty_match ty colty | |
| 1775 | val c' = USyntax.inst ty_theta c | |
| 1776 | val gvars = map (USyntax.inst ty_theta o gv) L | |
| 1777 | in (c', gvars) | |
| 1778 | end; | |
| 1779 | ||
| 1780 | ||
| 1781 | (*--------------------------------------------------------------------------- | |
| 1782 | * Goes through a list of rows and picks out the ones beginning with a | |
| 1783 | * pattern with constructor = name. | |
| 1784 | *---------------------------------------------------------------------------*) | |
| 1785 | fun mk_group name rows = | |
| 1786 | fold_rev (fn (row as ((prfx, p::rst), rhs)) => | |
| 1787 | fn (in_group,not_in_group) => | |
| 1788 | let val (pc,args) = USyntax.strip_comb p | |
| 1789 | in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false) | |
| 1790 | then (((prfx,args@rst), rhs)::in_group, not_in_group) | |
| 1791 | else (in_group, row::not_in_group) end) | |
| 1792 | rows ([],[]); | |
| 1793 | ||
| 1794 | (*--------------------------------------------------------------------------- | |
| 1795 | * Partition the rows. Not efficient: we should use hashing. | |
| 1796 | *---------------------------------------------------------------------------*) | |
| 1797 | fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows" | |
| 1798 | | partition gv ty_match | |
| 1799 | (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) = | |
| 1800 | let val fresh = fresh_constr ty_match colty gv | |
| 1801 |      fun part {constrs = [],      rows, A} = rev A
 | |
| 1802 |        | part {constrs = c::crst, rows, A} =
 | |
| 1803 | let val (c',gvars) = fresh c | |
| 1804 | val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows | |
| 1805 | val in_group' = | |
| 1806 | if (null in_group) (* Constructor not given *) | |
| 1807 | then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))] | |
| 1808 | else in_group | |
| 1809 | in | |
| 1810 |          part{constrs = crst,
 | |
| 1811 | rows = not_in_group, | |
| 1812 |               A = {constructor = c',
 | |
| 1813 | new_formals = gvars, | |
| 1814 | group = in_group'}::A} | |
| 1815 | end | |
| 1816 | in part{constrs=constructors, rows=rows, A=[]}
 | |
| 1817 | end; | |
| 1818 | ||
| 1819 | (*--------------------------------------------------------------------------- | |
| 1820 | * Misc. routines used in mk_case | |
| 1821 | *---------------------------------------------------------------------------*) | |
| 1822 | ||
| 1823 | fun mk_pat (c,l) = | |
| 1824 | let val L = length (binder_types (type_of c)) | |
| 1825 | fun build (prfx,tag,plist) = | |
| 1826 | let val (args, plist') = chop L plist | |
| 1827 | in (prfx,tag,list_comb(c,args)::plist') end | |
| 1828 | in map build l end; | |
| 1829 | ||
| 1830 | fun v_to_prfx (prfx, v::pats) = (v::prfx,pats) | |
| 1831 | | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx"; | |
| 1832 | ||
| 1833 | fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats) | |
| 1834 | | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats"; | |
| 1835 | ||
| 1836 | ||
| 1837 | (*---------------------------------------------------------------------------- | |
| 1838 | * Translation of pattern terms into nested case expressions. | |
| 1839 | * | |
| 1840 | * This performs the translation and also builds the full set of patterns. | |
| 1841 | * Thus it supports the construction of induction theorems even when an | |
| 1842 | * incomplete set of patterns is given. | |
| 1843 | *---------------------------------------------------------------------------*) | |
| 1844 | ||
| 1845 | fun mk_case ty_info ty_match usednames range_ty = | |
| 1846 | let | |
| 1847 | fun mk_case_fail s = raise TFL_ERR "mk_case" s | |
| 1848 | val fresh_var = gvvariant usednames | |
| 1849 | val divide = partition fresh_var ty_match | |
| 60521 | 1850 | fun expand _ ty ((_,[]), _) = mk_case_fail"expand_var_row" | 
| 60520 | 1851 | | expand constructors ty (row as ((prfx, p::rst), rhs)) = | 
| 1852 | if (is_Free p) | |
| 1853 | then let val fresh = fresh_constr ty_match ty fresh_var | |
| 1854 | fun expnd (c,gvs) = | |
| 1855 | let val capp = list_comb(c,gvs) | |
| 1856 | in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs) | |
| 1857 | end | |
| 1858 | in map expnd (map fresh constructors) end | |
| 1859 | else [row] | |
| 1860 |  fun mk{rows=[],...} = mk_case_fail"no rows"
 | |
| 1861 |    | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
 | |
| 1862 | ([(prfx,tag,[])], tm) | |
| 1863 |    | mk{path=[], rows = _::_} = mk_case_fail"blunder"
 | |
| 1864 |    | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
 | |
| 1865 |         mk{path = path,
 | |
| 1866 | rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst} | |
| 1867 |    | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
 | |
| 1868 | let val (pat_rectangle,rights) = ListPair.unzip rows | |
| 1869 | val col0 = map(hd o #2) pat_rectangle | |
| 1870 | in | |
| 1871 | if (forall is_Free col0) | |
| 1872 | then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e) | |
| 1873 | (ListPair.zip (col0, rights)) | |
| 1874 | val pat_rectangle' = map v_to_prfx pat_rectangle | |
| 1875 |               val (pref_patl,tm) = mk{path = rstp,
 | |
| 1876 | rows = ListPair.zip (pat_rectangle', | |
| 1877 | rights')} | |
| 1878 | in (map v_to_pats pref_patl, tm) | |
| 1879 | end | |
| 1880 | else | |
| 1881 | let val pty as Type (ty_name,_) = type_of p | |
| 1882 | in | |
| 1883 | case (ty_info ty_name) | |
| 1884 |      of NONE => mk_case_fail("Not a known datatype: "^ty_name)
 | |
| 1885 |       | SOME{case_const,constructors} =>
 | |
| 1886 | let | |
| 1887 | val case_const_name = #1(dest_Const case_const) | |
| 1888 | val nrows = maps (expand constructors pty) rows | |
| 1889 | val subproblems = divide(constructors, pty, range_ty, nrows) | |
| 1890 | val groups = map #group subproblems | |
| 1891 | and new_formals = map #new_formals subproblems | |
| 1892 | and constructors' = map #constructor subproblems | |
| 1893 |             val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
 | |
| 1894 | (ListPair.zip (new_formals, groups)) | |
| 1895 | val rec_calls = map mk news | |
| 1896 | val (pat_rect,dtrees) = ListPair.unzip rec_calls | |
| 1897 | val case_functions = map USyntax.list_mk_abs | |
| 1898 | (ListPair.zip (new_formals, dtrees)) | |
| 1899 | val types = map type_of (case_functions@[u]) @ [range_ty] | |
| 1900 | val case_const' = Const(case_const_name, list_mk_type types) | |
| 1901 | val tree = list_comb(case_const', case_functions@[u]) | |
| 1902 | val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect)) | |
| 1903 | in (pat_rect1,tree) | |
| 1904 | end | |
| 1905 | end end | |
| 1906 | in mk | |
| 1907 | end; | |
| 1908 | ||
| 1909 | ||
| 1910 | (* Repeated variable occurrences in a pattern are not allowed. *) | |
| 1911 | fun FV_multiset tm = | |
| 1912 | case (USyntax.dest_term tm) | |
| 1913 |      of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)]
 | |
| 1914 | | USyntax.CONST _ => [] | |
| 1915 |       | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
 | |
| 1916 | | USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda"; | |
| 1917 | ||
| 1918 | fun no_repeat_vars thy pat = | |
| 1919 | let fun check [] = true | |
| 1920 | | check (v::rst) = | |
| 1921 | if member (op aconv) rst v then | |
| 1922 | raise TFL_ERR "no_repeat_vars" | |
| 1923 | (quote (#1 (dest_Free v)) ^ | |
| 1924 | " occurs repeatedly in the pattern " ^ | |
| 1925 | quote (Syntax.string_of_term_global thy pat)) | |
| 1926 | else check rst | |
| 1927 | in check (FV_multiset pat) | |
| 1928 | end; | |
| 1929 | ||
| 1930 | fun dest_atom (Free p) = p | |
| 1931 | | dest_atom (Const p) = p | |
| 1932 | | dest_atom _ = raise TFL_ERR "dest_atom" "function name not an identifier"; | |
| 1933 | ||
| 1934 | fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q); | |
| 1935 | ||
| 1936 | local fun mk_functional_err s = raise TFL_ERR "mk_functional" s | |
| 1937 | fun single [_$_] = | |
| 1938 | mk_functional_err "recdef does not allow currying" | |
| 1939 | | single [f] = f | |
| 1940 | | single fs = | |
| 1941 | (*multiple function names?*) | |
| 1942 | if length (distinct same_name fs) < length fs | |
| 1943 | then mk_functional_err | |
| 1944 | "The function being declared appears with multiple types" | |
| 1945 | else mk_functional_err | |
| 1946 | (string_of_int (length fs) ^ | |
| 1947 | " distinct function names being declared") | |
| 1948 | in | |
| 1949 | fun mk_functional thy clauses = | |
| 1950 | let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses | |
| 1951 | handle TERM _ => raise TFL_ERR "mk_functional" | |
| 1952 | "recursion equations must use the = relation") | |
| 1953 | val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L) | |
| 1954 | val atom = single (distinct (op aconv) funcs) | |
| 1955 | val (fname,ftype) = dest_atom atom | |
| 60521 | 1956 | val _ = map (no_repeat_vars thy) pats | 
| 60520 | 1957 | val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats, | 
| 1958 | map_index (fn (i, t) => (t,(i,true))) R) | |
| 1959 | val names = List.foldr Misc_Legacy.add_term_names [] R | |
| 1960 | val atype = type_of(hd pats) | |
| 1961 | and aname = singleton (Name.variant_list names) "a" | |
| 1962 | val a = Free(aname,atype) | |
| 1963 | val ty_info = Thry.match_info thy | |
| 1964 | val ty_match = Thry.match_type thy | |
| 1965 | val range_ty = type_of (hd R) | |
| 1966 | val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty | |
| 1967 |                                     {path=[a], rows=rows}
 | |
| 1968 | val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts | |
| 1969 | handle Match => mk_functional_err "error in pattern-match translation" | |
| 1970 | val patts2 = Library.sort (Library.int_ord o apply2 row_of_pat) patts1 | |
| 1971 | val finals = map row_of_pat patts2 | |
| 1972 | val originals = map (row_of_pat o #2) rows | |
| 60521 | 1973 | val _ = case (subtract (op =) finals originals) | 
| 60520 | 1974 | of [] => () | 
| 1975 | | L => mk_functional_err | |
| 1976 |  ("The following clauses are redundant (covered by preceding clauses): " ^
 | |
| 1977 | commas (map (fn i => string_of_int (i + 1)) L)) | |
| 1978 |  in {functional = Abs(Long_Name.base_name fname, ftype,
 | |
| 1979 | abstract_over (atom, absfree (aname,atype) case_tm)), | |
| 1980 | pats = patts2} | |
| 1981 | end end; | |
| 1982 | ||
| 1983 | ||
| 1984 | (*---------------------------------------------------------------------------- | |
| 1985 | * | |
| 1986 | * PRINCIPLES OF DEFINITION | |
| 1987 | * | |
| 1988 | *---------------------------------------------------------------------------*) | |
| 1989 | ||
| 1990 | ||
| 1991 | (*For Isabelle, the lhs of a definition must be a constant.*) | |
| 1992 | fun const_def sign (c, Ty, rhs) = | |
| 1993 | singleton (Syntax.check_terms (Proof_Context.init_global sign)) | |
| 69593 | 1994 | (Const(\<^const_name>\<open>Pure.eq\<close>,dummyT) $ Const(c,Ty) $ rhs); | 
| 60520 | 1995 | |
| 1996 | (*Make all TVars available for instantiation by adding a ? to the front*) | |
| 1997 | fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts) | |
| 1998 |   | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
 | |
| 1999 |   | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
 | |
| 2000 | ||
| 2001 | local | |
| 2002 | val f_eq_wfrec_R_M = | |
| 60522 | 2003 |     #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl @{thm tfl_wfrec}))))
 | 
| 60520 | 2004 |   val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M
 | 
| 60521 | 2005 | val _ = dest_Free f | 
| 60520 | 2006 | val (wfrec,_) = USyntax.strip_comb rhs | 
| 2007 | in | |
| 2008 | ||
| 2009 | fun wfrec_definition0 fid R (functional as Abs(x, Ty, _)) thy = | |
| 2010 | let | |
| 2011 | val def_name = Thm.def_name (Long_Name.base_name fid) | |
| 2012 | val wfrec_R_M = map_types poly_tvars (wfrec $ map_types poly_tvars R) $ functional | |
| 2013 | val def_term = const_def thy (fid, Ty, wfrec_R_M) | |
| 2014 | val ([def], thy') = | |
| 2015 | Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy | |
| 2016 | in (def, thy') end; | |
| 2017 | ||
| 2018 | end; | |
| 2019 | ||
| 2020 | ||
| 2021 | ||
| 2022 | (*--------------------------------------------------------------------------- | |
| 2023 | * This structure keeps track of congruence rules that aren't derived | |
| 2024 | * from a datatype definition. | |
| 2025 | *---------------------------------------------------------------------------*) | |
| 2026 | fun extraction_thms thy = | |
| 2027 |  let val {case_rewrites,case_congs} = Thry.extract_info thy
 | |
| 2028 | in (case_rewrites, case_congs) | |
| 2029 | end; | |
| 2030 | ||
| 2031 | ||
| 2032 | (*--------------------------------------------------------------------------- | |
| 2033 | * Pair patterns with termination conditions. The full list of patterns for | |
| 2034 | * a definition is merged with the TCs arising from the user-given clauses. | |
| 2035 | * There can be fewer clauses than the full list, if the user omitted some | |
| 2036 | * cases. This routine is used to prepare input for mk_induction. | |
| 2037 | *---------------------------------------------------------------------------*) | |
| 2038 | fun merge full_pats TCs = | |
| 2039 | let fun insert (p,TCs) = | |
| 2040 | let fun insrt ((x as (h,[]))::rst) = | |
| 2041 | if (p aconv h) then (p,TCs)::rst else x::insrt rst | |
| 2042 | | insrt (x::rst) = x::insrt rst | |
| 2043 | | insrt[] = raise TFL_ERR "merge.insert" "pattern not found" | |
| 2044 | in insrt end | |
| 2045 | fun pass ([],ptcl_final) = ptcl_final | |
| 2046 | | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl) | |
| 2047 | in | |
| 2048 | pass (TCs, map (fn p => (p,[])) full_pats) | |
| 2049 | end; | |
| 2050 | ||
| 2051 | ||
| 2052 | fun post_definition ctxt meta_tflCongs (def, pats) = | |
| 2053 | let val thy = Proof_Context.theory_of ctxt | |
| 2054 | val tych = Thry.typecheck thy | |
| 2055 | val f = #lhs(USyntax.dest_eq(concl def)) | |
| 60522 | 2056 |      val corollary = Rules.MATCH_MP @{thm tfl_wfrec} def
 | 
| 60520 | 2057 | val pats' = filter given pats | 
| 2058 | val given_pats = map pat_of pats' | |
| 2059 | val rows = map row_of_pat pats' | |
| 2060 | val WFR = #ant(USyntax.dest_imp(concl corollary)) | |
| 2061 | val R = #Rand(USyntax.dest_comb WFR) | |
| 2062 | val corollary' = Rules.UNDISCH corollary (* put WF R on assums *) | |
| 2063 | val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats | |
| 2064 | val (case_rewrites,context_congs) = extraction_thms thy | |
| 2065 | (*case_ss causes minimal simplification: bodies of case expressions are | |
| 2066 | not simplified. Otherwise large examples (Red-Black trees) are too | |
| 2067 | slow.*) | |
| 2068 | val case_simpset = | |
| 2069 | put_simpset HOL_basic_ss ctxt | |
| 2070 | addsimps case_rewrites | |
| 2071 | |> fold (Simplifier.add_cong o #case_cong_weak o snd) | |
| 2072 | (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting])) | |
| 2073 | val corollaries' = map (Simplifier.simplify case_simpset) corollaries | |
| 2074 | val extract = | |
| 2075 |       Rules.CONTEXT_REWRITE_RULE ctxt (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs)
 | |
| 2076 | val (rules, TCs) = ListPair.unzip (map extract corollaries') | |
| 60522 | 2077 |      val rules0 = map (rewrite_rule ctxt @{thms tfl_cut_def}) rules
 | 
| 60520 | 2078 | val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR) | 
| 2079 | val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0) | |
| 2080 | in | |
| 2081 |  {rules = rules1,
 | |
| 2082 | rows = rows, | |
| 2083 | full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)), | |
| 2084 | TCs = TCs} | |
| 2085 | end; | |
| 2086 | ||
| 2087 | ||
| 2088 | (*---------------------------------------------------------------------------- | |
| 2089 | * | |
| 2090 | * INDUCTION THEOREM | |
| 2091 | * | |
| 2092 | *---------------------------------------------------------------------------*) | |
| 2093 | ||
| 2094 | ||
| 2095 | (*------------------------ Miscellaneous function -------------------------- | |
| 2096 | * | |
| 2097 | * [x_1,...,x_n] ?v_1...v_n. M[v_1,...,v_n] | |
| 2098 | * ----------------------------------------------------------- | |
| 2099 | * ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]), | |
| 2100 | * ... | |
| 2101 | * (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] ) | |
| 2102 | * | |
| 2103 | * This function is totally ad hoc. Used in the production of the induction | |
| 2104 | * theorem. The nchotomy theorem can have clauses that look like | |
| 2105 | * | |
| 2106 | * ?v1..vn. z = C vn..v1 | |
| 2107 | * | |
| 2108 | * in which the order of quantification is not the order of occurrence of the | |
| 2109 | * quantified variables as arguments to C. Since we have no control over this | |
| 2110 | * aspect of the nchotomy theorem, we make the correspondence explicit by | |
| 2111 | * pairing the incoming new variable with the term it gets beta-reduced into. | |
| 2112 | *---------------------------------------------------------------------------*) | |
| 2113 | ||
| 2114 | fun alpha_ex_unroll (xlist, tm) = | |
| 2115 | let val (qvars,body) = USyntax.strip_exists tm | |
| 2116 | val vlist = #2 (USyntax.strip_comb (USyntax.rhs body)) | |
| 2117 | val plist = ListPair.zip (vlist, xlist) | |
| 2118 | val args = map (the o AList.lookup (op aconv) plist) qvars | |
| 2119 | handle Option.Option => raise Fail "TFL.alpha_ex_unroll: no correspondence" | |
| 2120 | fun build ex [] = [] | |
| 2121 | | build (_$rex) (v::rst) = | |
| 2122 | let val ex1 = Term.betapply(rex, v) | |
| 2123 | in ex1 :: build ex1 rst | |
| 2124 | end | |
| 2125 | val (nex::exl) = rev (tm::build tm args) | |
| 2126 | in | |
| 2127 | (nex, ListPair.zip (args, rev exl)) | |
| 2128 | end; | |
| 2129 | ||
| 2130 | ||
| 2131 | ||
| 2132 | (*---------------------------------------------------------------------------- | |
| 2133 | * | |
| 2134 | * PROVING COMPLETENESS OF PATTERNS | |
| 2135 | * | |
| 2136 | *---------------------------------------------------------------------------*) | |
| 2137 | ||
| 60699 | 2138 | fun mk_case ctxt ty_info usednames = | 
| 60520 | 2139 | let | 
| 60699 | 2140 | val thy = Proof_Context.theory_of ctxt | 
| 60520 | 2141 | val divide = ipartition (gvvariant usednames) | 
| 2142 | val tych = Thry.typecheck thy | |
| 2143 | fun tych_binding(x,y) = (tych x, tych y) | |
| 2144 | fun fail s = raise TFL_ERR "mk_case" s | |
| 2145 |  fun mk{rows=[],...} = fail"no rows"
 | |
| 2146 |    | mk{path=[], rows = [([], (thm, bindings))]} =
 | |
| 2147 | Rules.IT_EXISTS ctxt (map tych_binding bindings) thm | |
| 2148 |    | mk{path = u::rstp, rows as (p::_, _)::_} =
 | |
| 2149 | let val (pat_rectangle,rights) = ListPair.unzip rows | |
| 2150 | val col0 = map hd pat_rectangle | |
| 2151 | val pat_rectangle' = map tl pat_rectangle | |
| 2152 | in | |
| 2153 | if (forall is_Free col0) (* column 0 is all variables *) | |
| 2154 | then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)])) | |
| 2155 | (ListPair.zip (rights, col0)) | |
| 2156 |           in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
 | |
| 2157 | end | |
| 2158 | else (* column 0 is all constructors *) | |
| 2159 | let val Type (ty_name,_) = type_of p | |
| 2160 | in | |
| 2161 | case (ty_info ty_name) | |
| 2162 |      of NONE => fail("Not a known datatype: "^ty_name)
 | |
| 2163 |       | SOME{constructors,nchotomy} =>
 | |
| 2164 | let val thm' = Rules.ISPEC (tych u) nchotomy | |
| 2165 | val disjuncts = USyntax.strip_disj (concl thm') | |
| 2166 | val subproblems = divide(constructors, rows) | |
| 2167 | val groups = map #group subproblems | |
| 2168 | and new_formals = map #new_formals subproblems | |
| 2169 | val existentials = ListPair.map alpha_ex_unroll | |
| 2170 | (new_formals, disjuncts) | |
| 2171 | val constraints = map #1 existentials | |
| 2172 | val vexl = map #2 existentials | |
| 2173 | fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS ctxt [Rules.ASSUME (tych tm)] th, b)) | |
| 2174 |             val news = map (fn (nf,rows,c) => {path = nf@rstp,
 | |
| 2175 | rows = map (expnd c) rows}) | |
| 2176 | (Utils.zip3 new_formals groups constraints) | |
| 2177 | val recursive_thms = map mk news | |
| 2178 | val build_exists = Library.foldr | |
| 2179 | (fn((x,t), th) => | |
| 2180 | Rules.CHOOSE ctxt (tych x, Rules.ASSUME (tych t)) th) | |
| 2181 | val thms' = ListPair.map build_exists (vexl, recursive_thms) | |
| 2182 | val same_concls = Rules.EVEN_ORS thms' | |
| 2183 | in Rules.DISJ_CASESL thm' same_concls | |
| 2184 | end | |
| 2185 | end end | |
| 2186 | in mk | |
| 2187 | end; | |
| 2188 | ||
| 2189 | ||
| 60699 | 2190 | fun complete_cases ctxt = | 
| 2191 | let val thy = Proof_Context.theory_of ctxt | |
| 60520 | 2192 | val tych = Thry.typecheck thy | 
| 2193 | val ty_info = Thry.induct_info thy | |
| 2194 | in fn pats => | |
| 2195 | let val names = List.foldr Misc_Legacy.add_term_names [] pats | |
| 2196 | val T = type_of (hd pats) | |
| 2197 | val aname = singleton (Name.variant_list names) "a" | |
| 2198 | val vname = singleton (Name.variant_list (aname::names)) "v" | |
| 2199 | val a = Free (aname, T) | |
| 2200 | val v = Free (vname, T) | |
| 2201 | val a_eq_v = HOLogic.mk_eq(a,v) | |
| 60781 | 2202 |      val ex_th0 = Rules.EXISTS ctxt (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
 | 
| 60520 | 2203 | (Rules.REFL (tych a)) | 
| 2204 | val th0 = Rules.ASSUME (tych a_eq_v) | |
| 2205 | val rows = map (fn x => ([x], (th0,[]))) pats | |
| 2206 | in | |
| 2207 | Rules.GEN ctxt (tych a) | |
| 2208 | (Rules.RIGHT_ASSOC ctxt | |
| 2209 | (Rules.CHOOSE ctxt (tych v, ex_th0) | |
| 60699 | 2210 | (mk_case ctxt ty_info (vname::aname::names) | 
| 2211 |                  {path=[v], rows=rows})))
 | |
| 60520 | 2212 | end end; | 
| 2213 | ||
| 2214 | ||
| 2215 | (*--------------------------------------------------------------------------- | |
| 2216 | * Constructing induction hypotheses: one for each recursive call. | |
| 2217 | * | |
| 2218 | * Note. R will never occur as a variable in the ind_clause, because | |
| 2219 | * to do so, it would have to be from a nested definition, and we don't | |
| 2220 | * allow nested defns to have R variable. | |
| 2221 | * | |
| 2222 | * Note. When the context is empty, there can be no local variables. | |
| 2223 | *---------------------------------------------------------------------------*) | |
| 2224 | ||
| 2225 | local infix 5 ==> | |
| 2226 |       fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
 | |
| 2227 | in | |
| 2228 | fun build_ih f (P,SV) (pat,TCs) = | |
| 2229 | let val pat_vars = USyntax.free_vars_lr pat | |
| 2230 | val globals = pat_vars@SV | |
| 2231 | fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm) | |
| 2232 | fun dest_TC tm = | |
| 2233 | let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm)) | |
| 2234 | val (R,y,_) = USyntax.dest_relation R_y_pat | |
| 2235 | val P_y = if (nested tm) then R_y_pat ==> P$y else P$y | |
| 2236 | in case cntxt | |
| 2237 | of [] => (P_y, (tm,[])) | |
| 2238 | | _ => let | |
| 2239 | val imp = USyntax.list_mk_conj cntxt ==> P_y | |
| 2240 | val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp) | |
| 2241 | val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs | |
| 2242 | in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end | |
| 2243 | end | |
| 2244 | in case TCs | |
| 2245 | of [] => (USyntax.list_mk_forall(pat_vars, P$pat), []) | |
| 2246 | | _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs) | |
| 2247 | val ind_clause = USyntax.list_mk_conj ihs ==> P$pat | |
| 2248 | in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals) | |
| 2249 | end | |
| 2250 | end | |
| 2251 | end; | |
| 2252 | ||
| 2253 | (*--------------------------------------------------------------------------- | |
| 2254 | * This function makes good on the promise made in "build_ih". | |
| 2255 | * | |
| 2256 | * Input is tm = "(!y. R y pat ==> P y) ==> P pat", | |
| 2257 | * TCs = TC_1[pat] ... TC_n[pat] | |
| 2258 | * thm = ih1 /\ ... /\ ih_n |- ih[pat] | |
| 2259 | *---------------------------------------------------------------------------*) | |
| 2260 | fun prove_case ctxt f (tm,TCs_locals,thm) = | |
| 2261 | let val tych = Thry.typecheck (Proof_Context.theory_of ctxt) | |
| 2262 | val antc = tych(#ant(USyntax.dest_imp tm)) | |
| 2263 | val thm' = Rules.SPEC_ALL thm | |
| 2264 | fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm) | |
| 2265 | fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC))))) | |
| 2266 | fun mk_ih ((TC,locals),th2,nested) = | |
| 2267 | Rules.GENL ctxt (map tych locals) | |
| 2268 | (if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2 | |
| 2269 | else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2 | |
| 2270 | else Rules.MP th2 TC) | |
| 2271 | in | |
| 2272 | Rules.DISCH antc | |
| 2273 | (if USyntax.is_imp(concl thm') (* recursive calls in this clause *) | |
| 2274 | then let val th1 = Rules.ASSUME antc | |
| 2275 | val TCs = map #1 TCs_locals | |
| 2276 | val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o | |
| 2277 | #2 o USyntax.strip_forall) TCs | |
| 2278 | val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs)) | |
| 2279 | TCs_locals | |
| 2280 | val th2list = map (fn t => Rules.SPEC (tych t) th1) ylist | |
| 2281 | val nlist = map nested TCs | |
| 2282 | val triples = Utils.zip3 TClist th2list nlist | |
| 2283 | val Pylist = map mk_ih triples | |
| 2284 | in Rules.MP thm' (Rules.LIST_CONJ Pylist) end | |
| 2285 | else thm') | |
| 2286 | end; | |
| 2287 | ||
| 2288 | ||
| 2289 | (*--------------------------------------------------------------------------- | |
| 2290 | * | |
| 2291 | * x = (v1,...,vn) |- M[x] | |
| 2292 | * --------------------------------------------- | |
| 2293 | * ?v1 ... vn. x = (v1,...,vn) |- M[x] | |
| 2294 | * | |
| 2295 | *---------------------------------------------------------------------------*) | |
| 2296 | fun LEFT_ABS_VSTRUCT ctxt tych thm = | |
| 2297 | let fun CHOOSER v (tm,thm) = | |
| 2298 |         let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm}
 | |
| 2299 | in (ex_tm, Rules.CHOOSE ctxt (tych v, Rules.ASSUME (tych ex_tm)) thm) | |
| 2300 | end | |
| 2301 | val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm)) | |
| 2302 |       val {lhs,rhs} = USyntax.dest_eq veq
 | |
| 2303 | val L = USyntax.free_vars_lr rhs | |
| 2304 | in #2 (fold_rev CHOOSER L (veq,thm)) end; | |
| 2305 | ||
| 2306 | ||
| 2307 | (*---------------------------------------------------------------------------- | |
| 2308 | * Input : f, R, and [(pat1,TCs1),..., (patn,TCsn)] | |
| 2309 | * | |
| 60522 | 2310 | * Instantiates tfl_wf_induct, getting Sinduct and then tries to prove | 
| 60520 | 2311 | * recursion induction (Rinduct) by proving the antecedent of Sinduct from | 
| 2312 | * the antecedent of Rinduct. | |
| 2313 | *---------------------------------------------------------------------------*) | |
| 60699 | 2314 | fun mk_induction ctxt {fconst, R, SV, pat_TCs_list} =
 | 
| 2315 | let | |
| 2316 | val thy = Proof_Context.theory_of ctxt | |
| 60520 | 2317 | val tych = Thry.typecheck thy | 
| 60522 | 2318 |     val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) @{thm tfl_wf_induct})
 | 
| 60520 | 2319 | val (pats,TCsl) = ListPair.unzip pat_TCs_list | 
| 60699 | 2320 | val case_thm = complete_cases ctxt pats | 
| 60520 | 2321 | val domain = (type_of o hd) pats | 
| 2322 | val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names) | |
| 2323 | [] (pats::TCsl))) "P" | |
| 2324 | val P = Free(Pname, domain --> HOLogic.boolT) | |
| 2325 | val Sinduct = Rules.SPEC (tych P) Sinduction | |
| 2326 | val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct) | |
| 2327 | val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list | |
| 2328 | val (Rassums,TCl') = ListPair.unzip Rassums_TCl' | |
| 2329 | val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums)) | |
| 2330 | val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats | |
| 2331 | val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum) | |
| 2332 | val proved_cases = map (prove_case ctxt fconst) tasks | |
| 2333 | val v = | |
| 2334 | Free (singleton | |
| 2335 | (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] (map concl proved_cases))) "v", | |
| 2336 | domain) | |
| 2337 | val vtyped = tych v | |
| 2338 | val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats | |
| 2339 | val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS ctxt [th]th') | |
| 2340 | (substs, proved_cases) | |
| 2341 | val abs_cases = map (LEFT_ABS_VSTRUCT ctxt tych) proved_cases1 | |
| 2342 | val dant = Rules.GEN ctxt vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases) | |
| 2343 | val dc = Rules.MP Sinduct dant | |
| 2344 | val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc))) | |
| 2345 | val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty) | |
| 2346 | val dc' = fold_rev (Rules.GEN ctxt o tych) vars | |
| 2347 | (Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc) | |
| 2348 | in | |
| 2349 | Rules.GEN ctxt (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc') | |
| 2350 | end | |
| 2351 | handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation"; | |
| 2352 | ||
| 2353 | ||
| 2354 | ||
| 2355 | (*--------------------------------------------------------------------------- | |
| 2356 | * | |
| 2357 | * POST PROCESSING | |
| 2358 | * | |
| 2359 | *---------------------------------------------------------------------------*) | |
| 2360 | ||
| 2361 | ||
| 2362 | fun simplify_induction thy hth ind = | |
| 2363 | let val tych = Thry.typecheck thy | |
| 2364 | val (asl,_) = Rules.dest_thm ind | |
| 2365 | val (_,tc_eq_tc') = Rules.dest_thm hth | |
| 2366 | val tc = USyntax.lhs tc_eq_tc' | |
| 2367 | fun loop [] = ind | |
| 2368 | | loop (asm::rst) = | |
| 2369 | if (can (Thry.match_term thy asm) tc) | |
| 2370 | then Rules.UNDISCH | |
| 2371 | (Rules.MATCH_MP | |
| 60522 | 2372 |                      (Rules.MATCH_MP @{thm tfl_simp_thm} (Rules.DISCH (tych asm) ind))
 | 
| 60520 | 2373 | hth) | 
| 2374 | else loop rst | |
| 2375 | in loop asl | |
| 2376 | end; | |
| 2377 | ||
| 2378 | ||
| 2379 | (*--------------------------------------------------------------------------- | |
| 2380 | * The termination condition is an antecedent to the rule, and an | |
| 2381 | * assumption to the theorem. | |
| 2382 | *---------------------------------------------------------------------------*) | |
| 2383 | fun elim_tc tcthm (rule,induction) = | |
| 2384 | (Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction) | |
| 2385 | ||
| 2386 | ||
| 2387 | fun trace_thms ctxt s L = | |
| 61268 | 2388 | if !trace then writeln (cat_lines (s :: map (Thm.string_of_thm ctxt) L)) | 
| 60520 | 2389 | else (); | 
| 2390 | ||
| 2391 | fun trace_cterm ctxt s ct = | |
| 2392 | if !trace then | |
| 2393 | writeln (cat_lines [s, Syntax.string_of_term ctxt (Thm.term_of ct)]) | |
| 2394 | else (); | |
| 2395 | ||
| 2396 | ||
| 2397 | fun postprocess ctxt strict {wf_tac, terminator, simplifier} {rules,induction,TCs} =
 | |
| 2398 | let | |
| 2399 | val thy = Proof_Context.theory_of ctxt; | |
| 2400 | val tych = Thry.typecheck thy; | |
| 2401 | ||
| 2402 | (*--------------------------------------------------------------------- | |
| 2403 | * Attempt to eliminate WF condition. It's the only assumption of rules | |
| 2404 | *---------------------------------------------------------------------*) | |
| 74532 | 2405 | val ((rules1, induction1), ctxt') = | 
| 2406 | let | |
| 2407 | val thm = | |
| 2408 | Rules.prove ctxt strict (HOLogic.mk_Trueprop (hd(#1(Rules.dest_thm rules)))) wf_tac | |
| 2409 | val ctxt' = Variable.declare_thm thm ctxt | |
| 2410 | in ((Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction), ctxt') | |
| 2411 | end handle Utils.ERR _ => ((rules, induction), ctxt); | |
| 60520 | 2412 | |
| 2413 | (*---------------------------------------------------------------------- | |
| 2414 | * The termination condition (tc) is simplified to |- tc = tc' (there | |
| 2415 | * might not be a change!) and then 3 attempts are made: | |
| 2416 | * | |
| 60522 | 2417 | * 1. if |- tc = T, then eliminate it with tfl_eq_True; otherwise, | 
| 60520 | 2418 | * 2. apply the terminator to tc'. If |- tc' = T then eliminate; else | 
| 2419 | * 3. replace tc by tc' in both the rules and the induction theorem. | |
| 2420 | *---------------------------------------------------------------------*) | |
| 2421 | ||
| 2422 | fun simplify_tc tc (r,ind) = | |
| 2423 | let val tc1 = tych tc | |
| 74532 | 2424 | val _ = trace_cterm ctxt' "TC before simplification: " tc1 | 
| 2425 | val tc_eq = simplifier ctxt' tc1 | |
| 2426 | val _ = trace_thms ctxt' "result: " [tc_eq] | |
| 60520 | 2427 | in | 
| 60522 | 2428 |        elim_tc (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq) (r,ind)
 | 
| 60520 | 2429 | handle Utils.ERR _ => | 
| 60522 | 2430 |         (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 74532 | 2431 | (Rules.prove ctxt' strict (HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq))) | 
| 2432 | terminator)) | |
| 60520 | 2433 | (r,ind) | 
| 2434 | handle Utils.ERR _ => | |
| 60522 | 2435 |           (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP @{thm tfl_simp_thm} r) tc_eq),
 | 
| 60520 | 2436 | simplify_induction thy tc_eq ind)) | 
| 2437 | end | |
| 2438 | ||
| 2439 | (*---------------------------------------------------------------------- | |
| 2440 | * Nested termination conditions are harder to get at, since they are | |
| 2441 | * left embedded in the body of the function (and in induction | |
| 2442 | * theorem hypotheses). Our "solution" is to simplify them, and try to | |
| 2443 | * prove termination, but leave the application of the resulting theorem | |
| 2444 | * to a higher level. So things go much as in "simplify_tc": the | |
| 2445 | * termination condition (tc) is simplified to |- tc = tc' (there might | |
| 2446 | * not be a change) and then 2 attempts are made: | |
| 2447 | * | |
| 2448 | * 1. if |- tc = T, then return |- tc; otherwise, | |
| 2449 | * 2. apply the terminator to tc'. If |- tc' = T then return |- tc; else | |
| 2450 | * 3. return |- tc = tc' | |
| 2451 | *---------------------------------------------------------------------*) | |
| 2452 | fun simplify_nested_tc tc = | |
| 74532 | 2453 | let val tc_eq = simplifier ctxt' (tych (#2 (USyntax.strip_forall tc))) | 
| 60520 | 2454 | in | 
| 74532 | 2455 | Rules.GEN_ALL ctxt' | 
| 60522 | 2456 |        (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq
 | 
| 60520 | 2457 | handle Utils.ERR _ => | 
| 60522 | 2458 |           (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 74532 | 2459 | (Rules.prove ctxt' strict (HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq))) | 
| 2460 | terminator) | |
| 60520 | 2461 | handle Utils.ERR _ => tc_eq)) | 
| 2462 | end | |
| 2463 | ||
| 2464 | (*------------------------------------------------------------------- | |
| 2465 | * Attempt to simplify the termination conditions in each rule and | |
| 2466 | * in the induction theorem. | |
| 2467 | *-------------------------------------------------------------------*) | |
| 2468 | fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm | |
| 2469 | fun loop ([],extras,R,ind) = (rev R, ind, extras) | |
| 2470 | | loop ((r,ftcs)::rst, nthms, R, ind) = | |
| 2471 | let val tcs = #1(strip_imp (concl r)) | |
| 2472 | val extra_tcs = subtract (op aconv) tcs ftcs | |
| 2473 | val extra_tc_thms = map simplify_nested_tc extra_tcs | |
| 2474 | val (r1,ind1) = fold simplify_tc tcs (r,ind) | |
| 2475 | val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1 | |
| 2476 | in loop(rst, nthms@extra_tc_thms, r2::R, ind1) | |
| 2477 | end | |
| 2478 | val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs) | |
| 2479 | val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1) | |
| 2480 | in | |
| 2481 |   {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras}
 | |
| 2482 | end; | |
| 2483 | ||
| 2484 | end; | |
| 2485 | ||
| 2486 | ||
| 60521 | 2487 | |
| 60520 | 2488 | (*** second part of main module (postprocessing of TFL definitions) ***) | 
| 2489 | ||
| 2490 | structure Tfl: TFL = | |
| 2491 | struct | |
| 2492 | ||
| 2493 | (* misc *) | |
| 2494 | ||
| 2495 | (*--------------------------------------------------------------------------- | |
| 2496 | * Extract termination goals so that they can be put it into a goalstack, or | |
| 2497 | * have a tactic directly applied to them. | |
| 2498 | *--------------------------------------------------------------------------*) | |
| 2499 | fun termination_goals rules = | |
| 2500 | map (Type.legacy_freeze o HOLogic.dest_Trueprop) | |
| 2501 | (fold_rev (union (op aconv) o Thm.prems_of) rules []); | |
| 2502 | ||
| 2503 | (*--------------------------------------------------------------------------- | |
| 2504 | * Three postprocessors are applied to the definition. It | |
| 2505 | * attempts to prove wellfoundedness of the given relation, simplifies the | |
| 2506 | * non-proved termination conditions, and finally attempts to prove the | |
| 2507 | * simplified termination conditions. | |
| 2508 | *--------------------------------------------------------------------------*) | |
| 2509 | fun std_postprocessor ctxt strict wfs = | |
| 2510 | Prim.postprocess ctxt strict | |
| 74532 | 2511 |    {wf_tac = fn ctxt' => REPEAT (ares_tac ctxt' wfs 1),
 | 
| 2512 | terminator = fn ctxt' => | |
| 2513 | asm_simp_tac ctxt' 1 | |
| 2514 | THEN TRY (Arith_Data.arith_tac ctxt' 1 ORELSE | |
| 2515 |         fast_force_tac (ctxt' addSDs @{thms not0_implies_Suc}) 1),
 | |
| 2516 | simplifier = fn ctxt' => Rules.simpl_conv ctxt' []}; | |
| 60520 | 2517 | |
| 2518 | ||
| 2519 | ||
| 2520 | val concl = #2 o Rules.dest_thm; | |
| 2521 | ||
| 2522 | (*--------------------------------------------------------------------------- | |
| 2523 | * Postprocess a definition made by "define". This is a separate stage of | |
| 2524 | * processing from the definition stage. | |
| 2525 | *---------------------------------------------------------------------------*) | |
| 2526 | local | |
| 2527 | ||
| 2528 | (* The rest of these local definitions are for the tricky nested case *) | |
| 2529 | val solved = not o can USyntax.dest_eq o #2 o USyntax.strip_forall o concl | |
| 2530 | ||
| 2531 | fun id_thm th = | |
| 2532 |    let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th))));
 | |
| 2533 | in lhs aconv rhs end | |
| 2534 | handle Utils.ERR _ => false; | |
| 2535 | ||
| 2536 | val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
 | |
| 2537 | fun mk_meta_eq r = | |
| 2538 | (case Thm.concl_of r of | |
| 69593 | 2539 | Const(\<^const_name>\<open>Pure.eq\<close>,_)$_$_ => r | 
| 2540 | | _ $(Const(\<^const_name>\<open>HOL.eq\<close>,_)$_$_) => r RS eq_reflection | |
| 60520 | 2541 | | _ => r RS P_imp_P_eq_True) | 
| 2542 | ||
| 2543 | (*Is this the best way to invoke the simplifier??*) | |
| 2544 | fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L)) | |
| 2545 | ||
| 2546 | fun join_assums ctxt th = | |
| 2547 | let val tych = Thm.cterm_of ctxt | |
| 2548 |       val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
 | |
| 2549 | val cntxtl = (#1 o USyntax.strip_imp) lhs (* cntxtl should = cntxtr *) | |
| 2550 | val cntxtr = (#1 o USyntax.strip_imp) rhs (* but union is solider *) | |
| 2551 | val cntxt = union (op aconv) cntxtl cntxtr | |
| 2552 | in | |
| 2553 | Rules.GEN_ALL ctxt | |
| 2554 | (Rules.DISCH_ALL | |
| 2555 | (rewrite ctxt (map (Rules.ASSUME o tych) cntxt) (Rules.SPEC_ALL th))) | |
| 2556 | end | |
| 2557 | val gen_all = USyntax.gen_all | |
| 2558 | in | |
| 2559 | fun proof_stage ctxt strict wfs {f, R, rules, full_pats_TCs, TCs} =
 | |
| 2560 | let | |
| 2561 | val _ = writeln "Proving induction theorem ..." | |
| 2562 | val ind = | |
| 60699 | 2563 | Prim.mk_induction ctxt | 
| 60520 | 2564 |         {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
 | 
| 2565 | val _ = writeln "Postprocessing ..."; | |
| 2566 |     val {rules, induction, nested_tcs} =
 | |
| 2567 |       std_postprocessor ctxt strict wfs {rules=rules, induction=ind, TCs=TCs}
 | |
| 2568 | in | |
| 2569 | case nested_tcs | |
| 2570 |   of [] => {induction=induction, rules=rules,tcs=[]}
 | |
| 60521 | 2571 | | L => let val _ = writeln "Simplifying nested TCs ..." | 
| 60520 | 2572 | val (solved,simplified,stubborn) = | 
| 2573 | fold_rev (fn th => fn (So,Si,St) => | |
| 2574 | if (id_thm th) then (So, Si, th::St) else | |
| 2575 | if (solved th) then (th::So, Si, St) | |
| 2576 | else (So, th::Si, St)) nested_tcs ([],[],[]) | |
| 2577 | val simplified' = map (join_assums ctxt) simplified | |
| 2578 | val dummy = (Prim.trace_thms ctxt "solved =" solved; | |
| 2579 | Prim.trace_thms ctxt "simplified' =" simplified') | |
| 74532 | 2580 | fun rewr th = | 
| 2581 | full_simplify (Variable.declare_thm th ctxt addsimps (solved @ simplified')) th; | |
| 60520 | 2582 | val dummy = Prim.trace_thms ctxt "Simplifying the induction rule..." [induction] | 
| 2583 | val induction' = rewr induction | |
| 2584 | val dummy = Prim.trace_thms ctxt "Simplifying the recursion rules..." [rules] | |
| 2585 | val rules' = rewr rules | |
| 2586 | val _ = writeln "... Postprocessing finished"; | |
| 2587 | in | |
| 2588 |           {induction = induction',
 | |
| 2589 | rules = rules', | |
| 2590 | tcs = map (gen_all o USyntax.rhs o #2 o USyntax.strip_forall o concl) | |
| 2591 | (simplified@stubborn)} | |
| 2592 | end | |
| 2593 | end; | |
| 2594 | ||
| 2595 | ||
| 2596 | (*lcp: curry the predicate of the induction rule*) | |
| 2597 | fun curry_rule ctxt rl = | |
| 2598 | Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl; | |
| 2599 | ||
| 2600 | (*lcp: put a theorem into Isabelle form, using meta-level connectives*) | |
| 2601 | fun meta_outer ctxt = | |
| 2602 | curry_rule ctxt o Drule.export_without_context o | |
| 60752 | 2603 | rule_by_tactic ctxt | 
| 2604 | (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' eresolve_tac ctxt [conjE]))); | |
| 60520 | 2605 | |
| 2606 | (*Strip off the outer !P*) | |
| 2607 | val spec'= | |
| 69593 | 2608 |   Rule_Insts.read_instantiate \<^context> [((("x", 0), Position.none), "P::'b=>bool")] [] spec;
 | 
| 60520 | 2609 | |
| 74532 | 2610 | fun rulify_no_asm ctxt th = | 
| 2611 | Object_Logic.rulify_no_asm (Variable.declare_thm th ctxt) th; | |
| 2612 | ||
| 60524 | 2613 | fun simplify_defn ctxt strict congs wfs pats def0 = | 
| 60520 | 2614 | let | 
| 60825 | 2615 | val thy = Proof_Context.theory_of ctxt; | 
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
66251diff
changeset | 2616 | val def = HOLogic.mk_obj_eq (Thm.unvarify_global thy def0) | 
| 60520 | 2617 |     val {rules, rows, TCs, full_pats_TCs} = Prim.post_definition ctxt congs (def, pats)
 | 
| 2618 |     val {lhs=f,rhs} = USyntax.dest_eq (concl def)
 | |
| 2619 | val (_,[R,_]) = USyntax.strip_comb rhs | |
| 60521 | 2620 | val _ = Prim.trace_thms ctxt "congs =" congs | 
| 60520 | 2621 | (*the next step has caused simplifier looping in some cases*) | 
| 2622 |     val {induction, rules, tcs} =
 | |
| 2623 | proof_stage ctxt strict wfs | |
| 2624 |        {f = f, R = R, rules = rules,
 | |
| 2625 | full_pats_TCs = full_pats_TCs, | |
| 2626 | TCs = TCs} | |
| 74532 | 2627 | val rules' = map (Drule.export_without_context o rulify_no_asm ctxt) (Rules.CONJUNCTS rules) | 
| 60520 | 2628 | in | 
| 74532 | 2629 |     {induct = meta_outer ctxt (rulify_no_asm ctxt (induction RS spec')),
 | 
| 60520 | 2630 | rules = ListPair.zip(rules', rows), | 
| 2631 | tcs = (termination_goals rules') @ tcs} | |
| 2632 | end | |
| 2633 |   handle Utils.ERR {mesg,func,module} =>
 | |
| 2634 | error (mesg ^ "\n (In TFL function " ^ module ^ "." ^ func ^ ")"); | |
| 2635 | ||
| 2636 | ||
| 2637 | (* Derive the initial equations from the case-split rules to meet the | |
| 2638 | users specification of the recursive function. *) | |
| 2639 | local | |
| 2640 | fun get_related_thms i = | |
| 2641 | map_filter ((fn (r,x) => if x = i then SOME r else NONE)); | |
| 2642 | ||
| 60521 | 2643 | fun solve_eq _ (_, [], _) = error "derive_init_eqs: missing rules" | 
| 2644 | | solve_eq _ (_, [a], i) = [(a, i)] | |
| 60520 | 2645 | | solve_eq ctxt (th, splitths, i) = | 
| 2646 | (writeln "Proving unsplit equation..."; | |
| 74532 | 2647 | [((Drule.export_without_context o rulify_no_asm ctxt) | 
| 60520 | 2648 | (CaseSplit.splitto ctxt splitths th), i)]) | 
| 2649 | handle ERROR s => | |
| 2650 |              (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths);
 | |
| 2651 | in | |
| 2652 | fun derive_init_eqs ctxt rules eqs = | |
| 2653 | map (Thm.trivial o Thm.cterm_of ctxt o HOLogic.mk_Trueprop) eqs | |
| 2654 | |> map_index (fn (i, e) => solve_eq ctxt (e, (get_related_thms i rules), i)) | |
| 2655 | |> flat; | |
| 2656 | end; | |
| 2657 | ||
| 2658 | ||
| 2659 | (*--------------------------------------------------------------------------- | |
| 2660 | * Defining a function with an associated termination relation. | |
| 2661 | *---------------------------------------------------------------------------*) | |
| 2662 | fun define_i strict congs wfs fid R eqs ctxt = | |
| 2663 | let | |
| 2664 | val thy = Proof_Context.theory_of ctxt | |
| 2665 |     val {functional, pats} = Prim.mk_functional thy eqs
 | |
| 2666 | val (def, thy') = Prim.wfrec_definition0 fid R functional thy | |
| 2667 | val ctxt' = Proof_Context.transfer thy' ctxt | |
| 2668 | val (lhs, _) = Logic.dest_equals (Thm.prop_of def) | |
| 60524 | 2669 |     val {induct, rules, tcs} = simplify_defn ctxt' strict congs wfs pats def
 | 
| 60520 | 2670 | val rules' = if strict then derive_init_eqs ctxt' rules eqs else rules | 
| 2671 |   in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, ctxt') end;
 | |
| 2672 | ||
| 2673 | fun define strict congs wfs fid R seqs ctxt = | |
| 2674 | define_i strict congs wfs fid | |
| 2675 | (Syntax.read_term ctxt R) (map (Syntax.read_term ctxt) seqs) ctxt | |
| 2676 |       handle Utils.ERR {mesg,...} => error mesg;
 | |
| 2677 | ||
| 2678 | end; | |
| 2679 | ||
| 2680 | end; | |
| 2681 | ||
| 2682 | ||
| 60521 | 2683 | |
| 60520 | 2684 | (*** wrappers for Isar ***) | 
| 2685 | ||
| 2686 | (** recdef hints **) | |
| 2687 | ||
| 2688 | (* type hints *) | |
| 2689 | ||
| 2690 | type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
 | |
| 2691 | ||
| 2692 | fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
 | |
| 2693 | fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
 | |
| 2694 | ||
| 2695 | fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs)); | |
| 2696 | fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs)); | |
| 2697 | fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs)); | |
| 2698 | ||
| 2699 | ||
| 2700 | (* congruence rules *) | |
| 2701 | ||
| 2702 | local | |
| 2703 | ||
| 2704 | val cong_head = | |
| 2705 | fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of; | |
| 2706 | ||
| 2707 | fun prep_cong raw_thm = | |
| 2708 | let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end; | |
| 2709 | ||
| 2710 | in | |
| 2711 | ||
| 2712 | fun add_cong raw_thm congs = | |
| 2713 | let | |
| 2714 | val (c, thm) = prep_cong raw_thm; | |
| 2715 | val _ = if AList.defined (op =) congs c | |
| 2716 |       then warning ("Overwriting recdef congruence rule for " ^ quote c)
 | |
| 2717 | else (); | |
| 2718 | in AList.update (op =) (c, thm) congs end; | |
| 2719 | ||
| 2720 | fun del_cong raw_thm congs = | |
| 2721 | let | |
| 60524 | 2722 | val (c, _) = prep_cong raw_thm; | 
| 60520 | 2723 | val _ = if AList.defined (op =) congs c | 
| 2724 | then () | |
| 2725 |       else warning ("No recdef congruence rule for " ^ quote c);
 | |
| 2726 | in AList.delete (op =) c congs end; | |
| 2727 | ||
| 2728 | end; | |
| 2729 | ||
| 2730 | ||
| 2731 | ||
| 2732 | (** global and local recdef data **) | |
| 2733 | ||
| 2734 | (* theory data *) | |
| 2735 | ||
| 2736 | type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
 | |
| 2737 | ||
| 2738 | structure Data = Generic_Data | |
| 2739 | ( | |
| 2740 | type T = recdef_info Symtab.table * hints; | |
| 2741 | val empty = (Symtab.empty, mk_hints ([], [], [])): T; | |
| 2742 | fun merge | |
| 2743 |    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
 | |
| 2744 |     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
 | |
| 2745 | (Symtab.merge (K true) (tab1, tab2), | |
| 2746 | mk_hints (Thm.merge_thms (simps1, simps2), | |
| 2747 | AList.merge (op =) (K true) (congs1, congs2), | |
| 2748 | Thm.merge_thms (wfs1, wfs2))); | |
| 2749 | ); | |
| 2750 | ||
| 2751 | val get_recdef = Symtab.lookup o #1 o Data.get o Context.Theory; | |
| 2752 | ||
| 2753 | fun put_recdef name info = | |
| 2754 | (Context.theory_map o Data.map o apfst) (fn tab => | |
| 2755 | Symtab.update_new (name, info) tab | |
| 2756 |       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name));
 | |
| 2757 | ||
| 2758 | val get_hints = #2 o Data.get o Context.Proof; | |
| 2759 | val map_hints = Data.map o apsnd; | |
| 2760 | ||
| 2761 | ||
| 2762 | (* attributes *) | |
| 2763 | ||
| 2764 | fun attrib f = Thm.declaration_attribute (map_hints o f); | |
| 2765 | ||
| 2766 | val simp_add = attrib (map_simps o Thm.add_thm); | |
| 2767 | val simp_del = attrib (map_simps o Thm.del_thm); | |
| 2768 | val cong_add = attrib (map_congs o add_cong); | |
| 2769 | val cong_del = attrib (map_congs o del_cong); | |
| 2770 | val wf_add = attrib (map_wfs o Thm.add_thm); | |
| 2771 | val wf_del = attrib (map_wfs o Thm.del_thm); | |
| 2772 | ||
| 2773 | ||
| 2774 | (* modifiers *) | |
| 2775 | ||
| 2776 | val recdef_simpN = "recdef_simp"; | |
| 2777 | val recdef_congN = "recdef_cong"; | |
| 2778 | val recdef_wfN = "recdef_wf"; | |
| 2779 | ||
| 2780 | val recdef_modifiers = | |
| 64556 | 2781 | [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add \<^here>), | 
| 2782 | Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), | |
| 2783 | Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), | |
| 2784 | Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add \<^here>), | |
| 2785 | Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>), | |
| 2786 | Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>), | |
| 2787 | Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add \<^here>), | |
| 2788 | Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add \<^here>), | |
| 2789 | Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del \<^here>)] @ | |
| 60520 | 2790 | Clasimp.clasimp_modifiers; | 
| 2791 | ||
| 2792 | ||
| 2793 | ||
| 2794 | (** prepare hints **) | |
| 2795 | ||
| 2796 | fun prepare_hints opt_src ctxt = | |
| 2797 | let | |
| 2798 | val ctxt' = | |
| 2799 | (case opt_src of | |
| 2800 | NONE => ctxt | |
| 2801 | | SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt)); | |
| 2802 |     val {simps, congs, wfs} = get_hints ctxt';
 | |
| 2803 |     val ctxt'' = ctxt' addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | |
| 2804 | in ((rev (map snd congs), wfs), ctxt'') end; | |
| 2805 | ||
| 2806 | fun prepare_hints_i () ctxt = | |
| 2807 | let | |
| 2808 |     val {simps, congs, wfs} = get_hints ctxt;
 | |
| 2809 |     val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | |
| 2810 | in ((rev (map snd congs), wfs), ctxt') end; | |
| 2811 | ||
| 2812 | ||
| 2813 | ||
| 2814 | (** add_recdef(_i) **) | |
| 2815 | ||
| 2816 | fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = | |
| 2817 | let | |
| 2818 | val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead"; | |
| 2819 | ||
| 2820 | val name = Sign.intern_const thy raw_name; | |
| 2821 | val bname = Long_Name.base_name name; | |
| 2822 |     val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
 | |
| 2823 | ||
| 2824 | val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); | |
| 2825 | val eq_atts = map (map (prep_att thy)) raw_eq_atts; | |
| 2826 | ||
| 2827 | val ((congs, wfs), ctxt) = prep_hints hints (Proof_Context.init_global thy); | |
| 2828 | (*We must remove imp_cong to prevent looping when the induction rule | |
| 2829 | is simplified. Many induction rules have nested implications that would | |
| 2830 | give rise to looping conditional rewriting.*) | |
| 2831 |     val ({lhs, rules = rules_idx, induct, tcs}, ctxt1) =
 | |
| 2832 | tfl_fn not_permissive congs wfs name R eqs ctxt; | |
| 2833 | val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); | |
| 2834 | val simp_att = | |
| 2835 | if null tcs then [Simplifier.simp_add, | |
| 69593 | 2836 | Named_Theorems.add \<^named_theorems>\<open>nitpick_simp\<close>] | 
| 60520 | 2837 | else []; | 
| 2838 | val ((simps' :: rules', [induct']), thy2) = | |
| 2839 | Proof_Context.theory_of ctxt1 | |
| 2840 | |> Sign.add_path bname | |
| 2841 | |> Global_Theory.add_thmss | |
| 2842 | (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) | |
| 2843 | ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])] | |
| 71214 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 wenzelm parents: 
71179diff
changeset | 2844 | ||> Spec_Rules.add_global (Binding.name bname) Spec_Rules.equational_recdef [lhs] (flat rules) | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
64556diff
changeset | 2845 | ||> null tcs ? Code.declare_default_eqns_global (map (rpair true) (flat rules)); | 
| 60520 | 2846 |     val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
 | 
| 2847 | val thy3 = | |
| 2848 | thy2 | |
| 2849 | |> put_recdef name result | |
| 2850 | |> Sign.parent_path; | |
| 2851 | in (thy3, result) end; | |
| 2852 | ||
| 2853 | val add_recdef = gen_add_recdef Tfl.define Attrib.attribute_cmd_global prepare_hints; | |
| 2854 | fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); | |
| 2855 | ||
| 2856 | ||
| 2857 | ||
| 2858 | (** package setup **) | |
| 2859 | ||
| 2860 | (* setup theory *) | |
| 2861 | ||
| 2862 | val _ = | |
| 2863 | Theory.setup | |
| 69593 | 2864 | (Attrib.setup \<^binding>\<open>recdef_simp\<close> (Attrib.add_del simp_add simp_del) | 
| 60520 | 2865 | "declaration of recdef simp rule" #> | 
| 69593 | 2866 | Attrib.setup \<^binding>\<open>recdef_cong\<close> (Attrib.add_del cong_add cong_del) | 
| 60520 | 2867 | "declaration of recdef cong rule" #> | 
| 69593 | 2868 | Attrib.setup \<^binding>\<open>recdef_wf\<close> (Attrib.add_del wf_add wf_del) | 
| 60520 | 2869 | "declaration of recdef wf rule"); | 
| 2870 | ||
| 2871 | ||
| 2872 | (* outer syntax *) | |
| 2873 | ||
| 2874 | val hints = | |
| 69593 | 2875 | \<^keyword>\<open>(\<close> |-- | 
| 2876 | Parse.!!! ((Parse.token \<^keyword>\<open>hints\<close> ::: Parse.args) --| \<^keyword>\<open>)\<close>); | |
| 60520 | 2877 | |
| 2878 | val recdef_decl = | |
| 2879 | Scan.optional | |
| 69593 | 2880 | (\<^keyword>\<open>(\<close> -- Parse.!!! (\<^keyword>\<open>permissive\<close> -- \<^keyword>\<open>)\<close>) >> K false) true -- | 
| 60520 | 2881 | Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop) | 
| 2882 | -- Scan.option hints | |
| 61466 | 2883 | >> (fn ((((p, f), R), eqs), src) => | 
| 2884 | #1 o add_recdef p f R (map (fn ((x, y), z) => ((x, z), y)) eqs) src); | |
| 60520 | 2885 | |
| 2886 | val _ = | |
| 69593 | 2887 | Outer_Syntax.command \<^command_keyword>\<open>recdef\<close> "define general recursive functions (obsolete TFL)" | 
| 60520 | 2888 | (recdef_decl >> Toplevel.theory); | 
| 2889 | ||
| 2890 | end; |