| author | huffman | 
| Tue, 02 Jun 2009 20:35:04 -0700 | |
| changeset 31400 | d671d74b2d1d | 
| parent 30607 | c3d1590debd8 | 
| permissions | -rw-r--r-- | 
| 22166 | 1 | (* Title: HOL/Tools/function_package/fundef_core.ML | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | ||
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
28965diff
changeset | 4 | A package for general recursive function definitions: | 
| 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
28965diff
changeset | 5 | Main functionality. | 
| 22166 | 6 | *) | 
| 7 | ||
| 8 | signature FUNDEF_CORE = | |
| 9 | sig | |
| 10 | val prepare_fundef : FundefCommon.fundef_config | |
| 11 | -> string (* defname *) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 12 | -> ((bstring * typ) * mixfix) list (* defined symbol *) | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 13 | -> ((bstring * typ) list * term list * term * term) list (* specification *) | 
| 22166 | 14 | -> local_theory | 
| 15 | ||
| 16 | -> (term (* f *) | |
| 17 | * thm (* goalstate *) | |
| 18 | * (thm -> FundefCommon.fundef_result) (* continuation *) | |
| 19 | ) * local_theory | |
| 20 | ||
| 21 | end | |
| 22 | ||
| 23 | structure FundefCore : FUNDEF_CORE = | |
| 24 | struct | |
| 25 | ||
| 23819 | 26 | val boolT = HOLogic.boolT | 
| 27 | val mk_eq = HOLogic.mk_eq | |
| 22166 | 28 | |
| 29 | open FundefLib | |
| 30 | open FundefCommon | |
| 31 | ||
| 32 | datatype globals = | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 33 |    Globals of {
 | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 34 | fvar: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 35 | domT: typ, | 
| 22166 | 36 | ranT: typ, | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 37 | h: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 38 | y: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 39 | x: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 40 | z: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 41 | a: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 42 | P: term, | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 43 | D: term, | 
| 22166 | 44 | Pbool:term | 
| 45 | } | |
| 46 | ||
| 47 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 48 | datatype rec_call_info = | 
| 22166 | 49 | RCInfo of | 
| 50 |   {
 | |
| 51 | RIvs: (string * typ) list, (* Call context: fixes and assumes *) | |
| 52 | CCas: thm list, | |
| 53 | rcarg: term, (* The recursive argument *) | |
| 54 | ||
| 55 | llRI: thm, | |
| 56 | h_assum: term | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 57 | } | 
| 22166 | 58 | |
| 59 | ||
| 60 | datatype clause_context = | |
| 61 | ClauseContext of | |
| 62 |   {
 | |
| 63 | ctxt : Proof.context, | |
| 64 | ||
| 65 | qs : term list, | |
| 66 | gs : term list, | |
| 67 | lhs: term, | |
| 68 | rhs: term, | |
| 69 | ||
| 70 | cqs: cterm list, | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 71 | ags: thm list, | 
| 22166 | 72 | case_hyp : thm | 
| 73 | } | |
| 74 | ||
| 75 | ||
| 76 | fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
 | |
| 77 |     ClauseContext { ctxt = ProofContext.transfer thy ctxt,
 | |
| 78 | qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp } | |
| 79 | ||
| 80 | ||
| 81 | datatype clause_info = | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 82 | ClauseInfo of | 
| 22166 | 83 |      {
 | 
| 84 | no: int, | |
| 85 | qglr : ((string * typ) list * term list * term * term), | |
| 86 | cdata : clause_context, | |
| 87 | ||
| 23819 | 88 | tree: FundefCtxTree.ctx_tree, | 
| 22166 | 89 | lGI: thm, | 
| 90 | RCs: rec_call_info list | |
| 91 | } | |
| 92 | ||
| 93 | ||
| 94 | (* Theory dependencies. *) | |
| 23880 | 95 | val Pair_inject = @{thm Product_Type.Pair_inject};
 | 
| 22166 | 96 | |
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 97 | val acc_induct_rule = @{thm accp_induct_rule};
 | 
| 22166 | 98 | |
| 23880 | 99 | val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
 | 
| 100 | val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
 | |
| 101 | val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
 | |
| 22166 | 102 | |
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 103 | val acc_downward = @{thm accp_downward};
 | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 104 | val accI = @{thm accp.accI};
 | 
| 27115 
0dcafa5c9e3f
eliminated obsolete case_split_thm -- use case_split;
 wenzelm parents: 
26748diff
changeset | 105 | val case_split = @{thm HOL.case_split};
 | 
| 23880 | 106 | val fundef_default_value = @{thm FunDef.fundef_default_value};
 | 
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 107 | val not_acc_down = @{thm not_accp_down};
 | 
