| author | eberlm | 
| Tue, 05 Jan 2016 17:54:21 +0100 | |
| changeset 62066 | 4db2d39aa76c | 
| parent 59595 | 2d90b85b9264 | 
| child 62687 | 1c4842b32bfb | 
| permissions | -rw-r--r-- | 
| 55061 | 1 | (* Title: HOL/Tools/BNF/bnf_fp_rec_sugar_util.ML | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 2 | Author: Lorenz Panny, TU Muenchen | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 4 | Copyright 2013 | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 5 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 6 | Library for recursor and corecursor sugar. | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 7 | *) | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 8 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 9 | signature BNF_FP_REC_SUGAR_UTIL = | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 10 | sig | 
| 59595 | 11 | val error_at: Proof.context -> term list -> string -> 'a | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 12 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 13 | datatype fp_kind = Least_FP | Greatest_FP | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 14 | |
| 55966 | 15 | val case_fp: fp_kind -> 'a -> 'a -> 'a | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 16 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 17 | type fp_rec_sugar = | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 18 |     {transfers: bool list,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 19 | fun_names: string list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 20 | funs: term list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 21 | fun_defs: thm list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 22 | fpTs: typ list} | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 23 | |
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 24 | val morph_fp_rec_sugar: morphism -> fp_rec_sugar -> fp_rec_sugar | 
| 59281 | 25 | val transfer_fp_rec_sugar: theory -> fp_rec_sugar -> fp_rec_sugar | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 26 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 27 | val flat_rec_arg_args: 'a list list -> 'a list | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 28 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 29 | val indexed: 'a list -> int -> int list * int | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 30 | val indexedd: 'a list list -> int -> int list list * int | 
| 54299 | 31 | val indexeddd: 'a list list list -> int -> int list list list * int | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 32 | val indexedddd: 'a list list list list -> int -> int list list list list * int | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 33 | val find_index_eq: ''a list -> ''a -> int | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 34 |   val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
 | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 35 |   val find_indices: ('b * 'a -> bool) -> 'a list -> 'b list -> int list
 | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 36 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 37 | val mk_common_name: string list -> string | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 38 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 39 | val num_binder_types: typ -> int | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 40 | val exists_subtype_in: typ list -> typ -> bool | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 41 | val exists_strict_subtype_in: typ list -> typ -> bool | 
| 55571 | 42 | val tvar_subst: theory -> typ list -> typ list -> ((string * int) * typ) list | 
| 43 | ||
| 56651 | 44 | val retype_const_or_free: typ -> term -> term | 
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 45 | val drop_all: term -> term | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 46 | val permute_args: int -> term -> term | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 47 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 48 | val mk_partial_compN: int -> typ -> term -> term | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 49 | val mk_compN: int -> typ list -> term * term -> term | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 50 | val mk_comp: typ list -> term * term -> term | 
| 54202 
0a06b51ffa56
handle applied ctor arguments gracefully when computing 'callssss' (for recursion through functions)
 blanchet parents: 
54200diff
changeset | 51 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
56651diff
changeset | 52 | val mk_co_rec: theory -> fp_kind -> typ list -> typ -> term -> term | 
| 55571 | 53 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 54 | val mk_conjunctN: int -> int -> thm | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 55 | val conj_dests: int -> thm -> thm list | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 56 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 57 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 58 | structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL = | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 59 | struct | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 60 | |
| 59595 | 61 | fun error_at ctxt ts str = | 
| 62 | error (str ^ (if null ts then "" | |
| 63 | else " at\n " ^ space_implode "\n " (map (quote o Syntax.string_of_term ctxt) ts))); | |
| 64 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 65 | datatype fp_kind = Least_FP | Greatest_FP; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 66 | |
| 55966 | 67 | fun case_fp Least_FP l _ = l | 
| 68 | | case_fp Greatest_FP _ g = g; | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 69 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 70 | type fp_rec_sugar = | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 71 |   {transfers: bool list,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 72 | fun_names: string list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 73 | funs: term list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 74 | fun_defs: thm list, | 
| 59281 | 75 | fpTs: typ list}; | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 76 | |
| 59314 | 77 | fun morph_fp_rec_sugar phi ({transfers, fun_names, funs, fun_defs, fpTs} : fp_rec_sugar) =
 | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 78 |   {transfers = transfers,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 79 | fun_names = fun_names, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 80 | funs = map (Morphism.term phi) funs, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 81 | fun_defs = map (Morphism.thm phi) fun_defs, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 82 | fpTs = map (Morphism.typ phi) fpTs}; | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 83 | |
