| author | wenzelm | 
| Sun, 19 Mar 2017 14:43:54 +0100 | |
| changeset 65326 | cb7cb57c7ce1 | 
| parent 64705 | 7596b0736ab9 | 
| child 67313 | a2d7c0987f19 | 
| 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 | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 12 | val warning_at: Proof.context -> term list -> string -> unit | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 13 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 14 | val excess_equations: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 15 | val extra_variable_in_rhs: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 16 | val ill_formed_corec_call: Proof.context -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 17 | val ill_formed_equation_head: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 18 | val ill_formed_equation_lhs_rhs: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 19 | val ill_formed_equation: Proof.context -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 20 | val ill_formed_formula: Proof.context -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 21 | val ill_formed_rec_call: Proof.context -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 22 | val inconstant_pattern_pos_for_fun: Proof.context -> term list -> string -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 23 | val invalid_map: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 24 | val missing_args_to_fun_on_lhs: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 25 | val missing_equations_for_const: string -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 26 | val missing_equations_for_fun: string -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 27 | val missing_pattern: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 28 | val more_than_one_nonvar_in_lhs: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 29 | val multiple_equations_for_ctr: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 30 | val nonprimitive_corec: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 31 | val nonprimitive_pattern_in_lhs: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 32 | val not_codatatype: Proof.context -> typ -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 33 | val not_datatype: Proof.context -> typ -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 34 | val not_constructor_in_pattern: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 35 | val not_constructor_in_rhs: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 36 | val rec_call_not_apply_to_ctr_arg: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 37 | val partially_applied_ctr_in_pattern: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 38 | val partially_applied_ctr_in_rhs: Proof.context -> term list -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 39 | val too_few_args_in_rec_call: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 40 | val unexpected_rec_call_in: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 41 | val unexpected_corec_call_in: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 42 | val unsupported_case_around_corec_call: Proof.context -> term list -> term -> 'a | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 43 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 44 | val no_equation_for_ctr_warning: Proof.context -> term list -> term -> unit | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 45 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 46 | val check_all_fun_arg_frees: Proof.context -> term list -> term list -> unit | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 47 | val check_duplicate_const_names: binding list -> unit | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 48 | val check_duplicate_variables_in_lhs: Proof.context -> term list -> term list -> unit | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 49 | val check_top_sort: Proof.context -> binding -> typ -> unit | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 50 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 51 | datatype fp_kind = Least_FP | Greatest_FP | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 52 | |
| 55966 | 53 | 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 | 54 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 55 | type fp_rec_sugar = | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 56 |     {transfers: bool list,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 57 | fun_names: string list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 58 | funs: term list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 59 | fun_defs: thm list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 60 | fpTs: typ list} | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 61 | |
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 62 | val morph_fp_rec_sugar: morphism -> fp_rec_sugar -> fp_rec_sugar | 
| 59281 | 63 | 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 | 64 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 65 | 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 | 66 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 67 | 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 | 68 | val indexedd: 'a list list -> int -> int list list * int | 
| 54299 | 69 | 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 | 70 | 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 | 71 | 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 | 72 |   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 | 73 |   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 | 74 | |
| 62687 | 75 |   val order_strong_conn: ('a * 'a -> bool) -> ((('a * unit) * 'a list) list -> 'b) ->
 | 
| 76 |     ('b -> 'a list) -> ('a * 'a list) list -> 'a list list -> 'a list list
 | |
| 77 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 78 | val mk_common_name: string list -> string | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 79 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 80 | val num_binder_types: typ -> int | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 81 | 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 | 82 | val exists_strict_subtype_in: typ list -> typ -> bool | 
| 55571 | 83 | val tvar_subst: theory -> typ list -> typ list -> ((string * int) * typ) list | 
| 84 | ||
| 56651 | 85 | val retype_const_or_free: typ -> term -> term | 
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 86 | val drop_all: term -> term | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 87 | val permute_args: int -> term -> term | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 88 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 89 | 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 | 90 | 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 | 91 | 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 | 92 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
56651diff
changeset | 93 | val mk_co_rec: theory -> fp_kind -> typ list -> typ -> term -> term | 
| 55571 | 94 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 95 | val mk_conjunctN: int -> int -> thm | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 96 | val conj_dests: int -> thm -> thm list | 
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 97 | |
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 98 | val print_def_consts: bool -> (term * (string * thm)) list -> Proof.context -> unit | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 99 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 100 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 101 | 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 | 102 | struct | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 103 | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 104 | fun error_at ctxt ats str = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 105 | error (str ^ (if null ats then "" | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 106 | else " at\n" ^ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats))); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 107 | fun warning_at ctxt ats str = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 108 | warning (str ^ (if null ats then "" | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 109 | else " at\n" ^ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats))); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 110 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 111 | fun excess_equations ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 112 |   error ("Excess equation(s):\n" ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 113 | cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats)); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 114 | fun extra_variable_in_rhs ctxt ats var = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 115 |   error_at ctxt ats ("Extra variable " ^ quote (Syntax.string_of_term ctxt var) ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 116 | " in right-hand side"); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 117 | fun ill_formed_corec_call ctxt t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 118 |   error ("Ill-formed corecursive call " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 119 | fun ill_formed_equation_head ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 120 | error_at ctxt ats "Ill-formed function equation (expected function name on left-hand side)"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 121 | fun ill_formed_equation_lhs_rhs ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 122 | error_at ctxt ats "Ill-formed equation (expected \"lhs = rhs\")"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 123 | fun ill_formed_equation ctxt t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 124 |   error_at ctxt [] ("Ill-formed equation:\n  " ^ Syntax.string_of_term ctxt t);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 125 | fun ill_formed_formula ctxt t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 126 |   error_at ctxt [] ("Ill-formed formula:\n  " ^ Syntax.string_of_term ctxt t);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 127 | fun ill_formed_rec_call ctxt t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 128 |   error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 129 | fun inconstant_pattern_pos_for_fun ctxt ats fun_name = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 130 |   error_at ctxt ats ("Inconstant constructor pattern position for function " ^ quote fun_name);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 131 | fun invalid_map ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 132 |   error_at ctxt ats ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 133 | fun missing_args_to_fun_on_lhs ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 134 | error_at ctxt ats "Expected more arguments to function on left-hand side"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 135 | fun missing_equations_for_const fun_name = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 136 |   error ("Missing equations for constant " ^ quote fun_name);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 137 | fun missing_equations_for_fun fun_name = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 138 |   error ("Missing equations for function " ^ quote fun_name);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 139 | fun missing_pattern ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 140 | error_at ctxt ats "Constructor pattern missing in left-hand side"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 141 | fun more_than_one_nonvar_in_lhs ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 142 | error_at ctxt ats "More than one non-variable argument in left-hand side"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 143 | fun multiple_equations_for_ctr ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 144 |   error ("Multiple equations for constructor:\n" ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 145 | cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats)); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 146 | fun nonprimitive_corec ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 147 | error_at ctxt ats "Nonprimitive corecursive specification"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 148 | fun nonprimitive_pattern_in_lhs ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 149 | error_at ctxt ats "Nonprimitive pattern in left-hand side"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 150 | fun not_codatatype ctxt T = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 151 |   error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 152 | fun not_datatype ctxt T = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 153 |   error ("Not a datatype: " ^ Syntax.string_of_typ ctxt T);
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 154 | fun not_constructor_in_pattern ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 155 |   error_at ctxt ats ("Not a constructor " ^ quote (Syntax.string_of_term ctxt t) ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 156 | " in pattern"); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 157 | fun not_constructor_in_rhs ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 158 |   error_at ctxt ats ("Not a constructor " ^ quote (Syntax.string_of_term ctxt t) ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 159 | " in right-hand side"); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 160 | fun rec_call_not_apply_to_ctr_arg ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 161 |   error_at ctxt ats ("Recursive call not directly applied to constructor argument in " ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 162 | quote (Syntax.string_of_term ctxt t)); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 163 | fun partially_applied_ctr_in_pattern ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 164 | error_at ctxt ats "Partially applied constructor in pattern"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 165 | fun partially_applied_ctr_in_rhs ctxt ats = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 166 | error_at ctxt ats "Partially applied constructor in right-hand side"; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 167 | fun too_few_args_in_rec_call ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 168 |   error_at ctxt ats ("Too few arguments in recursive call " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 169 | fun unexpected_rec_call_in ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 170 |   error_at ctxt ats ("Unexpected recursive call in " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 171 | fun unexpected_corec_call_in ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 172 |   error_at ctxt ats ("Unexpected corecursive call in " ^ quote (Syntax.string_of_term ctxt t));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 173 | fun unsupported_case_around_corec_call ctxt ats t = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 174 |   error_at ctxt ats ("Unsupported corecursive call under case expression " ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 175 | quote (Syntax.string_of_term ctxt t) ^ "\n(Define " ^ | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 176 | quote (Syntax.string_of_typ ctxt (domain_type (fastype_of t))) ^ | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 177 | " with discriminators and selectors to circumvent this limitation)"); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 178 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 179 | fun no_equation_for_ctr_warning ctxt ats ctr = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 180 |   warning_at ctxt ats ("No equation for constructor " ^ quote (Syntax.string_of_term ctxt ctr));
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 181 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 182 | fun check_all_fun_arg_frees ctxt ats fun_args = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 183 | (case find_first (not o is_Free) fun_args of | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 184 |     SOME t => error_at ctxt ats ("Non-variable function argument on left-hand side " ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 185 | quote (Syntax.string_of_term ctxt t)) | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 186 | | NONE => | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 187 | (case find_first (Variable.is_fixed ctxt o fst o dest_Free) fun_args of | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 188 |       SOME t => error_at ctxt ats ("Function argument " ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 189 | quote (Syntax.string_of_term ctxt t) ^ " is fixed in context") | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 190 | | NONE => ())); | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 191 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 192 | fun check_duplicate_const_names bs = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 193 | let val dups = duplicates (op =) (map Binding.name_of bs) in | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 194 |     ignore (null dups orelse error ("Duplicate constant name " ^ quote (hd dups)))
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 195 | end; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 196 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 197 | fun check_duplicate_variables_in_lhs ctxt ats vars = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 198 | let val dups = duplicates (op aconv) vars in | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 199 | ignore (null dups orelse | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 200 |       error_at ctxt ats ("Duplicable variable " ^ quote (Syntax.string_of_term ctxt (hd dups)) ^
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 201 | " in left-hand side")) | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 202 | end; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 203 | |
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 204 | fun check_top_sort ctxt b T = | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 205 |   ignore (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort type}) orelse
 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 206 |     error ("Type of " ^ Binding.print b ^ " contains top sort"));
 | 