| 22166 | 108 | |
| 109 | ||
| 110 | ||
| 111 | fun find_calls tree = | |
| 112 | let | |
| 113 | fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs) | |
| 114 | | add_Ri _ _ _ _ = raise Match | |
| 115 | in | |
| 116 | rev (FundefCtxTree.traverse_tree add_Ri tree []) | |
| 117 | end | |
| 118 | ||
| 119 | ||
| 120 | (** building proof obligations *) | |
| 121 | ||
| 122 | fun mk_compat_proof_obligations domT ranT fvar f glrs = | |
| 123 | let | |
| 124 | fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) = | |
| 125 | let | |
| 126 | val shift = incr_boundvars (length qs') | |
| 127 | in | |
| 27336 | 128 | Logic.mk_implies | 
| 129 | (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'), | |
| 130 | HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs')) | |
| 22166 | 131 | |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs') | 
| 27336 | 132 | |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs') | 
| 22166 | 133 | |> curry abstract_over fvar | 
| 134 | |> curry subst_bound f | |
| 135 | end | |
| 136 | in | |
| 137 | map mk_impl (unordered_pairs glrs) | |
| 138 | end | |
| 139 | ||
| 140 | ||
| 141 | fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
 | |
| 142 | let | |
| 143 |         fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
 | |
| 23819 | 144 | HOLogic.mk_Trueprop Pbool | 
| 145 | |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs))) | |
| 22166 | 146 | |> fold_rev (curry Logic.mk_implies) gs | 
| 147 | |> fold_rev mk_forall_rename (map fst oqs ~~ qs) | |
| 148 | in | |
| 23819 | 149 | HOLogic.mk_Trueprop Pbool | 
| 22166 | 150 | |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs) | 
| 151 |                  |> mk_forall_rename ("x", x)
 | |
| 152 |                  |> mk_forall_rename ("P", Pbool)
 | |
| 153 | end | |
| 154 | ||
| 155 | (** making a context with it's own local bindings **) | |
| 156 | ||
| 157 | fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) = | |
| 158 | let | |
| 159 | val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt | |
| 160 | |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs | |
| 161 | ||
| 162 | val thy = ProofContext.theory_of ctxt' | |
| 163 | ||
| 164 | fun inst t = subst_bounds (rev qs, t) | |
| 165 | val gs = map inst pre_gs | |
| 166 | val lhs = inst pre_lhs | |
| 167 | val rhs = inst pre_rhs | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 168 | |
| 22166 | 169 | val cqs = map (cterm_of thy) qs | 
| 170 | val ags = map (assume o cterm_of thy) gs | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 171 | |
| 23819 | 172 | val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs)))) | 
| 22166 | 173 | in | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 174 |       ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
 | 
| 22166 | 175 | cqs = cqs, ags = ags, case_hyp = case_hyp } | 
| 176 | end | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 177 | |
| 22166 | 178 | |
| 179 | (* lowlevel term function *) | |
| 180 | fun abstract_over_list vs body = | |
| 181 | let | |
| 182 | exception SAME; | |
| 183 | fun abs lev v tm = | |
| 184 | if v aconv tm then Bound lev | |
| 185 | else | |
| 186 | (case tm of | |
| 187 | Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t) | |
| 188 | | t $ u => (abs lev v t $ (abs lev v u handle SAME => u) handle SAME => t $ abs lev v u) | |
| 189 | | _ => raise SAME); | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 190 | in | 
| 22166 | 191 | fold_index (fn (i,v) => fn t => abs i v t handle SAME => t) vs body | 
| 192 | end | |
| 193 | ||
| 194 | ||
| 195 | ||
| 196 | fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms = | |
| 197 | let | |
| 198 |         val Globals {h, fvar, x, ...} = globals
 | |
| 199 | ||
| 200 |         val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
 | |
| 201 | val cert = Thm.cterm_of (ProofContext.theory_of ctxt) | |
| 202 | ||
| 203 | (* Instantiate the GIntro thm with "f" and import into the clause context. *) | |
| 204 | val lGI = GIntro_thm | |
| 205 | |> forall_elim (cert f) | |
| 206 | |> fold forall_elim cqs | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 207 | |> fold Thm.elim_implies ags | 
| 22166 | 208 | |
| 209 | fun mk_call_info (rcfix, rcassm, rcarg) RI = | |
| 210 | let | |
| 211 | val llRI = RI | |
| 212 | |> fold forall_elim cqs | |
| 213 | |> fold (forall_elim o cert o Free) rcfix | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 214 | |> fold Thm.elim_implies ags | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 215 | |> fold Thm.elim_implies rcassm | 
| 22166 | 216 | |
| 217 | val h_assum = | |
| 23819 | 218 | HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg)) | 
| 22166 | 219 | |> fold_rev (curry Logic.mk_implies o prop_of) rcassm | 
| 27336 | 220 | |> fold_rev (Logic.all o Free) rcfix | 
| 22166 | 221 | |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] [] | 
| 222 | |> abstract_over_list (rev qs) | |
| 223 | in | |
| 224 |                 RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
 | |
| 225 | end | |
| 226 | ||
| 227 | val RC_infos = map2 mk_call_info RCs RIntro_thms | |
| 228 | in | |
| 229 | ClauseInfo | |
| 230 |             {
 | |
| 231 | no=no, | |
| 232 | cdata=cdata, | |
| 233 | qglr=qglr, | |
| 234 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 235 | lGI=lGI, | 
| 22166 | 236 | RCs=RC_infos, | 
| 237 | tree=tree | |
| 238 | } | |
| 239 | end | |
| 240 | ||
| 241 | ||
| 242 | ||
| 243 | ||
| 244 | ||
| 245 | ||
| 246 | ||
| 247 | (* replace this by a table later*) | |
| 248 | fun store_compat_thms 0 thms = [] | |
| 249 | | store_compat_thms n thms = | |
| 250 | let | |
| 251 | val (thms1, thms2) = chop n thms | |
| 252 | in | |
| 253 | (thms1 :: store_compat_thms (n - 1) thms2) | |
| 254 | end | |
| 255 | ||
| 256 | (* expects i <= j *) | |
| 257 | fun lookup_compat_thm i j cts = | |
| 258 | nth (nth cts (i - 1)) (j - i) | |
| 259 | ||
| 260 | (* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *) | |
| 261 | (* if j < i, then turn around *) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 262 | fun get_compat_thm thy cts i j ctxi ctxj = | 
| 22166 | 263 | let | 
| 264 |       val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
 | |