| 59281 | 84 | val transfer_fp_rec_sugar = morph_fp_rec_sugar o Morphism.transfer_morphism; | 
| 85 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 86 | fun flat_rec_arg_args xss = | 
| 59281 | 87 | (* FIXME (once the old datatype package is completely phased out): The first line below gives the | 
| 88 | preferred order. The second line is for compatibility with the old datatype package. *) | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 89 | (* flat xss *) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 90 | map hd xss @ maps tl xss; | 
| 55571 | 91 | |
| 54614 | 92 | fun indexe _ h = (h, h + 1); | 
| 93 | fun indexed xs = fold_map indexe xs; | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 94 | fun indexedd xss = fold_map indexed xss; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 95 | fun indexeddd xsss = fold_map indexedd xsss; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 96 | fun indexedddd xssss = fold_map indexeddd xssss; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 97 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 98 | fun find_index_eq hs h = find_index (curry (op =) h) hs; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 99 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 100 | fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 101 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 102 | fun find_indices eq xs ys = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 103 | map_filter I (map_index (fn (i, y) => if member eq xs y then SOME i else NONE) ys); | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 104 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 105 | val mk_common_name = space_implode "_"; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 106 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 107 | (*stolen from ~~/src/HOL/Tools/Nitpick/nitpick_hol.ML*) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 108 | fun num_binder_types (Type (@{type_name fun}, [_, T])) = 1 + num_binder_types T
 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 109 | | num_binder_types _ = 0; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 110 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 111 | val exists_subtype_in = Term.exists_subtype o member (op =); | 
| 58337 | 112 | fun exists_strict_subtype_in Ts T = exists_subtype_in (remove (op =) T Ts) T; | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 113 | |
| 55571 | 114 | fun tvar_subst thy Ts Us = | 
| 115 | Vartab.fold (cons o apsnd snd) (fold (Sign.typ_match thy) (Ts ~~ Us) Vartab.empty) []; | |
| 116 | ||
| 56651 | 117 | fun retype_const_or_free T (Const (s, _)) = Const (s, T) | 
| 118 | | retype_const_or_free T (Free (s, _)) = Free (s, T) | |
| 119 |   | retype_const_or_free _ t = raise TERM ("retype_const_or_free", [t]);
 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 120 | |
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 121 | fun drop_all t = | 
| 56245 | 122 |   subst_bounds (strip_qnt_vars @{const_name Pure.all} t |> map Free |> rev,
 | 
| 123 |     strip_qnt_body @{const_name Pure.all} t);
 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 124 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 125 | fun permute_args n t = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 126 | list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n); | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 127 | |
| 55856 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 blanchet parents: 
55772diff
changeset | 128 | fun mk_partial_comp fT g = fst (Term.dest_comb (HOLogic.mk_comp (g, Free (Name.uu, fT)))); | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 129 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 130 | fun mk_partial_compN 0 _ g = g | 
| 55856 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 blanchet parents: 
55772diff
changeset | 131 | | mk_partial_compN n fT g = mk_partial_comp fT (mk_partial_compN (n - 1) (range_type fT) g); | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 132 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 133 | fun mk_compN n bound_Ts (g, f) = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 134 | let val typof = curry fastype_of1 bound_Ts in | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 135 | mk_partial_compN n (typof f) g $ f | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 136 | end; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 137 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 138 | val mk_comp = mk_compN 1; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 139 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
56651diff
changeset | 140 | fun mk_co_rec thy fp Cs fpT t = | 
| 55571 | 141 | let | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 142 | val ((f_Cs, prebody), body) = strip_type (fastype_of t) |>> split_last; | 
| 55966 | 143 | val fpT0 = case_fp fp prebody body; | 
| 144 | val Cs0 = distinct (op =) (map (case_fp fp body_type domain_type) f_Cs); | |
| 55571 | 145 | val rho = tvar_subst thy (fpT0 :: Cs0) (fpT :: Cs); | 
| 146 | in | |
| 147 | Term.subst_TVars rho t | |
| 148 | end; | |
| 149 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 150 | fun mk_conjunctN 1 1 = @{thm TrueE[OF TrueI]}
 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 151 | | mk_conjunctN _ 1 = conjunct1 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 152 | | mk_conjunctN 2 2 = conjunct2 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 153 | | mk_conjunctN n m = conjunct2 RS (mk_conjunctN (n - 1) (m - 1)); | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 154 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 155 | fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n); | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 156 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 157 | end; |