| 59595 | 207 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 208 | datatype fp_kind = Least_FP | Greatest_FP; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 209 | |
| 55966 | 210 | fun case_fp Least_FP l _ = l | 
| 211 | | case_fp Greatest_FP _ g = g; | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 212 | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 213 | type fp_rec_sugar = | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 214 |   {transfers: bool list,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 215 | fun_names: string list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 216 | funs: term list, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 217 | fun_defs: thm list, | 
| 59281 | 218 | fpTs: typ list}; | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 219 | |
| 59314 | 220 | 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 | 221 |   {transfers = transfers,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 222 | fun_names = fun_names, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 223 | funs = map (Morphism.term phi) funs, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 224 | fun_defs = map (Morphism.thm phi) fun_defs, | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 225 | fpTs = map (Morphism.typ phi) fpTs}; | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
58337diff
changeset | 226 | |
| 59281 | 227 | val transfer_fp_rec_sugar = morph_fp_rec_sugar o Morphism.transfer_morphism; | 
| 228 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 229 | fun flat_rec_arg_args xss = | 
| 59281 | 230 | (* FIXME (once the old datatype package is completely phased out): The first line below gives the | 
| 231 | 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 | 232 | (* flat xss *) | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 233 | map hd xss @ maps tl xss; | 
| 55571 | 234 | |
| 54614 | 235 | fun indexe _ h = (h, h + 1); | 
| 236 | 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 | 237 | 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 | 238 | 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 | 239 | 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 | 240 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 241 | 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 | 242 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 243 | 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 | 244 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 245 | fun find_indices eq xs ys = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 246 | 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 | 247 | |
| 62687 | 248 | fun order_strong_conn eq make_graph topological_order deps sccs = | 
| 249 | let | |
| 250 | val normals = maps (fn x :: xs => map (fn y => (y, x)) xs) sccs; | |
| 251 | fun normal s = AList.lookup eq normals s |> the_default s; | |
| 252 | ||
| 253 | val normal_deps = deps | |
| 254 | |> map (fn (x, xs) => let val x' = normal x in | |
| 255 | (x', fold (insert eq o normal) xs [] |> remove eq x') | |
| 256 | end) | |
| 257 | |> AList.group eq | |
| 258 | |> map (apsnd (fn xss => fold (union eq) xss [])); | |
| 259 | ||
| 260 | val normal_G = make_graph (map (apfst (rpair ())) normal_deps); | |
| 261 | val ordered_normals = rev (topological_order normal_G); | |
| 262 | in | |
| 263 | map (fn x => the (find_first (fn (y :: _) => eq (y, x)) sccs)) ordered_normals | |
| 264 | end; | |
| 265 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 266 | val mk_common_name = space_implode "_"; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 267 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 268 | 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 | 269 | | num_binder_types _ = 0; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 270 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 271 | val exists_subtype_in = Term.exists_subtype o member (op =); | 
| 58337 | 272 | 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 | 273 | |
| 55571 | 274 | fun tvar_subst thy Ts Us = | 
| 275 | Vartab.fold (cons o apsnd snd) (fold (Sign.typ_match thy) (Ts ~~ Us) Vartab.empty) []; | |
| 276 | ||
| 56651 | 277 | fun retype_const_or_free T (Const (s, _)) = Const (s, T) | 
| 278 | | retype_const_or_free T (Free (s, _)) = Free (s, T) | |
| 279 |   | 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 | 280 | |
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54614diff
changeset | 281 | fun drop_all t = | 
| 56245 | 282 |   subst_bounds (strip_qnt_vars @{const_name Pure.all} t |> map Free |> rev,
 | 
| 283 |     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 | 284 | |
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 285 | fun permute_args n t = | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 286 | 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 | 287 | |
| 55856 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 blanchet parents: 
55772diff
changeset | 288 | 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 | 289 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 290 | 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 | 291 | | 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 | 292 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 293 | 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 | 294 | 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 | 295 | 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 | 296 | end; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 297 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 298 | val mk_comp = mk_compN 1; | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 299 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
56651diff
changeset | 300 | fun mk_co_rec thy fp Cs fpT t = | 
| 55571 | 301 | let | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 302 | val ((f_Cs, prebody), body) = strip_type (fastype_of t) |>> split_last; | 
| 55966 | 303 | val fpT0 = case_fp fp prebody body; | 
| 304 | val Cs0 = distinct (op =) (map (case_fp fp body_type domain_type) f_Cs); | |
| 55571 | 305 | val rho = tvar_subst thy (fpT0 :: Cs0) (fpT :: Cs); | 
| 306 | in | |
| 307 | Term.subst_TVars rho t | |
| 308 | end; | |
| 309 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 310 | 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 | 311 | | mk_conjunctN _ 1 = conjunct1 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 312 | | mk_conjunctN 2 2 = conjunct2 | 
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 313 | | 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 | 314 | |
| 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 315 | 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 | 316 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 317 | fun print_def_consts int defs ctxt = | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 318 | Proof_Display.print_consts int (Position.thread_data ()) ctxt (K false) | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 319 | (map_filter (try (dest_Free o fst)) defs); | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
62687diff
changeset | 320 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 321 | end; |