| 265 |       val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
 | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 266 | |
| 23819 | 267 | val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj))) | 
| 22166 | 268 | in if j < i then | 
| 269 | let | |
| 270 | val compat = lookup_compat_thm j i cts | |
| 271 | in | |
| 272 | compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *) | |
| 273 | |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *) | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 274 | |> fold Thm.elim_implies agsj | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 275 | |> fold Thm.elim_implies agsi | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 276 | |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *) | 
| 22166 | 277 | end | 
| 278 | else | |
| 279 | let | |
| 280 | val compat = lookup_compat_thm i j cts | |
| 281 | in | |
| 282 | compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *) | |
| 283 | |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *) | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 284 | |> fold Thm.elim_implies agsi | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 285 | |> fold Thm.elim_implies agsj | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 286 | |> Thm.elim_implies (assume lhsi_eq_lhsj) | 
| 22166 | 287 | |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *) | 
| 288 | end | |
| 289 | end | |
| 290 | ||
| 291 | ||
| 292 | ||
| 293 | ||
| 294 | (* Generates the replacement lemma in fully quantified form. *) | |
| 295 | fun mk_replacement_lemma thy h ih_elim clause = | |
| 296 | let | |
| 297 |         val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
 | |
| 26196 | 298 | local open Conv in | 
| 299 | val ih_conv = arg1_conv o arg_conv o arg_conv | |
| 300 | end | |
| 22166 | 301 | |
| 26196 | 302 | val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim | 
| 22166 | 303 | |
| 304 |         val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
 | |
| 305 |         val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
 | |
| 306 | ||
| 26196 | 307 | val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree | 
| 22166 | 308 | |
| 309 | val replace_lemma = (eql RS meta_eq_to_obj_eq) | |
| 310 | |> implies_intr (cprop_of case_hyp) | |
| 311 | |> fold_rev (implies_intr o cprop_of) h_assums | |
| 312 | |> fold_rev (implies_intr o cprop_of) ags | |
| 313 | |> fold_rev forall_intr cqs | |
| 26628 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 wenzelm parents: 
26569diff
changeset | 314 | |> Thm.close_derivation | 
| 22166 | 315 | in | 
| 316 | replace_lemma | |
| 317 | end | |
| 318 | ||
| 319 | ||
| 320 | fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj = | |
| 321 | let | |
| 322 |         val Globals {h, y, x, fvar, ...} = globals
 | |
| 323 |         val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
 | |
| 324 |         val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
 | |
| 325 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 326 |         val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
 | 
| 22166 | 327 | = mk_clause_context x ctxti cdescj | 
| 328 | ||
| 329 | val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj' | |
| 330 | val compat = get_compat_thm thy compat_store i j cctxi cctxj | |
| 331 |         val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
 | |
| 332 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 333 | val RLj_import = | 
| 22166 | 334 | RLj |> fold forall_elim cqsj' | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 335 | |> fold Thm.elim_implies agsj' | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 336 | |> fold Thm.elim_implies Ghsj' | 
| 22166 | 337 | |
| 23819 | 338 | val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h)))) | 
| 339 | val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *) | |
| 22166 | 340 | in | 
| 341 | (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *) | |
| 342 | |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *) | |
| 343 | |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *) | |
| 344 | |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *) | |
| 345 | |> fold_rev (implies_intr o cprop_of) Ghsj' | |
| 346 | |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *) | |
| 347 | |> implies_intr (cprop_of y_eq_rhsj'h) | |
| 348 | |> implies_intr (cprop_of lhsi_eq_lhsj') | |
| 349 | |> fold_rev forall_intr (cterm_of thy h :: cqsj') | |
| 350 | end | |
| 351 | ||
| 352 | ||
| 353 | ||
| 22617 | 354 | fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei = | 
| 22166 | 355 | let | 
| 356 |         val Globals {x, y, ranT, fvar, ...} = globals
 | |
| 357 |         val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
 | |
| 358 | val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs | |
| 359 | ||
| 360 | val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro | |
| 361 | ||
| 362 |         fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
 | |
| 363 | |> fold_rev (implies_intr o cprop_of) CCas | |
| 364 | |> fold_rev (forall_intr o cterm_of thy o Free) RIvs | |
| 365 | ||
| 366 | val existence = fold (curry op COMP o prep_RC) RCs lGI | |
| 367 | ||
| 368 | val P = cterm_of thy (mk_eq (y, rhsC)) | |
| 23819 | 369 | val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y))) | 
| 22166 | 370 | |
| 371 | val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas | |
| 372 | ||
| 373 | val uniqueness = G_cases | |
| 374 | |> forall_elim (cterm_of thy lhs) | |
| 375 | |> forall_elim (cterm_of thy y) | |
| 376 | |> forall_elim P | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 377 | |> Thm.elim_implies G_lhs_y | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 378 | |> fold Thm.elim_implies unique_clauses | 
| 22166 | 379 | |> implies_intr (cprop_of G_lhs_y) | 
| 380 | |> forall_intr (cterm_of thy y) | |
| 381 | ||
| 382 | val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *) | |
| 383 | ||
| 384 | val exactly_one = | |
| 385 | ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)] | |
| 386 | |> curry (op COMP) existence | |
| 387 | |> curry (op COMP) uniqueness | |
| 388 | |> simplify (HOL_basic_ss addsimps [case_hyp RS sym]) | |
| 389 | |> implies_intr (cprop_of case_hyp) | |
| 390 | |> fold_rev (implies_intr o cprop_of) ags | |
| 391 | |> fold_rev forall_intr cqs | |
| 392 | ||
| 393 | val function_value = | |
| 394 | existence | |
| 395 | |> implies_intr ihyp | |
| 396 | |> implies_intr (cprop_of case_hyp) | |
| 397 | |> forall_intr (cterm_of thy x) | |
| 398 | |> forall_elim (cterm_of thy lhs) | |
| 399 | |> curry (op RS) refl | |
| 400 | in | |
| 401 | (exactly_one, function_value) | |
| 402 | end | |
| 403 | ||
| 404 | ||
| 405 | ||
| 406 | ||
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23880diff
changeset | 407 | fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def = | 
| 22166 | 408 | let | 
| 409 |         val Globals {h, domT, ranT, x, ...} = globals
 | |
