| author | blanchet | 
| Tue, 11 Mar 2014 15:34:38 +0100 | |
| changeset 56046 | 683148f3ae48 | 
| parent 55966 | 972f0aa7091b | 
| child 56245 | 84fc7dfa3cd4 | 
| 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 | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 11 | datatype fp_kind = Least_FP | Greatest_FP | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 12 | |
| 55966 | 13 | 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 | 14 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 15 | 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 | 16 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 17 | 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 | 18 | val indexedd: 'a list list -> int -> int list list * int | 
| 54299 | 19 | 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 | 20 | 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 | 21 | 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 | 22 |   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 | 23 |   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 | 24 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 25 | val mk_common_name: string list -> string | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 26 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 27 | val num_binder_types: typ -> int | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 28 | 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 | 29 | val exists_strict_subtype_in: typ list -> typ -> bool | 
| 55571 | 30 | val tvar_subst: theory -> typ list -> typ list -> ((string * int) * typ) list | 
| 31 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 32 | val retype_free: typ -> term -> term | 
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 33 | val drop_all: term -> term | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 34 | val permute_args: int -> term -> term | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 35 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 36 | 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 | 37 | 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 | 38 | 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 | 39 | |
| 55869 | 40 | val mk_co_rec: theory -> fp_kind -> typ -> typ list -> term -> term | 
| 55571 | 41 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 42 | val mk_conjunctN: int -> int -> thm | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 43 | 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 | 44 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 45 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 46 | 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 | 47 | struct | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 48 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 49 | datatype fp_kind = Least_FP | Greatest_FP; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 50 | |
| 55966 | 51 | fun case_fp Least_FP l _ = l | 
| 52 | | case_fp Greatest_FP _ g = g; | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 53 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 54 | fun flat_rec_arg_args xss = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 55 | (* FIXME (once the old datatype package is phased out): The first line below gives the preferred | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 56 | order. The second line is for compatibility with the old datatype package. *) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 57 | (* flat xss *) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 58 | map hd xss @ maps tl xss; | 
| 55571 | 59 | |
| 54614 | 60 | fun indexe _ h = (h, h + 1); | 
| 61 | 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 | 62 | 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 | 63 | 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 | 64 | 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 | 65 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 66 | 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 | 67 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 68 | 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 | 69 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 70 | fun find_indices eq xs ys = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 71 | 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 | 72 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 73 | val mk_common_name = space_implode "_"; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 74 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 75 | (*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 | 76 | 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 | 77 | | num_binder_types _ = 0; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 78 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 79 | val exists_subtype_in = Term.exists_subtype o member (op =); | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 80 | fun exists_strict_subtype_in Ts T = exists_subtype_in (filter_out (curry (op =) T) Ts) T; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 81 | |
| 55571 | 82 | fun tvar_subst thy Ts Us = | 
| 83 | Vartab.fold (cons o apsnd snd) (fold (Sign.typ_match thy) (Ts ~~ Us) Vartab.empty) []; | |
| 84 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 85 | fun retype_free T (Free (s, _)) = Free (s, T) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 86 |   | retype_free _ t = raise TERM ("retype_free", [t]);
 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 87 | |
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 88 | fun drop_all t = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 89 |   subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 90 |     strip_qnt_body @{const_name all} t);
 | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 91 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 92 | fun permute_args n t = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 93 | 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 | 94 | |
| 55856 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 blanchet parents: 
55772diff
changeset | 95 | 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 | 96 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 97 | 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 | 98 | | 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 | 99 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 100 | 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 | 101 | 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 | 102 | 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 | 103 | end; | 
| 
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_comp = mk_compN 1; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 106 | |
| 55869 | 107 | fun mk_co_rec thy fp fpT Cs t = | 
| 55571 | 108 | let | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 109 | val ((f_Cs, prebody), body) = strip_type (fastype_of t) |>> split_last; | 
| 55966 | 110 | val fpT0 = case_fp fp prebody body; | 
| 111 | val Cs0 = distinct (op =) (map (case_fp fp body_type domain_type) f_Cs); | |
| 55571 | 112 | val rho = tvar_subst thy (fpT0 :: Cs0) (fpT :: Cs); | 
| 113 | in | |
| 114 | Term.subst_TVars rho t | |
| 115 | end; | |
| 116 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 117 | 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 | 118 | | mk_conjunctN _ 1 = conjunct1 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 119 | | mk_conjunctN 2 2 = conjunct2 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 120 | | 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 | 121 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 122 | 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 | 123 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 124 | end; |