| author | webertj | 
| Fri, 15 Sep 2006 18:06:51 +0200 | |
| changeset 20544 | 893e7a9546ff | 
| parent 20534 | b147d0c13f6e | 
| child 20654 | d80502f0d701 | 
| permissions | -rw-r--r-- | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 1 | (* Title: HOL/Tools/function_package/mutual.ML | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 2 | ID: $Id$ | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 3 | Author: Alexander Krauss, TU Muenchen | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 4 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 5 | A package for general recursive function definitions. | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 6 | Tools for mutual recursive definitions. | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 7 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 8 | *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 9 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 10 | signature FUNDEF_MUTUAL = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 11 | sig | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 12 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 13 | val prepare_fundef_mutual : ((string * typ) * mixfix) list | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 14 | -> term list | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 15 | -> local_theory | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 16 | -> ((FundefCommon.mutual_info * string * FundefCommon.prep_result) * local_theory) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 17 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 18 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 19 | val mk_partial_rules_mutual : Proof.context -> FundefCommon.mutual_info -> FundefCommon.prep_result -> thm -> | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 20 | FundefCommon.fundef_mresult | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 21 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 22 | val sort_by_function : FundefCommon.mutual_info -> string list -> 'a list -> 'a list list | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 23 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 24 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 25 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 26 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 27 | structure FundefMutual: FUNDEF_MUTUAL = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 28 | struct | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 29 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 30 | open FundefCommon | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 31 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 32 | (* Theory dependencies *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 33 | val sum_case_rules = thms "Datatype.sum.cases" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 34 | val split_apply = thm "Product_Type.split" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 35 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 36 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 37 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 38 | fun mutual_induct_Pnames n = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 39 | if n < 5 then fst (chop n ["P","Q","R","S"]) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 40 | else map (fn i => "P" ^ string_of_int i) (1 upto n) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 41 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 42 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 43 | fun check_head fs t = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 44 | if (case t of | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 45 | (Free (n, _)) => n mem fs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 46 | | _ => false) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 47 | then dest_Free t | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 48 | else raise ERROR "Head symbol of every left hand side must be the new function." (* FIXME: Output the equation here *) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 49 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 50 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 51 | fun open_all_all (Const ("all", _) $ Abs (n, T, b)) = apfst (cons (n, T)) (open_all_all b)
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 52 | | open_all_all t = ([], t) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 53 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 54 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 55 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 56 | (* Builds a curried clause description in abstracted form *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 57 | fun split_def fnames geq = | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 58 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 59 | val (qs, imp) = open_all_all geq | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 60 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 61 | val gs = Logic.strip_imp_prems imp | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 62 | val eq = Logic.strip_imp_concl imp | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 63 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 64 | val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 65 | val (fc, args) = strip_comb f_args | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 66 | val f as (fname, _) = check_head fnames fc | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 67 | |
| 20534 
b147d0c13f6e
Fixed Subscript Exception occurring with Higher-Order recursion
 krauss parents: 
20523diff
changeset | 68 | fun add_bvs t is = add_loose_bnos (t, 0, is) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 69 | val rhs_only = (add_bvs rhs [] \\ fold add_bvs args []) | 
| 20534 
b147d0c13f6e
Fixed Subscript Exception occurring with Higher-Order recursion
 krauss parents: 
20523diff
changeset | 70 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 71 | |> map (fst o nth (rev qs)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 72 | val _ = if null rhs_only then () | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 73 | else raise ERROR "Variables occur on right hand side only." (* FIXME: Output vars *) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 74 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 75 | ((f, length args), (fname, qs, gs, args, rhs)) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 76 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 77 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 78 | fun get_part fname = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 79 |     the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 80 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 81 | (* FIXME *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 82 | fun mk_prod_abs e (t1, t2) = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 83 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 84 | val bTs = rev (map snd e) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 85 | val T1 = fastype_of1 (bTs, t1) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 86 | val T2 = fastype_of1 (bTs, t2) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 87 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 88 | HOLogic.pair_const T1 T2 $ t1 $ t2 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 89 | end; | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 90 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 91 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 92 | fun analyze_eqs ctxt fnames eqs = | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 93 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 94 | (* FIXME: Add check for number of arguments | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 95 | fun all_equal ((x as ((n:string,T), k:int))::xs) = if forall (fn ((n',_),k') => n = n' andalso k = k') xs then x | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 96 | 							   else raise ERROR ("All equations in a block must describe the same "
 | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 97 | ^ "function and have the same number of arguments.") | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 98 | *) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 99 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 100 | val (consts, fqgars) = split_list (map (split_def fnames) eqs) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 101 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 102 | val different_consts = distinct (eq_fst (eq_fst eq_str)) consts | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 103 | val cnames = map (fst o fst) different_consts | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 104 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 105 | val check_rcs = exists_subterm (fn Free (n, _) => if n mem cnames | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 106 | then raise ERROR "Recursive Calls not allowed in premises." else false | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 107 | | _ => false) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 108 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 109 | val _ = forall (fn (_, _, gs, _, _) => forall check_rcs gs) fqgars | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 110 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 111 | fun curried_types ((_,T), k) = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 112 | let | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 113 | val (caTs, uaTs) = chop k (binder_types T) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 114 | in | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 115 | (caTs, uaTs ---> body_type T) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 116 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 117 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 118 | val (caTss, resultTs) = split_list (map curried_types different_consts) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 119 | val argTs = map (foldr1 HOLogic.mk_prodT) caTss | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 120 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 121 | val (RST,streeR, pthsR) = SumTools.mk_tree_distinct resultTs | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 122 | val (ST, streeA, pthsA) = SumTools.mk_tree argTs | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 123 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 124 | val def_name = foldr1 (fn (a,b) => a ^ "_" ^ b) (map Sign.base_name cnames) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 125 | val fsum_type = ST --> RST | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 126 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 127 | val ([fsum_var_name], _) = Variable.add_fixes [ def_name ^ "_sum" ] ctxt | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 128 | val fsum_var = (fsum_var_name, fsum_type) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 129 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 130 | fun define (fvar as (n, T), _) caTs pthA pthR = | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 131 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 132 | 		val vars = map_index (fn (i,T) => Free ("x" ^ string_of_int i, T)) caTs (* FIXME: Bind xs properly *)
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 133 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 134 | val f_exp = SumTools.mk_proj streeR pthR (Free fsum_var $ SumTools.mk_inj streeA pthA (foldr1 HOLogic.mk_prod vars)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 135 | val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 136 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 137 | val rew = (n, fold_rev lambda vars f_exp) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 138 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 139 | 		(MutualPart {fvar=fvar,cargTs=caTs,pthA=pthA,pthR=pthR,f_def=def,f=NONE,f_defthm=NONE}, rew)
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 140 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 141 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 142 | val (parts, rews) = split_list (map4 define different_consts caTss pthsA pthsR) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 143 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 144 | fun convert_eqs (f, qs, gs, args, rhs) = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 145 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 146 |               val MutualPart {pthA, pthR, ...} = get_part f parts
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 147 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 148 | (qs, gs, SumTools.mk_inj streeA pthA (foldr1 (mk_prod_abs qs) args), | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 149 | SumTools.mk_inj streeR pthR (replace_frees rews rhs) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 150 | |> Envir.norm_term (Envir.empty 0)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 151 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 152 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 153 | val qglrs = map convert_eqs fqgars | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 154 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 155 | 	Mutual {defname=def_name,fsum_var=fsum_var, ST=ST, RST=RST, streeA=streeA, streeR=streeR, 
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 156 | parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE} | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 157 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 158 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 159 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 160 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 161 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 162 | fun define_projections fixes mutual fsum lthy = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 163 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 164 |       fun def ((MutualPart {fvar=(fname, fT), cargTs, pthA, pthR, f_def, ...}), (_, mixfix)) lthy =
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 165 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 166 | val ((f, (_, f_defthm)), lthy') = LocalTheory.def ((fname, mixfix), | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 167 | ((fname ^ "_def", []), Term.subst_bound (fsum, f_def))) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 168 | lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 169 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 170 |             (MutualPart {fvar=(fname, fT), cargTs=cargTs, pthA=pthA, pthR=pthR, f_def=f_def, 
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 171 | f=SOME f, f_defthm=SOME f_defthm }, | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 172 | lthy') | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 173 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 174 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 175 |       val Mutual { defname, fsum_var, ST, RST, streeA, streeR, parts, fqgars, qglrs, ... } = mutual
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 176 | val (parts', lthy') = fold_map def (parts ~~ fixes) lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 177 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 178 |       (Mutual { defname=defname, fsum_var=fsum_var, ST=ST, RST=RST, streeA=streeA, streeR=streeR, parts=parts', 
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 179 | fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum }, | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 180 | lthy') | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 181 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 182 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 183 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 184 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 185 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 186 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 187 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 188 | fun prepare_fundef_mutual fixes eqss lthy = | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 189 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 190 | val mutual = analyze_eqs lthy (map (fst o fst) fixes) eqss | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 191 | 	val Mutual {defname, fsum_var=(n, T), qglrs, ...} = mutual
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 192 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 193 | val (prep_result, fsum, lthy') = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 194 | FundefPrep.prepare_fundef defname (n, T, NoSyn) qglrs lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 195 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 196 | val (mutual', lthy'') = define_projections fixes mutual fsum lthy' | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 197 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 198 | ((mutual', defname, prep_result), lthy'') | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 199 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 200 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 201 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 202 | (* Beta-reduce both sides of a meta-equality *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 203 | fun beta_norm_eq thm = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 204 | let | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 205 | val (lhs, rhs) = dest_equals (cprop_of thm) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 206 | val lhs_conv = beta_conversion false lhs | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 207 | val rhs_conv = beta_conversion false rhs | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 208 | in | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 209 | transitive (symmetric lhs_conv) (transitive thm rhs_conv) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 210 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 211 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 212 | fun beta_reduce thm = Thm.equal_elim (Thm.beta_conversion true (cprop_of thm)) thm | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 213 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 214 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 215 | fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F = | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 216 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 217 | val thy = ProofContext.theory_of ctxt | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19781diff
changeset | 218 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 219 | val oqnames = map fst pre_qs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 220 | val (qs, ctxt') = Variable.invent_fixes oqnames ctxt | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 221 | |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 222 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 223 | fun inst t = subst_bounds (rev qs, t) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 224 | val gs = map inst pre_gs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 225 | val args = map inst pre_args | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 226 | val rhs = inst pre_rhs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 227 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 228 | val cqs = map (cterm_of thy) qs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 229 | val ags = map (assume o cterm_of thy) gs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 230 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 231 | val import = fold forall_elim cqs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 232 | #> fold implies_elim_swp ags | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 233 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 234 | val export = fold_rev (implies_intr o cprop_of) ags | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 235 | #> fold_rev forall_intr_rename (oqnames ~~ cqs) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 236 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 237 | F (f, qs, gs, args, rhs) import export | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 238 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 239 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 240 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 241 | fun recover_mutual_psimp thy RST streeR all_f_defs parts (f, _, _, args, _) import (export : thm -> thm) sum_psimp_eq = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 242 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 243 |       val (MutualPart {f_defthm=SOME f_def, pthR, ...}) = get_part f parts
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 244 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 245 | val psimp = import sum_psimp_eq | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 246 | val (simp, restore_cond) = case cprems_of psimp of | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 247 | [] => (psimp, I) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 248 | | [cond] => (implies_elim psimp (assume cond), implies_intr cond) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 249 | | _ => sys_error "Too many conditions" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 250 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 251 |       val x = Free ("x", RST)
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 252 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 253 | val f_def_inst = fold (fn arg => fn thm => combination thm (reflexive (cterm_of thy arg))) args (Thm.freezeT f_def) (* FIXME *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 254 | |> beta_reduce | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 255 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 256 | reflexive (cterm_of thy (lambda x (SumTools.mk_proj streeR pthR x))) (* PR(x) == PR(x) *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 257 | |> (fn it => combination it (simp RS eq_reflection)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 258 | |> beta_norm_eq (* PR(S(I(as))) == PR(IR(...)) *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 259 | |> transitive f_def_inst (* f ... == PR(IR(...)) *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 260 | |> simplify (HOL_basic_ss addsimps [SumTools.projl_inl, SumTools.projr_inr]) (* f ... == ... *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 261 | |> simplify (HOL_basic_ss addsimps all_f_defs) (* f ... == ... *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 262 | |> (fn it => it RS meta_eq_to_obj_eq) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 263 | |> restore_cond | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 264 | |> export | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 265 | end | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 266 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 267 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 268 | (* FIXME HACK *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 269 | fun mk_applied_form ctxt caTs thm = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 270 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 271 | val thy = ProofContext.theory_of ctxt | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 272 |       val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 273 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 274 | fold (fn x => fn thm => combination thm (reflexive x)) xs thm | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 275 | |> beta_reduce | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 276 | |> fold_rev forall_intr xs | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 277 | |> forall_elim_vars 0 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 278 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 279 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 280 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 281 | fun mutual_induct_rules thy induct all_f_defs (Mutual {RST, parts, streeA, ...}) =
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 282 | let | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 283 | 	fun mk_P (MutualPart {cargTs, ...}) Pname =
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 284 | let | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 285 | 		val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 286 | val atup = foldr1 HOLogic.mk_prod avars | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 287 | in | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 288 | tupled_lambda atup (list_comb (Free (Pname, cargTs ---> HOLogic.boolT), avars)) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 289 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 290 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 291 | val Ps = map2 mk_P parts (mutual_induct_Pnames (length parts)) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 292 | val case_exp = SumTools.mk_sumcases streeA HOLogic.boolT Ps | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 293 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 294 | val induct_inst = | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 295 | forall_elim (cterm_of thy case_exp) induct | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 296 | |> full_simplify (HOL_basic_ss addsimps (split_apply :: sum_case_rules)) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 297 | |> full_simplify (HOL_basic_ss addsimps all_f_defs) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 298 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 299 | 	fun mk_proj rule (MutualPart {cargTs, pthA, ...}) =
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 300 | let | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 301 | 		val afs = map_index (fn (i,T) => Free ("a" ^ string_of_int i, T)) cargTs
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 302 | val inj = SumTools.mk_inj streeA pthA (foldr1 HOLogic.mk_prod afs) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 303 | in | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 304 | rule | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 305 | |> forall_elim (cterm_of thy inj) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 306 | |> full_simplify (HOL_basic_ss addsimps (split_apply :: sum_case_rules)) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 307 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 308 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 309 | in | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 310 | map (mk_proj induct_inst) parts | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 311 | end | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 312 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 313 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 314 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 315 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 316 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 317 | fun mk_partial_rules_mutual lthy (m as Mutual {RST, parts, streeR, fqgars, ...}) data prep_result =
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 318 | let | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 319 | val thy = ProofContext.theory_of lthy | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 320 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 321 | val result = FundefProof.mk_partial_rules thy data prep_result | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 322 |       val FundefResult {f, G, R, completeness, psimps, subset_pinduct,simple_pinduct,total_intro,dom_intros} = result
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 323 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 324 |       val all_f_defs = map (fn MutualPart {f_defthm = SOME f_def, cargTs, ...} => 
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 325 | mk_applied_form lthy cargTs (symmetric (Thm.freezeT f_def))) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 326 | parts | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 327 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 328 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 329 | fun mk_mpsimp fqgar sum_psimp = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 330 | in_context lthy fqgar (recover_mutual_psimp thy RST streeR all_f_defs parts) sum_psimp | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 331 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 332 | val mpsimps = map2 mk_mpsimp fqgars psimps | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 333 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 334 | val minducts = mutual_induct_rules thy simple_pinduct all_f_defs m | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 335 | val termination = full_simplify (HOL_basic_ss addsimps all_f_defs) total_intro | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 336 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 337 |       FundefMResult { f=f, G=G, R=R,
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 338 | psimps=mpsimps, subset_pinducts=[subset_pinduct], simple_pinducts=minducts, | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 339 | cases=completeness, termination=termination, domintros=dom_intros } | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 340 | end | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 341 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 342 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 343 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 344 | (* puts an object in the "right bucket" *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 345 | fun store_grouped P x [] = [] | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 346 | | store_grouped P x ((l, xs)::bs) = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 347 | if P (x, l) then ((l, x::xs)::bs) else ((l, xs)::store_grouped P x bs) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 348 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 349 | fun sort_by_function (Mutual {fqgars, ...}) names xs =
 | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 350 | fold_rev (store_grouped (eq_str o apfst fst)) (* fill *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 351 | (map name_of_fqgar fqgars ~~ xs) (* the name-thm pairs *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 352 | (map (rpair []) names) (* in the empty buckets labeled with names *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 353 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 354 | |> map (snd #> map snd) (* and remove the labels afterwards *) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 355 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19922diff
changeset | 356 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 357 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 358 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 359 | end | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 360 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 361 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 362 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 363 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 364 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 365 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 366 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 367 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 368 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 369 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 370 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 371 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 372 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 373 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 374 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 375 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 376 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 377 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 378 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 379 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 380 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 381 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 382 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 383 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: diff
changeset | 384 |