| 22617 | 410 | val thy = ProofContext.theory_of ctxt | 
| 22166 | 411 | |
| 412 | (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *) | |
| 27336 | 413 |         val ihyp = Term.all domT $ Abs ("z", domT,
 | 
| 414 | Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x), | |
| 415 |                                      HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
 | |
| 416 |                                                              Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
 | |
| 22166 | 417 | |> cterm_of thy | 
| 418 | ||
| 26653 | 419 | val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0 | 
| 22166 | 420 | val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex) | 
| 421 | val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un) | |
| 26196 | 422 | |> instantiate' [] [NONE, SOME (cterm_of thy h)] | 
| 22166 | 423 | |
| 424 | val _ = Output.debug (K "Proving Replacement lemmas...") | |
| 425 | val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses | |
| 426 | ||
| 427 | val _ = Output.debug (K "Proving cases for unique existence...") | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 428 | val (ex1s, values) = | 
| 22617 | 429 | split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses) | 
| 22166 | 430 | |
| 22617 | 431 | val _ = Output.debug (K "Proving: Graph is a function") | 
| 22166 | 432 | val graph_is_function = complete | 
| 26653 | 433 | |> Thm.forall_elim_vars 0 | 
| 22166 | 434 | |> fold (curry op COMP) ex1s | 
| 435 | |> implies_intr (ihyp) | |
| 23819 | 436 | |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x))) | 
| 22166 | 437 | |> forall_intr (cterm_of thy x) | 
| 438 | |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *) | |
| 29329 | 439 | |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it) | 
| 22166 | 440 | |
| 441 | val goalstate = Conjunction.intr graph_is_function complete | |
| 26628 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 wenzelm parents: 
26569diff
changeset | 442 | |> Thm.close_derivation | 
| 22166 | 443 | |> Goal.protect | 
| 444 | |> fold_rev (implies_intr o cprop_of) compat | |
| 445 | |> implies_intr (cprop_of complete) | |
| 446 | in | |
| 447 | (goalstate, values) | |
| 448 | end | |
| 449 | ||
| 450 | ||
| 451 | fun define_graph Gname fvar domT ranT clauses RCss lthy = | |
| 452 | let | |
| 453 | val GT = domT --> ranT --> boolT | |
| 454 | val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)])) | |
| 455 | ||
| 456 |       fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
 | |
| 457 | let | |
| 458 | fun mk_h_assm (rcfix, rcassm, rcarg) = | |
| 23819 | 459 | HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg)) | 
| 22166 | 460 | |> fold_rev (curry Logic.mk_implies o prop_of) rcassm | 
| 27336 | 461 | |> fold_rev (Logic.all o Free) rcfix | 
| 22166 | 462 | in | 
| 23819 | 463 | HOLogic.mk_Trueprop (Gvar $ lhs $ rhs) | 
| 22166 | 464 | |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs | 
| 465 | |> fold_rev (curry Logic.mk_implies) gs | |
| 27336 | 466 | |> fold_rev Logic.all (fvar :: qs) | 
| 22166 | 467 | end | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 468 | |
| 22166 | 469 | val G_intros = map2 mk_GIntro clauses RCss | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 470 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 471 | val (GIntro_thms, (G, G_elim, G_induct, lthy)) = | 
| 22166 | 472 | FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy) | 
| 473 | in | |
| 474 | ((G, GIntro_thms, G_elim, G_induct), lthy) | |
| 475 | end | |
| 476 | ||
| 477 | ||
| 478 | ||
| 479 | fun define_function fdefname (fname, mixfix) domT ranT G default lthy = | |
| 480 | let | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 481 | val f_def = | 
| 22166 | 482 |           Abs ("x", domT, Const ("FunDef.THE_default", ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
 | 
| 483 |                                 Abs ("y", ranT, G $ Bound 1 $ Bound 0))
 | |
| 24510 
6c612357cf3d
replaced ProofContext.cert_term/prop by general Syntax.check_term/prop (which also includes type-inference);
 wenzelm parents: 
24168diff
changeset | 484 | |> Syntax.check_term lthy | 
| 22496 | 485 | |
| 22166 | 486 | val ((f, (_, f_defthm)), lthy) = | 
| 28965 | 487 | LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy | 
| 22166 | 488 | in | 
| 489 | ((f, f_defthm), lthy) | |
| 490 | end | |
| 491 | ||
| 492 | ||
| 493 | fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy = | |
| 494 | let | |
| 495 | ||
| 496 | val RT = domT --> domT --> boolT | |
| 497 | val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)])) | |
| 498 | ||
| 499 |       fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
 | |
| 23819 | 500 | HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs) | 
| 22166 | 501 | |> fold_rev (curry Logic.mk_implies o prop_of) rcassm | 
| 502 | |> fold_rev (curry Logic.mk_implies) gs | |
| 27336 | 503 | |> fold_rev (Logic.all o Free) rcfix | 
| 22166 | 504 | |> fold_rev mk_forall_rename (map fst oqs ~~ qs) | 
| 505 | (* "!!qs xs. CS ==> G => (r, lhs) : R" *) | |
| 506 | ||
| 507 | val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss | |
| 508 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 509 | val (RIntro_thmss, (R, R_elim, _, lthy)) = | 
| 22166 | 510 | fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy) | 
| 511 | in | |
| 512 | ((R, RIntro_thmss, R_elim), lthy) | |
| 513 | end | |
| 514 | ||
| 515 | ||
| 516 | fun fix_globals domT ranT fvar ctxt = | |
| 517 | let | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 518 | val ([h, y, x, z, a, D, P, Pbool],ctxt') = | 
| 22166 | 519 | Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt | 
| 520 | in | |
| 521 |       (Globals {h = Free (h, domT --> ranT),
 | |
| 522 | y = Free (y, ranT), | |
| 523 | x = Free (x, domT), | |
| 524 | z = Free (z, domT), | |
| 525 | a = Free (a, domT), | |
| 526 | D = Free (D, domT --> boolT), | |
| 527 | P = Free (P, domT --> boolT), | |
| 528 | Pbool = Free (Pbool, boolT), | |
| 529 | fvar = fvar, | |
| 530 | domT = domT, | |
| 531 | ranT = ranT | |
| 532 | }, | |
| 533 | ctxt') | |
| 534 | end | |
| 535 | ||
| 536 | ||
| 537 | ||
| 538 | fun inst_RC thy fvar f (rcfix, rcassm, rcarg) = | |
| 539 | let | |
| 540 | fun inst_term t = subst_bound(f, abstract_over (fvar, t)) | |
| 541 | in | |
| 542 | (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg) | |
| 543 | end | |
| 544 | ||
| 545 | ||
| 546 | ||
| 547 | (********************************************************** | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 548 | * PROVING THE RULES | 
| 22166 | 549 | **********************************************************) | 
| 550 | ||
| 551 | fun mk_psimps thy globals R clauses valthms f_iff graph_is_function = | |
| 552 | let | |
| 553 |       val Globals {domT, z, ...} = globals
 | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 554 | |
| 22166 | 555 |       fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
 | 
| 556 | let | |
| 23819 | 557 | val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *) | 
| 558 | val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *) | |
| 22166 | 559 | in | 
| 560 | ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward)) | |
| 561 | |> (fn it => it COMP graph_is_function) | |
| 562 | |> implies_intr z_smaller | |
| 563 | |> forall_intr (cterm_of thy z) | |
| 564 | |> (fn it => it COMP valthm) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 565 | |> implies_intr lhs_acc | 
| 22166 | 566 | |> asm_simplify (HOL_basic_ss addsimps [f_iff]) | 
| 567 | |> fold_rev (implies_intr o cprop_of) ags | |
| 568 | |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) | |
| 569 | end | |
| 570 | in | |
| 571 | map2 mk_psimp clauses valthms | |
| 572 | end | |
| 573 | ||
| 574 | ||
| 575 | (** Induction rule **) | |
| 576 | ||
| 577 | ||
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 578 | val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
 | 
| 22166 | 579 | |
| 26196 | 580 | |
| 26569 | 581 | fun binder_conv cv ctxt = Conv.arg_conv (Conv.abs_conv (K cv) ctxt); | 
| 26196 | 582 | |
| 22166 | 583 | fun mk_partial_induct_rule thy globals R complete_thm clauses = | 
| 584 | let | |
| 585 |       val Globals {domT, x, z, a, P, D, ...} = globals
 | |
| 586 | val acc_R = mk_acc domT R | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 587 | |
| 23819 | 588 | val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x))) | 
| 589 | val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a)) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 590 | |
| 27336 | 591 | val D_subset = cterm_of thy (Logic.all x | 
| 592 | (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x)))) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 593 | |
| 22166 | 594 | val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *) | 
| 27336 | 595 | Logic.all x | 
| 596 | (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), | |
| 23819 | 597 | Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x), | 
| 598 | HOLogic.mk_Trueprop (D $ z))))) | |
| 22166 | 599 | |> cterm_of thy | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 600 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 601 | |
| 22166 | 602 | (* Inductive Hypothesis: !!z. (z,x):R ==> P z *) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 603 |       val ihyp = Term.all domT $ Abs ("z", domT,
 | 
| 27336 | 604 | Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x), | 
| 605 | HOLogic.mk_Trueprop (P $ Bound 0))) | |
| 22166 | 606 | |> cterm_of thy | 
| 607 | ||
| 608 | val aihyp = assume ihyp | |
| 609 | ||
| 610 | fun prove_case clause = | |
| 611 | let | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 612 |     val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
 | 
| 22166 | 613 | qglr = (oqs, _, _, _), ...} = clause | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 614 | |
| 26196 | 615 | val case_hyp_conv = K (case_hyp RS eq_reflection) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 616 | local open Conv in | 
| 26196 | 617 | val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D | 
| 26569 | 618 | val sih = fconv_rule (binder_conv (arg1_conv (arg_conv (arg_conv case_hyp_conv))) ctxt) aihyp | 
| 26196 | 619 | end | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 620 | |
| 22166 | 621 |     fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
 | 
| 622 | sih |> forall_elim (cterm_of thy rcarg) | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 623 | |> Thm.elim_implies llRI | 
| 22166 | 624 | |> fold_rev (implies_intr o cprop_of) CCas | 
| 625 | |> fold_rev (forall_intr o cterm_of thy o Free) RIvs | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 626 | |
| 22166 | 627 | val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 628 | |
| 23819 | 629 | val step = HOLogic.mk_Trueprop (P $ lhs) | 
| 22166 | 630 | |> fold_rev (curry Logic.mk_implies o prop_of) P_recs | 
| 631 | |> fold_rev (curry Logic.mk_implies) gs | |
| 23819 | 632 | |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs)) | 
| 22166 | 633 | |> fold_rev mk_forall_rename (map fst oqs ~~ qs) | 
| 634 | |> cterm_of thy | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 635 | |
| 22166 | 636 | val P_lhs = assume step | 
| 637 | |> fold forall_elim cqs | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 638 | |> Thm.elim_implies lhs_D | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 639 | |> fold Thm.elim_implies ags | 
| 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 640 | |> fold Thm.elim_implies P_recs | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 641 | |
| 23819 | 642 | val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x)) | 
| 26196 | 643 | |> Conv.arg_conv (Conv.arg_conv case_hyp_conv) | 
| 22166 | 644 | |> symmetric (* P lhs == P x *) | 
| 645 | |> (fn eql => equal_elim eql P_lhs) (* "P x" *) | |
| 646 | |> implies_intr (cprop_of case_hyp) | |
| 647 | |> fold_rev (implies_intr o cprop_of) ags | |
| 648 | |> fold_rev forall_intr cqs | |
| 649 | in | |
| 650 | (res, step) | |
| 651 | end | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 652 | |
| 22166 | 653 | val (cases, steps) = split_list (map prove_case clauses) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 654 | |
| 22166 | 655 | val istep = complete_thm | 
| 26653 | 656 | |> Thm.forall_elim_vars 0 | 
| 22166 | 657 | |> fold (curry op COMP) cases (* P x *) | 
| 658 | |> implies_intr ihyp | |
| 659 | |> implies_intr (cprop_of x_D) | |
| 660 | |> forall_intr (cterm_of thy x) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 661 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 662 | val subset_induct_rule = | 
| 22166 | 663 | acc_subset_induct | 
| 664 | |> (curry op COMP) (assume D_subset) | |
| 665 | |> (curry op COMP) (assume D_dcl) | |
| 666 | |> (curry op COMP) (assume a_D) | |
| 667 | |> (curry op COMP) istep | |
| 668 | |> fold_rev implies_intr steps | |
| 669 | |> implies_intr a_D | |
| 670 | |> implies_intr D_dcl | |
| 671 | |> implies_intr D_subset | |
| 672 | ||
| 673 | val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule | |
| 674 | ||
| 675 | val simple_induct_rule = | |
| 676 | subset_induct_rule | |
| 677 | |> forall_intr (cterm_of thy D) | |
| 678 | |> forall_elim (cterm_of thy acc_R) | |
| 679 | |> assume_tac 1 |> Seq.hd | |
| 680 | |> (curry op COMP) (acc_downward | |
| 681 | |> (instantiate' [SOME (ctyp_of thy domT)] | |
| 682 | (map (SOME o cterm_of thy) [R, x, z])) | |
| 683 | |> forall_intr (cterm_of thy z) | |
| 684 | |> forall_intr (cterm_of thy x)) | |
| 685 | |> forall_intr (cterm_of thy a) | |
| 686 | |> forall_intr (cterm_of thy P) | |
| 687 | in | |
| 26529 | 688 | simple_induct_rule | 
| 22166 | 689 | end | 
| 690 | ||
| 691 | ||
| 692 | ||
| 693 | (* FIXME: This should probably use fixed goals, to be more reliable and faster *) | |
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 694 | fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
 | 
| 22166 | 695 | let | 
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 696 | val thy = ProofContext.theory_of ctxt | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 697 |       val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
 | 
| 22166 | 698 | qglr = (oqs, _, _, _), ...} = clause | 
| 23819 | 699 | val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs) | 
| 22166 | 700 | |> fold_rev (curry Logic.mk_implies) gs | 
| 701 | |> cterm_of thy | |
| 702 | in | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 703 | Goal.init goal | 
| 22166 | 704 | |> (SINGLE (resolve_tac [accI] 1)) |> the | 
| 26653 | 705 | |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the | 
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 706 | |> (SINGLE (auto_tac (local_clasimpset_of ctxt))) |> the | 
| 22166 | 707 | |> Goal.conclude | 
| 708 | |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) | |
| 709 | end | |
| 710 | ||
| 711 | ||
| 712 | ||
| 713 | (** Termination rule **) | |
| 714 | ||
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 715 | val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
 | 
| 23880 | 716 | val wf_in_rel = @{thm FunDef.wf_in_rel};
 | 
| 717 | val in_rel_def = @{thm FunDef.in_rel_def};
 | |
| 22166 | 718 | |
| 719 | fun mk_nest_term_case thy globals R' ihyp clause = | |
| 720 | let | |
| 721 |       val Globals {x, z, ...} = globals
 | |
| 722 |       val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
 | |
| 723 | qglr=(oqs, _, _, _), ...} = clause | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 724 | |
| 22166 | 725 | val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 726 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 727 | fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) = | 
| 22166 | 728 | let | 
| 26115 | 729 | val used = map (fn (ctx,thm) => FundefCtxTree.export_thm thy ctx thm) (u @ sub) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 730 | |
| 23819 | 731 | val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs) | 
| 26196 | 732 | |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 733 | |> FundefCtxTree.export_term (fixes, assumes) | 
| 22166 | 734 | |> fold_rev (curry Logic.mk_implies o prop_of) ags | 
| 735 | |> fold_rev mk_forall_rename (map fst oqs ~~ qs) | |
| 736 | |> cterm_of thy | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 737 | |
| 22166 | 738 | val thm = assume hyp | 
| 739 | |> fold forall_elim cqs | |
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24960diff
changeset | 740 | |> fold Thm.elim_implies ags | 
| 26196 | 741 | |> FundefCtxTree.import_thm thy (fixes, assumes) | 
| 742 | |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *) | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 743 | |
| 26196 | 744 | val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg)))) | 
| 22166 | 745 | |
| 26196 | 746 | val acc = thm COMP ih_case | 
| 747 | val z_acc_local = acc | |
| 748 | |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection))))) | |
| 749 | ||
| 750 | val ethm = z_acc_local | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 751 | |> FundefCtxTree.export_thm thy (fixes, | 
| 26196 | 752 | z_eq_arg :: case_hyp :: ags @ assumes) | 
| 22166 | 753 | |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) | 
| 754 | ||
| 755 | val sub' = sub @ [(([],[]), acc)] | |
| 756 | in | |
| 757 | (sub', (hyp :: hyps, ethm :: thms)) | |
| 758 | end | |
| 759 | | step _ _ _ _ = raise Match | |
| 760 | in | |
| 761 | FundefCtxTree.traverse_tree step tree | |
| 762 | end | |
| 22325 | 763 | |
| 764 | ||
| 22166 | 765 | fun mk_nest_term_rule thy globals R R_cases clauses = | 
| 766 | let | |
| 767 |       val Globals { domT, x, z, ... } = globals
 | |
| 768 | val acc_R = mk_acc domT R | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 769 | |
| 22166 | 770 |       val R' = Free ("R", fastype_of R)
 | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 771 | |
| 23819 | 772 |       val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
 | 
| 773 |       val inrel_R = Const ("FunDef.in_rel", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
 | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 774 | |
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
26653diff
changeset | 775 |       val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name "Wellfounded.wfP"}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
 | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 776 | |
| 22166 | 777 | (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 778 |       val ihyp = Term.all domT $ Abs ("z", domT,
 | 
| 27336 | 779 | Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x), | 
| 780 | HOLogic.mk_Trueprop (acc_R $ Bound 0))) | |
| 22166 | 781 | |> cterm_of thy | 
| 782 | ||
| 26653 | 783 | val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0 | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 784 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 785 | val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x)) | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 786 | |
| 22166 | 787 | val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[]) | 
| 788 | in | |
| 789 | R_cases | |
| 790 | |> forall_elim (cterm_of thy z) | |
| 791 | |> forall_elim (cterm_of thy x) | |
| 792 | |> forall_elim (cterm_of thy (acc_R $ z)) | |
| 793 | |> curry op COMP (assume R_z_x) | |
| 794 | |> fold_rev (curry op COMP) cases | |
| 795 | |> implies_intr R_z_x | |
| 796 | |> forall_intr (cterm_of thy z) | |
| 797 | |> (fn it => it COMP accI) | |
| 798 | |> implies_intr ihyp | |
| 799 | |> forall_intr (cterm_of thy x) | |
| 800 | |> (fn it => Drule.compose_single(it,2,wf_induct_rule)) | |
| 801 | |> curry op RS (assume wfR') | |
| 22325 | 802 | |> forall_intr_vars | 
| 803 | |> (fn it => it COMP allI) | |
| 22166 | 804 | |> fold implies_intr hyps | 
| 805 | |> implies_intr wfR' | |
| 806 | |> forall_intr (cterm_of thy R') | |
| 807 | |> forall_elim (cterm_of thy (inrel_R)) | |
| 808 | |> curry op RS wf_in_rel | |
| 809 | |> full_simplify (HOL_basic_ss addsimps [in_rel_def]) | |
| 810 | |> forall_intr (cterm_of thy Rrel) | |
| 811 | end | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 812 | |
| 22166 | 813 | |
| 814 | ||
| 815 | (* Tail recursion (probably very fragile) | |
| 816 | * | |
| 817 | * FIXME: | |
| 818 | * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context. | |
| 819 | * - Must we really replace the fvar by f here? | |
| 820 | * - Splitting is not configured automatically: Problems with case? | |
| 821 | *) | |
| 822 | fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps = | |
| 823 | let | |
| 824 |       val Globals {domT, ranT, fvar, ...} = globals
 | |
| 825 | ||
| 26653 | 826 | val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *) | 
| 22166 | 827 | |
| 828 | val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *) | |
| 829 |           Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
 | |
| 830 |                      (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
 | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 831 |                      (fn {prems=[a], ...} =>
 | 
| 22166 | 832 | ((rtac (G_induct OF [a])) | 
| 833 | THEN_ALL_NEW (rtac accI) | |
| 834 | THEN_ALL_NEW (etac R_cases) | |
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 835 | THEN_ALL_NEW (asm_full_simp_tac (local_simpset_of octxt))) 1) | 
| 22166 | 836 | |
| 837 | val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value) | |
| 838 | ||
| 839 | fun mk_trsimp clause psimp = | |
| 840 | let | |
| 841 |             val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
 | |
| 842 | val thy = ProofContext.theory_of ctxt | |
| 843 | val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 844 | |
| 22166 | 845 | val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *) | 
| 846 | val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *) | |
| 23819 | 847 | fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def]) | 
| 22166 | 848 | in | 
| 849 | Goal.prove ctxt [] [] trsimp | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 850 | (fn _ => | 
| 22166 | 851 | rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1 | 
| 26653 | 852 | THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1 | 
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 853 | THEN (simp_default_tac (local_simpset_of ctxt) 1) | 
| 22279 
b0d482a9443f
Adapted to changes in Accessible_Part and Wellfounded_Recursion theories.
 berghofe parents: 
22166diff
changeset | 854 | THEN (etac not_acc_down 1) | 
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 855 | THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (local_simpset_of ctxt))) 1) | 
| 22166 | 856 | |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) | 
| 857 | end | |
| 858 | in | |
| 859 | map2 mk_trsimp clauses psimps | |
| 860 | end | |
| 861 | ||
| 862 | ||
| 23189 | 863 | fun prepare_fundef config defname [((fname, fT), mixfix)] abstract_qglrs lthy = | 
| 22166 | 864 | let | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 865 |       val FundefConfig {domintros, tailrec, default=default_str, ...} = config
 | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 866 | |
| 22166 | 867 | val fvar = Free (fname, fT) | 
| 868 | val domT = domain_type fT | |
| 869 | val ranT = range_type fT | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 870 | |
| 24693 | 871 | val default = Syntax.parse_term lthy default_str | 
| 872 | |> TypeInfer.constrain fT |> Syntax.check_term lthy | |
| 22166 | 873 | |
| 874 | val (globals, ctxt') = fix_globals domT ranT fvar lthy | |
| 875 | ||
| 876 |       val Globals { x, h, ... } = globals
 | |
| 877 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 878 | val clauses = map (mk_clause_context x ctxt') abstract_qglrs | 
| 22166 | 879 | |
| 880 | val n = length abstract_qglrs | |
| 881 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 882 |       fun build_tree (ClauseContext { ctxt, rhs, ...}) =
 | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23880diff
changeset | 883 | FundefCtxTree.mk_tree (fname, fT) h ctxt rhs | 
| 22166 | 884 | |
| 885 | val trees = map build_tree clauses | |
| 886 | val RCss = map find_calls trees | |
| 887 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 888 | val ((G, GIntro_thms, G_elim, G_induct), lthy) = | 
| 22166 | 889 | PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy | 
| 890 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 891 | val ((f, f_defthm), lthy) = | 
| 28004 
c8642f498aa3
function package: name primitive defs "f_sumC_def" instead of "f_sum_def" to avoid clashes
 krauss parents: 
27336diff
changeset | 892 | PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy | 
| 22166 | 893 | |
| 894 | val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss | |
| 895 | val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees | |
| 896 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 897 | val ((R, RIntro_thmss, R_elim), lthy) = | 
| 22166 | 898 | PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy | 
| 899 | ||
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 900 | val (_, lthy) = | 
| 28965 | 901 | LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy | 
| 22166 | 902 | |
| 903 | val newthy = ProofContext.theory_of lthy | |
| 904 | val clauses = map (transfer_clause_ctx newthy) clauses | |
| 905 | ||
| 906 | val cert = cterm_of (ProofContext.theory_of lthy) | |
| 907 | ||
| 908 | val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss | |
| 909 | ||
| 910 | val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume | |
| 911 | val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume) | |
| 912 | ||
| 913 | val compat_store = store_compat_thms n compat | |
| 914 | ||
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23880diff
changeset | 915 | val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm | 
| 22166 | 916 | |
| 917 | val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses | |
| 918 | ||
| 919 | fun mk_partial_rules provedgoal = | |
| 920 | let | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 921 | val newthy = theory_of_thm provedgoal (*FIXME*) | 
| 22166 | 922 | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 923 | val (graph_is_function, complete_thm) = | 
| 22166 | 924 | provedgoal | 
| 925 | |> Conjunction.elim | |
| 26653 | 926 | |> apfst (Thm.forall_elim_vars 0) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 927 | |
| 22166 | 928 | val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 929 | |
| 22166 | 930 | val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 931 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 932 | val simple_pinduct = PROFILE "Proving partial induction rule" | 
| 22166 | 933 | (mk_partial_induct_rule newthy globals R complete_thm) xclauses | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 934 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 935 | |
| 22166 | 936 | val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 937 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 938 | val dom_intros = if domintros | 
| 30607 
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
 wenzelm parents: 
29329diff
changeset | 939 | then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses) | 
| 22166 | 940 | else NONE | 
| 941 | val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE | |
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 942 | |
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 943 | in | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 944 |             FundefResult {fs=[f], G=G, R=R, cases=complete_thm,
 | 
| 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28004diff
changeset | 945 | psimps=psimps, simple_pinducts=[simple_pinduct], | 
| 22166 | 946 | termination=total_intro, trsimps=trsimps, | 
| 947 | domintros=dom_intros} | |
| 948 | end | |
| 949 | in | |
| 950 | ((f, goalstate, mk_partial_rules), lthy) | |
| 951 | end | |
| 952 | ||
| 953 | ||
| 954 | end |