| author | wenzelm | 
| Sun, 19 Mar 2017 14:43:54 +0100 | |
| changeset 65326 | cb7cb57c7ce1 | 
| parent 64705 | 7596b0736ab9 | 
| child 66251 | cd935b7cb3fb | 
| permissions | -rw-r--r-- | 
| 55061 | 1 | (* Title: HOL/Tools/BNF/bnf_gfp_rec_sugar.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 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 53303 
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 | |
| 55538 | 6 | Corecursor sugar ("primcorec" and "primcorecursive").
 | 
| 53303 
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 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 9 | signature BNF_GFP_REC_SUGAR = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 10 | sig | 
| 59281 | 11 | datatype corec_option = | 
| 12 | Plugins_Option of Proof.context -> Plugin_Name.filter | | |
| 13 | Sequential_Option | | |
| 14 | Exhaustive_Option | | |
| 15 | Transfer_Option | |
| 54899 
7a01387c47d5
added tactic to prove 'disc_iff' properties in 'primcorec'
 blanchet parents: 
54883diff
changeset | 16 | |
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 17 | datatype corec_call = | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 18 | Dummy_No_Corec of int | | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 19 | No_Corec of int | | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 20 | Mutual_Corec of int * int * int | | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 21 | Nested_Corec of int | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 22 | |
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 23 | type corec_ctr_spec = | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 24 |     {ctr: term,
 | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 25 | disc: term, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 26 | sels: term list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 27 | pred: int option, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 28 | calls: corec_call list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 29 | discI: thm, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 30 | sel_thms: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 31 | distinct_discss: thm list list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 32 | collapse: thm, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 33 | corec_thm: thm, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 34 | corec_disc: thm, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 35 | corec_sels: thm list} | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 36 | |
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 37 | type corec_spec = | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 38 |     {T: typ,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 39 | corec: term, | 
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 40 | exhaust_discs: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 41 | sel_defs: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 42 | fp_nesting_maps: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 43 | fp_nesting_map_ident0s: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 44 | fp_nesting_map_comps: thm list, | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 45 | ctr_specs: corec_ctr_spec list} | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 46 | |
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 47 | val abstract_over_list: term list -> term -> term | 
| 59945 | 48 | val abs_tuple_balanced: term list -> term -> term | 
| 49 | ||
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 50 | val mk_conjs: term list -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 51 | val mk_disjs: term list -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 52 | val mk_dnf: term list list -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 53 | val conjuncts_s: term -> term list | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 54 | val s_not: term -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 55 | val s_not_conj: term list -> term list | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 56 | val s_conjs: term list -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 57 | val s_disjs: term list -> term | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 58 | val s_dnf: term list list -> term list | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 59 | |
| 62686 | 60 | val case_of: Proof.context -> string -> (string * bool) option | 
| 59674 | 61 | val fold_rev_let_if_case: Proof.context -> (term list -> term -> 'a -> 'a) -> typ list -> | 
| 62 | term -> 'a -> 'a | |
| 63 | val massage_let_if_case: Proof.context -> (term -> bool) -> (typ list -> term -> term) -> | |
| 62318 | 64 | (typ list -> term -> unit) -> (typ list -> term -> term) -> typ list -> term -> term | 
| 59674 | 65 | val massage_nested_corec_call: Proof.context -> (term -> bool) -> | 
| 59989 | 66 | (typ list -> typ -> typ -> term -> term) -> (typ list -> typ -> typ -> term -> term) -> | 
| 67 | typ list -> typ -> typ -> term -> term | |
| 60683 | 68 | val expand_to_ctr_term: Proof.context -> typ -> term -> term | 
| 59674 | 69 | val massage_corec_code_rhs: Proof.context -> (typ list -> term -> term list -> term) -> | 
| 70 | typ list -> term -> term | |
| 71 | val fold_rev_corec_code_rhs: Proof.context -> (term list -> term -> term list -> 'a -> 'a) -> | |
| 72 | typ list -> term -> 'a -> 'a | |
| 73 | val case_thms_of_term: Proof.context -> term -> | |
| 74 | thm list * thm list * thm list * thm list * thm list | |
| 75 | val map_thms_of_type: Proof.context -> typ -> thm list | |
| 76 | ||
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 77 | val corec_specs_of: binding list -> typ list -> typ list -> term list -> | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 78 | (term * term list list) list list -> local_theory -> | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 79 | corec_spec list * typ list * thm * thm * thm list * thm list * (Token.src list * Token.src list) | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 80 | * bool * local_theory | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 81 | |
| 59281 | 82 | val gfp_rec_sugar_interpretation: string -> | 
| 83 | (BNF_FP_Rec_Sugar_Util.fp_rec_sugar -> local_theory -> local_theory) -> theory -> theory | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 84 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 85 | val primcorec_ursive: bool -> bool -> corec_option list -> ((binding * typ) * mixfix) list -> | 
| 63719 | 86 | ((binding * Token.T list list) * term) list -> term option list -> Proof.context -> | 
| 87 | (term * 'a list) list list * (thm list list -> local_theory -> local_theory) * local_theory | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 88 | val primcorec_ursive_cmd: bool -> bool -> corec_option list -> | 
| 61301 | 89 | (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list -> | 
| 90 | Proof.context -> | |
| 91 | (term * 'a list) list list * (thm list list -> local_theory -> local_theory) * local_theory | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 92 | val primcorecursive_cmd: bool -> corec_option list -> | 
| 53831 
80423b9080cf
support "of" syntax to disambiguate selector equations
 panny parents: 
53830diff
changeset | 93 | (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list -> | 
| 
80423b9080cf
support "of" syntax to disambiguate selector equations
 panny parents: 
53830diff
changeset | 94 | Proof.context -> Proof.state | 
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 95 | val primcorec_cmd: bool -> corec_option list -> | 
| 53831 
80423b9080cf
support "of" syntax to disambiguate selector equations
 panny parents: 
53830diff
changeset | 96 | (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list -> | 
| 
80423b9080cf
support "of" syntax to disambiguate selector equations
 panny parents: 
53830diff
changeset | 97 | local_theory -> local_theory | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 98 | end; | 
| 
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 | structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 101 | struct | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 102 | |
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 103 | open Ctr_Sugar_General_Tactics | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 104 | open Ctr_Sugar | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 105 | open BNF_Util | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 106 | open BNF_Def | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 107 | open BNF_FP_Util | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 108 | open BNF_FP_Def_Sugar | 
| 54243 | 109 | open BNF_FP_N2M_Sugar | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 110 | open BNF_FP_Rec_Sugar_Util | 
| 59276 | 111 | open BNF_FP_Rec_Sugar_Transfer | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 112 | open BNF_GFP_Rec_Sugar_Tactics | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 113 | |
| 58387 | 114 | val codeN = "code"; | 
| 115 | val ctrN = "ctr"; | |
| 116 | val discN = "disc"; | |
| 117 | val disc_iffN = "disc_iff"; | |
| 118 | val excludeN = "exclude"; | |
| 119 | val selN = "sel"; | |
| 53791 | 120 | |
| 54145 
297d1c603999
make sure that registered code equations are actually equations
 blanchet parents: 
54133diff
changeset | 121 | val nitpicksimp_attrs = @{attributes [nitpick_simp]};
 | 
| 53794 | 122 | val simp_attrs = @{attributes [simp]};
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 123 | |
| 59597 | 124 | fun use_primcorecursive () = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59873diff
changeset | 125 |   error ("\"auto\" failed (try " ^ quote (#1 @{command_keyword primcorecursive}) ^ " instead of " ^
 | 
| 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59873diff
changeset | 126 |     quote (#1 @{command_keyword primcorec}) ^ ")");
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 127 | |
| 59281 | 128 | datatype corec_option = | 
| 129 | Plugins_Option of Proof.context -> Plugin_Name.filter | | |
| 130 | Sequential_Option | | |
| 131 | Exhaustive_Option | | |
| 132 | Transfer_Option; | |
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 133 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 134 | datatype corec_call = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 135 | Dummy_No_Corec of int | | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 136 | No_Corec of int | | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 137 | Mutual_Corec of int * int * int | | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 138 | Nested_Corec of int; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 139 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 140 | type basic_corec_ctr_spec = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 141 |   {ctr: term,
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 142 | disc: term, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 143 | sels: term list}; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 144 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 145 | type corec_ctr_spec = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 146 |   {ctr: term,
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 147 | disc: term, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 148 | sels: term list, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 149 | pred: int option, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 150 | calls: corec_call list, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 151 | discI: thm, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 152 | sel_thms: thm list, | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 153 | distinct_discss: thm list list, | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 154 | collapse: thm, | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 155 | corec_thm: thm, | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 156 | corec_disc: thm, | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 157 | corec_sels: thm list}; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 158 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 159 | type corec_spec = | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 160 |   {T: typ,
 | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 161 | corec: term, | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 162 | exhaust_discs: thm list, | 
| 56858 | 163 | sel_defs: thm list, | 
| 57397 | 164 | fp_nesting_maps: thm list, | 
| 57399 | 165 | fp_nesting_map_ident0s: thm list, | 
| 57397 | 166 | fp_nesting_map_comps: thm list, | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 167 | ctr_specs: corec_ctr_spec list}; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 168 | |
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 169 | exception NO_MAP of term; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 170 | |
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 171 | fun abstract_over_list rev_vs = | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 172 | let | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 173 | val vs = rev rev_vs; | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 174 | |
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 175 | fun abs n (t $ u) = abs n t $ abs n u | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 176 | | abs n (Abs (s, T, t)) = Abs (s, T, abs (n + 1) t) | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 177 | | abs n t = | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 178 | let val j = find_index (curry (op =) t) vs in | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 179 | if j < 0 then t else Bound (n + j) | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 180 | end; | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 181 | in | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 182 | abs 0 | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 183 | end; | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 184 | |
| 59945 | 185 | val abs_tuple_balanced = HOLogic.tupled_lambda o mk_tuple_balanced; | 
| 186 | ||
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 187 | fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) =
 | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 188 | Ts ---> T; | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 189 | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 190 | fun sort_list_duplicates xs = map snd (sort (int_ord o apply2 fst) xs); | 
| 54948 | 191 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 192 | val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 193 | val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
 | 
| 54900 | 194 | val mk_dnf = mk_disjs o map mk_conjs; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 195 | |
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 196 | val conjuncts_s = filter_out (curry (op aconv) @{const True}) o HOLogic.conjuncts;
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 197 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 198 | fun s_not @{const True} = @{const False}
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 199 |   | s_not @{const False} = @{const True}
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 200 |   | s_not (@{const Not} $ t) = t
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 201 |   | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 202 |   | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 203 |   | s_not t = @{const Not} $ t;
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 204 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 205 | val s_not_conj = conjuncts_s o s_not o mk_conjs; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 206 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 207 | fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 208 | fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 209 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 210 | fun propagate_units css = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 211 | (case List.partition (can the_single) css of | 
| 58393 | 212 | ([], _) => css | 
| 213 | | ([u] :: uss, css') => | |
| 214 | [u] :: propagate_units (map (propagate_unit_neg (s_not u)) | |
| 215 | (map (propagate_unit_pos u) (uss @ css')))); | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 216 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 217 | fun s_conjs cs = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 218 |   if member (op aconv) cs @{const False} then @{const False}
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 219 |   else mk_conjs (remove (op aconv) @{const True} cs);
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 220 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 221 | fun s_disjs ds = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 222 |   if member (op aconv) ds @{const True} then @{const True}
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 223 |   else mk_disjs (remove (op aconv) @{const False} ds);
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 224 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 225 | fun s_dnf css0 = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 226 | let val css = propagate_units css0 in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 227 | if null css then | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 228 |       [@{const False}]
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 229 | else if exists null css then | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 230 | [] | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 231 | else | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 232 | map (fn c :: cs => (c, cs)) css | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 233 | |> AList.coalesce (op =) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 234 | |> map (fn (c, css) => c :: s_dnf css) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 235 | |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)]) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 236 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 237 | |
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 238 | fun fold_rev_let_if_case ctxt f bound_Ts = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 239 | let | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 240 | val thy = Proof_Context.theory_of ctxt; | 
| 53794 | 241 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 242 | fun fld conds t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 243 | (case Term.strip_comb t of | 
| 55343 | 244 |         (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_lets_splits t)
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 245 |       | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 246 | fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 247 | | (Const (c, _), args as _ :: _ :: _) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 248 | let val n = num_binder_types (Sign.the_const_type thy c) - 1 in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 249 | if n >= 0 andalso n < length args then | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 250 | (case fastype_of1 (bound_Ts, nth args n) of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 251 | Type (s, Ts) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 252 | (case dest_case ctxt s Ts t of | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 253 |                 SOME ({split_sels = _ :: _, ...}, conds', branches) =>
 | 
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 254 | fold_rev (uncurry fld) (map (append conds o conjuncts_s) conds' ~~ branches) | 
| 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 255 | | _ => f conds t) | 
| 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 256 | | _ => f conds t) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 257 | else | 
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 258 | f conds t | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 259 | end | 
| 55342 | 260 | | _ => f conds t); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 261 | in | 
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 262 | fld [] | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 263 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 264 | |
| 54970 
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
 blanchet parents: 
54969diff
changeset | 265 | fun case_of ctxt s = | 
| 
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
 blanchet parents: 
54969diff
changeset | 266 | (case ctr_sugar_of ctxt s of | 
| 62318 | 267 |     SOME {casex = Const (s', _), split_sels, ...} => SOME (s', not (null split_sels))
 | 
| 54970 
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
 blanchet parents: 
54969diff
changeset | 268 | | _ => NONE); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 269 | |
| 62318 | 270 | fun massage_let_if_case ctxt has_call massage_leaf unexpected_call unsupported_case bound_Ts t0 = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 271 | let | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 272 | val thy = Proof_Context.theory_of ctxt; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 273 | |
| 60001 | 274 | fun check_no_call bound_Ts t = if has_call t then unexpected_call bound_Ts t else (); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 275 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 276 | fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 277 | | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 278 | | massage_abs bound_Ts m t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 279 | let val T = domain_type (fastype_of1 (bound_Ts, t)) in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 280 | Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0)) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 281 | end | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 282 | and massage_rec bound_Ts t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 283 | let val typof = curry fastype_of1 bound_Ts in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 284 | (case Term.strip_comb t of | 
| 55343 | 285 |           (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_lets_splits t)
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 286 |         | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
 | 
| 59612 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 287 | (case List.partition Term.is_dummy_pattern (map (massage_rec bound_Ts) branches) of | 
| 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 288 | (dummy_branch' :: _, []) => dummy_branch' | 
| 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 289 | | (_, [branch']) => branch' | 
| 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 290 | | (_, branches') => | 
| 60001 | 291 | Term.list_comb (If_const (typof (hd branches')) $ tap (check_no_call bound_Ts) obj, | 
| 292 | branches')) | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61348diff
changeset | 293 |         | (c as Const (@{const_name case_prod}, _), arg :: args) =>
 | 
| 55343 | 294 | massage_rec bound_Ts | 
| 57895 | 295 | (unfold_splits_lets (Term.list_comb (c $ Envir.eta_long bound_Ts arg, args))) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 296 | | (Const (c, _), args as _ :: _ :: _) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 297 | (case try strip_fun_type (Sign.the_const_type thy c) of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 298 | SOME (gen_branch_Ts, gen_body_fun_T) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 299 | let | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 300 | val gen_branch_ms = map num_binder_types gen_branch_Ts; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 301 | val n = length gen_branch_ms; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 302 | in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 303 | if n < length args then | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 304 | (case gen_body_fun_T of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 305 | Type (_, [Type (T_name, _), _]) => | 
| 62318 | 306 | (case case_of ctxt T_name of | 
| 307 | SOME (c', has_split_sels) => | |
| 308 | if c' = c then | |
| 309 | if has_split_sels then | |
| 310 | let | |
| 311 | val (branches, obj_leftovers) = chop n args; | |
| 312 | val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches; | |
| 313 | val branch_Ts' = map typof branches'; | |
| 314 | val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts')); | |
| 315 | val casex' = | |
| 316 | Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T'); | |
| 317 | in | |
| 318 | Term.list_comb (casex', | |
| 319 | branches' @ tap (List.app (check_no_call bound_Ts)) obj_leftovers) | |
| 320 | end | |
| 321 | else | |
| 322 | unsupported_case bound_Ts t | |
| 323 | else | |
| 324 | massage_leaf bound_Ts t | |
| 325 | | NONE => massage_leaf bound_Ts t) | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 326 | | _ => massage_leaf bound_Ts t) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 327 | else | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 328 | massage_leaf bound_Ts t | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 329 | end | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 330 | | NONE => massage_leaf bound_Ts t) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 331 | | _ => massage_leaf bound_Ts t) | 
| 55342 | 332 | end; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 333 | in | 
| 59596 
c067eba942e7
no quick_and_dirty for goals that may fail + tuned messages
 blanchet parents: 
59595diff
changeset | 334 | massage_rec bound_Ts t0 | 
| 59612 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 335 | |> Term.map_aterms (fn t => | 
| 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 336 |       if Term.is_dummy_pattern t then Const (@{const_name undefined}, fastype_of t) else t)
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 337 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 338 | |
| 60001 | 339 | fun massage_let_if_case_corec ctxt has_call massage_leaf bound_Ts t0 = | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 340 | massage_let_if_case ctxt has_call massage_leaf (K (unexpected_corec_call_in ctxt [t0])) | 
| 62318 | 341 | (K (unsupported_case_around_corec_call ctxt [t0])) bound_Ts t0; | 
| 60001 | 342 | |
| 59989 | 343 | fun massage_nested_corec_call ctxt has_call massage_call massage_noncall bound_Ts U T t0 = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 344 | let | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 345 | fun check_no_call t = if has_call t then unexpected_corec_call_in ctxt [t0] t else (); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 346 | |
| 59609 | 347 |     fun massage_mutual_call bound_Ts (Type (@{type_name fun}, [_, U2]))
 | 
| 348 |         (Type (@{type_name fun}, [T1, T2])) t =
 | |
| 62582 | 349 | Abs (Name.uu, T1, massage_mutual_call (T1 :: bound_Ts) U2 T2 (incr_boundvars 1 t $ Bound 0)) | 
| 59609 | 350 | | massage_mutual_call bound_Ts U T t = | 
| 59989 | 351 | (if has_call t then massage_call else massage_noncall) bound_Ts U T t; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 352 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 353 | fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 354 | (case try (dest_map ctxt s) t of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 355 | SOME (map0, fs) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 356 | let | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 357 | val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t)); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 358 | val map' = mk_map (length fs) dom_Ts Us map0; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 359 | val fs' = | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 360 |               map_flattened_map_args ctxt s (@{map 3} (massage_map_or_map_arg bound_Ts) Us Ts) fs;
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 361 | in | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 362 | Term.list_comb (map', fs') | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 363 | end | 
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 364 | | NONE => raise NO_MAP t) | 
| 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 365 | | massage_map _ _ _ t = raise NO_MAP t | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 366 | and massage_map_or_map_arg bound_Ts U T t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 367 | if T = U then | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 368 | tap check_no_call t | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 369 | else | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 370 | massage_map bound_Ts U T t | 
| 58223 
ba7a2d19880c
export useful functions for users of (co)recursors
 blanchet parents: 
58211diff
changeset | 371 | handle NO_MAP _ => massage_mutual_fun bound_Ts U T t | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 372 | and massage_mutual_fun bound_Ts U T t = | 
| 60741 | 373 | let | 
| 374 | val j = Term.maxidx_of_term t + 1; | |
| 375 | val var = Var ((Name.uu, j), domain_type (fastype_of1 (bound_Ts, t))); | |
| 376 | ||
| 377 | fun massage_body () = | |
| 59947 | 378 | Term.lambda var (Term.incr_boundvars 1 (massage_any_call bound_Ts U T | 
| 60741 | 379 | (betapply (t, var)))); | 
| 380 | in | |
| 381 | (case t of | |
| 382 |           Const (@{const_name comp}, _) $ t1 $ t2 =>
 | |
| 383 | if has_call t2 then massage_body () | |
| 384 | else mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, t2) | |
| 385 | | _ => massage_body ()) | |
| 386 | end | |
| 59947 | 387 | and massage_any_call bound_Ts U T = | 
| 60001 | 388 | massage_let_if_case_corec ctxt has_call (fn bound_Ts => fn t => | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 389 | if has_call t then | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 390 | (case U of | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 391 | Type (s, Us) => | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 392 | (case try (dest_ctr ctxt s) t of | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 393 | SOME (f, args) => | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 394 | let | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 395 | val typof = curry fastype_of1 bound_Ts; | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 396 | val f' = mk_ctr Us f | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 397 | val f'_T = typof f'; | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 398 | val arg_Ts = map typof args; | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 399 | in | 
| 59598 | 400 | Term.list_comb (f', | 
| 59947 | 401 |                   @{map 3} (massage_any_call bound_Ts) (binder_types f'_T) arg_Ts args)
 | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 402 | end | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 403 | | NONE => | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 404 | (case t of | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61348diff
changeset | 405 |                 Const (@{const_name case_prod}, _) $ t' =>
 | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 406 | let | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 407 | val U' = curried_type U; | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 408 | val T' = curried_type T; | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 409 | in | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61348diff
changeset | 410 |                   Const (@{const_name case_prod}, U' --> U) $ massage_any_call bound_Ts U' T' t'
 | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 411 | end | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 412 | | t1 $ t2 => | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 413 | (if has_call t2 then | 
| 59989 | 414 | massage_mutual_call bound_Ts U T t | 
| 415 | else | |
| 416 | massage_map bound_Ts U T t1 $ t2 | |
| 417 | handle NO_MAP _ => massage_mutual_call bound_Ts U T t) | |
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 418 | | Abs (s, T', t') => | 
| 59947 | 419 | Abs (s, T', massage_any_call (T' :: bound_Ts) (range_type U) (range_type T) t') | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 420 | | _ => massage_mutual_call bound_Ts U T t)) | 
| 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 421 | | _ => ill_formed_corec_call ctxt t) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 422 | else | 
| 59989 | 423 | massage_noncall bound_Ts U T t) bound_Ts; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 424 | in | 
| 59989 | 425 | (if has_call t0 then massage_any_call else massage_noncall) bound_Ts U T t0 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 426 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 427 | |
| 60683 | 428 | fun expand_to_ctr_term ctxt (T as Type (s, Ts)) t = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 429 | (case ctr_sugar_of ctxt s of | 
| 60683 | 430 |     SOME {ctrs, casex, ...} => Term.list_comb (mk_case Ts T casex, map (mk_ctr Ts) ctrs) $ t
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 431 | | NONE => raise Fail "expand_to_ctr_term"); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 432 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 433 | fun expand_corec_code_rhs ctxt has_call bound_Ts t = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 434 | (case fastype_of1 (bound_Ts, t) of | 
| 60683 | 435 | T as Type (s, _) => | 
| 60001 | 436 | massage_let_if_case_corec ctxt has_call (fn _ => fn t => | 
| 60683 | 437 | if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt T t) bound_Ts t | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 438 | | _ => raise Fail "expand_corec_code_rhs"); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 439 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 440 | fun massage_corec_code_rhs ctxt massage_ctr = | 
| 60001 | 441 | massage_let_if_case_corec ctxt (K false) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 442 | (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 443 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 444 | fun fold_rev_corec_code_rhs ctxt f = | 
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 445 | fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 446 | |
| 55400 | 447 | fun case_thms_of_term ctxt t = | 
| 448 | let val ctr_sugars = map_filter (Ctr_Sugar.ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in | |
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 449 | (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #exhaust_discs ctr_sugars, | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 450 | maps #split_sels ctr_sugars, maps #split_sel_asms ctr_sugars) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 451 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 452 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 453 | fun basic_corec_specs_of ctxt res_T = | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 454 | (case res_T of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 455 | Type (T_name, _) => | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 456 | (case Ctr_Sugar.ctr_sugar_of ctxt T_name of | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 457 | NONE => not_codatatype ctxt res_T | 
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 458 |     | SOME {T = fpT, ctrs, discs, selss, ...} =>
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 459 | let | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 460 | val thy = Proof_Context.theory_of ctxt; | 
| 54272 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 461 | |
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 462 | val As_rho = tvar_subst thy [fpT] [res_T]; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 463 | val substA = Term.subst_TVars As_rho; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 464 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 465 |         fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 466 | in | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 467 |         @{map 3} mk_spec ctrs discs selss
 | 
| 54911 | 468 | handle ListPair.UnequalLengths => not_codatatype ctxt res_T | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 469 | end) | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 470 | | _ => not_codatatype ctxt res_T); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 471 | |
| 59674 | 472 | fun map_thms_of_type ctxt (Type (s, _)) = | 
| 58462 | 473 |     (case fp_sugar_of ctxt s of SOME {fp_bnf_sugar = {map_thms, ...}, ...} => map_thms | NONE => [])
 | 
| 59674 | 474 | | map_thms_of_type _ _ = []; | 
| 54955 
cf8d429dc24e
reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
 blanchet parents: 
54954diff
changeset | 475 | |
| 59281 | 476 | structure GFP_Rec_Sugar_Plugin = Plugin(type T = fp_rec_sugar); | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 477 | |
| 59281 | 478 | fun gfp_rec_sugar_interpretation name f = | 
| 479 | GFP_Rec_Sugar_Plugin.interpretation name (fn fp_rec_sugar => fn lthy => | |
| 480 | f (transfer_fp_rec_sugar (Proof_Context.theory_of lthy) fp_rec_sugar) lthy); | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 481 | |
| 59281 | 482 | val interpret_gfp_rec_sugar = GFP_Rec_Sugar_Plugin.data; | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 483 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 484 | fun corec_specs_of bs arg_Ts res_Ts callers callssss0 lthy0 = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 485 | let | 
| 55005 | 486 | val thy = Proof_Context.theory_of lthy0; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 487 | |
| 57397 | 488 |     val ((missing_res_Ts, perm0_kks, fp_sugars as {fp_nesting_bnfs,
 | 
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 489 |           fp_co_induct_sugar = SOME {common_co_inducts = common_coinduct_thms, ...}, ...} :: _,
 | 
| 58461 | 490 | (_, gfp_sugar_thms)), lthy) = | 
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58286diff
changeset | 491 | nested_to_mutual_fps (K true) Greatest_FP bs res_Ts callers callssss0 lthy0; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 492 | |
| 58286 | 493 | val coinduct_attrs_pair = | 
| 494 | (case gfp_sugar_thms of SOME ((_, attrs_pair), _, _, _, _) => attrs_pair | NONE => ([], [])); | |
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 495 | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 496 | val perm_fp_sugars = sort (int_ord o apply2 #fp_res_index) fp_sugars; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 497 | |
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55538diff
changeset | 498 | val indices = map #fp_res_index fp_sugars; | 
| 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55538diff
changeset | 499 | val perm_indices = map #fp_res_index perm_fp_sugars; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 500 | |
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 501 | val perm_fpTs = map #T perm_fp_sugars; | 
| 58460 | 502 | val perm_ctrXs_Tsss' = | 
| 503 | map (repair_nullary_single_ctr o #ctrXs_Tss o #fp_ctr_sugar) perm_fp_sugars; | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 504 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 505 | val nn0 = length res_Ts; | 
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 506 | val nn = length perm_fpTs; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 507 | val kks = 0 upto nn - 1; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 508 | val perm_ns' = map length perm_ctrXs_Tsss'; | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 509 | |
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 510 | val perm_Ts = map #T perm_fp_sugars; | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 511 | val perm_Xs = map #X perm_fp_sugars; | 
| 58461 | 512 | val perm_Cs = | 
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 513 | map (domain_type o body_fun_type o fastype_of o #co_rec o the o #fp_co_induct_sugar) | 
| 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 514 | perm_fp_sugars; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 515 | val Xs_TCs = perm_Xs ~~ (perm_Ts ~~ perm_Cs); | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 516 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 517 | fun zip_corecT (Type (s, Us)) = [Type (s, map (mk_sumTN o zip_corecT) Us)] | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 518 | | zip_corecT U = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 519 | (case AList.lookup (op =) Xs_TCs U of | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 520 | SOME (T, C) => [T, C] | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 521 | | NONE => [U]); | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 522 | |
| 55869 | 523 | val perm_p_Tss = mk_corec_p_pred_types perm_Cs perm_ns'; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 524 | val perm_f_Tssss = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 525 | map2 (fn C => map (map (map (curry (op -->) C) o zip_corecT))) perm_Cs perm_ctrXs_Tsss'; | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 526 | val perm_q_Tssss = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 527 | map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) perm_f_Tssss; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 528 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 529 | val (perm_p_hss, h) = indexedd perm_p_Tss 0; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 530 | val (perm_q_hssss, h') = indexedddd perm_q_Tssss h; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 531 | val (perm_f_hssss, _) = indexedddd perm_f_Tssss h'; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 532 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 533 | val fun_arg_hs = | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 534 |       flat (@{map 3} flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
 | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 535 | |
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55462diff
changeset | 536 | fun unpermute0 perm0_xs = permute_like_unique (op =) perm0_kks kks perm0_xs; | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55462diff
changeset | 537 | fun unpermute perm_xs = permute_like_unique (op =) perm_indices indices perm_xs; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 538 | |
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55538diff
changeset | 539 | val coinduct_thmss = map (unpermute0 o conj_dests nn) common_coinduct_thms; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 540 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 541 | val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 542 | val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 543 | val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss); | 
| 53358 | 544 | |
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 545 | val f_Tssss = unpermute perm_f_Tssss; | 
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 546 | val fpTs = unpermute perm_fpTs; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 547 | val Cs = unpermute perm_Cs; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 548 | |
| 58187 
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
 blanchet parents: 
58131diff
changeset | 549 | val As_rho = tvar_subst thy (take nn0 fpTs) res_Ts; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 550 | val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 551 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 552 | val substA = Term.subst_TVars As_rho; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 553 | val substAT = Term.typ_subst_TVars As_rho; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 554 | val substCT = Term.typ_subst_TVars Cs_rho; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 555 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 556 | val perm_Cs' = map substCT perm_Cs; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 557 | |
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 558 |     fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
 | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 559 | (if exists_subtype_in Cs T then Nested_Corec | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 560 | else if nullary then Dummy_No_Corec | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 561 | else No_Corec) g_i | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 562 | | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i'); | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 563 | |
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 564 | fun mk_ctr_spec ctr disc sels p_io q_iss f_iss f_Tss discI sel_thms distinct_discss collapse | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 565 | corec_thm corec_disc corec_sels = | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 566 | let val nullary = not (can dest_funT (fastype_of ctr)) in | 
| 54900 | 567 |         {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_io,
 | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 568 |          calls = @{map 3} (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
 | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 569 | distinct_discss = distinct_discss, collapse = collapse, corec_thm = corec_thm, | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 570 | corec_disc = corec_disc, corec_sels = corec_sels} | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 571 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 572 | |
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 573 |     fun mk_ctr_specs ({ctrs, discs, selss, discIs, sel_thmss, distinct_discsss, collapses, ...}
 | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 574 | : ctr_sugar) p_is q_isss f_isss f_Tsss corec_thms corec_discs corec_selss = | 
| 55863 | 575 | let val p_ios = map SOME p_is @ [NONE] in | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 576 |         @{map 14} mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss
 | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 577 | distinct_discsss collapses corec_thms corec_discs corec_selss | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 578 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 579 | |
| 58461 | 580 |     fun mk_spec ({T, fp_ctr_sugar = {ctr_sugar as {exhaust_discs, sel_defs, ...}, ...},
 | 
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 581 |         fp_co_induct_sugar = SOME {co_rec = corec, co_rec_thms = corec_thms,
 | 
| 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 582 | co_rec_discs = corec_discs, co_rec_selss = corec_selss, ...}, ...} : fp_sugar) p_is q_isss | 
| 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63727diff
changeset | 583 | f_isss f_Tsss = | 
| 59598 | 584 |       {T = T, corec = mk_co_rec thy Greatest_FP perm_Cs' (substAT T) corec,
 | 
| 585 | exhaust_discs = exhaust_discs, sel_defs = sel_defs, | |
| 59674 | 586 | fp_nesting_maps = maps (map_thms_of_type lthy o T_of_bnf) fp_nesting_bnfs, | 
| 57399 | 587 | fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs, | 
| 57397 | 588 | fp_nesting_map_comps = map map_comp_of_bnf fp_nesting_bnfs, | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 589 | ctr_specs = mk_ctr_specs ctr_sugar p_is q_isss f_isss f_Tsss corec_thms corec_discs | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 590 | corec_selss}; | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 591 | in | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 592 |     (@{map 5} mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
 | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 593 | co_induct_of common_coinduct_thms, strong_co_induct_of common_coinduct_thms, | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 594 | co_induct_of coinduct_thmss, strong_co_induct_of coinduct_thmss, coinduct_attrs_pair, | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 595 | is_some gfp_sugar_thms, lthy) | 
| 54246 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 596 | end; | 
| 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 blanchet parents: 
54243diff
changeset | 597 | |
| 53358 | 598 | val undef_const = Const (@{const_name undefined}, dummyT);
 | 
| 53357 | 599 | |
| 58393 | 600 | type coeqn_data_disc = | 
| 601 |   {fun_name: string,
 | |
| 602 | fun_T: typ, | |
| 603 | fun_args: term list, | |
| 604 | ctr: term, | |
| 605 | ctr_no: int, | |
| 606 | disc: term, | |
| 607 | prems: term list, | |
| 608 | auto_gen: bool, | |
| 609 | ctr_rhs_opt: term option, | |
| 610 | code_rhs_opt: term option, | |
| 611 | eqn_pos: int, | |
| 612 | user_eqn: term}; | |
| 54001 | 613 | |
| 58393 | 614 | type coeqn_data_sel = | 
| 615 |   {fun_name: string,
 | |
| 616 | fun_T: typ, | |
| 617 | fun_args: term list, | |
| 618 | ctr: term, | |
| 619 | sel: term, | |
| 620 | rhs_term: term, | |
| 621 | ctr_rhs_opt: term option, | |
| 622 | code_rhs_opt: term option, | |
| 623 | eqn_pos: int, | |
| 624 | user_eqn: term}; | |
| 54001 | 625 | |
| 59602 | 626 | fun ctr_sel_of ({ctr, sel, ...} : coeqn_data_sel) = (ctr, sel);
 | 
| 627 | ||
| 54153 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 628 | datatype coeqn_data = | 
| 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 629 | Disc of coeqn_data_disc | | 
| 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 630 | Sel of coeqn_data_sel; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 631 | |
| 59948 | 632 | fun is_free_in frees (Free (s, _)) = member (op =) frees s | 
| 59594 | 633 | | is_free_in _ _ = false; | 
| 634 | ||
| 59948 | 635 | fun is_catch_all_prem (Free (s, _)) = s = Name.uu_ | 
| 59799 
0b21e85fd9ba
clarified role of Name.uu_, which happens to be the internal replacement of the first underscore under certain assumptions about the context;
 wenzelm parents: 
59674diff
changeset | 636 | | is_catch_all_prem _ = false; | 
| 
0b21e85fd9ba
clarified role of Name.uu_, which happens to be the internal replacement of the first underscore under certain assumptions about the context;
 wenzelm parents: 
59674diff
changeset | 637 | |
| 59599 | 638 | fun add_extra_frees ctxt frees names = | 
| 59948 | 639 | fold_aterms (fn x as Free (s, _) => | 
| 640 | (not (member (op =) frees x) andalso not (member (op =) names s) andalso | |
| 641 | not (Variable.is_fixed ctxt s) andalso not (is_catch_all_prem x)) | |
| 59599 | 642 | ? cons x | _ => I); | 
| 643 | ||
| 644 | fun check_extra_frees ctxt frees names t = | |
| 645 | let val bads = add_extra_frees ctxt frees names t [] in | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 646 | null bads orelse extra_variable_in_rhs ctxt [t] (hd bads) | 
| 56152 | 647 | end; | 
| 648 | ||
| 59604 | 649 | fun check_fun_args ctxt eqn fun_args = | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 650 | (check_duplicate_variables_in_lhs ctxt [eqn] fun_args; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 651 | check_all_fun_arg_frees ctxt [eqn] fun_args); | 
| 59604 | 652 | |
| 59594 | 653 | fun dissect_coeqn_disc ctxt fun_names sequentials | 
| 654 | (basic_ctr_specss : basic_corec_ctr_spec list list) eqn_pos ctr_rhs_opt code_rhs_opt prems0 | |
| 655 | concl matchedsss = | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 656 | let | 
| 54272 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 657 | fun find_subterm p = | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 658 | let (* FIXME \<exists>? *) | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 659 | fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v) | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 660 | | find t = if p t then SOME t else NONE; | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 661 | in find end; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 662 | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 663 | val applied_fun = concl | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 664 | |> find_subterm (member (op = o apsnd SOME) fun_names o try (fst o dest_Free o head_of)) | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 665 | |> the | 
| 59594 | 666 | handle Option.Option => error_at ctxt [concl] "Ill-formed discriminator formula"; | 
| 53720 | 667 | val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free; | 
| 56152 | 668 | |
| 59604 | 669 | val _ = check_fun_args ctxt concl fun_args; | 
| 56152 | 670 | |
| 59594 | 671 | val bads = filter (Term.exists_subterm (is_free_in fun_names)) prems0; | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 672 | val _ = null bads orelse unexpected_rec_call_in ctxt [] (hd bads); | 
| 56152 | 673 | |
| 58393 | 674 | val (sequential, basic_ctr_specs) = | 
| 675 | the (AList.lookup (op =) (fun_names ~~ (sequentials ~~ basic_ctr_specss)) fun_name); | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 676 | |
| 54160 | 677 | val discs = map #disc basic_ctr_specs; | 
| 678 | val ctrs = map #ctr basic_ctr_specs; | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 679 |     val not_disc = head_of concl = @{term Not};
 | 
| 53401 | 680 | val _ = not_disc andalso length ctrs <> 2 andalso | 
| 59594 | 681 | error_at ctxt [concl] "Negated discriminator for a type with \<noteq> 2 constructors"; | 
| 54160 | 682 | val disc' = find_subterm (member (op =) discs o head_of) concl; | 
| 54209 | 683 | val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd) | 
| 59598 | 684 | |> (fn SOME t => let val n = find_index (curry (op =) t) ctrs in | 
| 685 | if n >= 0 then SOME n else NONE end | _ => NONE); | |
| 56152 | 686 | |
| 59594 | 687 | val _ = is_none disc' orelse perhaps (try HOLogic.dest_not) concl = the disc' orelse | 
| 688 | error_at ctxt [concl] "Ill-formed discriminator formula"; | |
| 689 | val _ = is_some disc' orelse is_some eq_ctr0 orelse | |
| 690 | error_at ctxt [concl] "No discriminator in equation"; | |
| 56152 | 691 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 692 | val ctr_no' = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 693 | if is_none disc' then the eq_ctr0 else find_index (curry (op =) (head_of (the disc'))) discs; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 694 | val ctr_no = if not_disc then 1 - ctr_no' else ctr_no'; | 
| 54160 | 695 |     val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
 | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 696 | |
| 59605 | 697 | val catch_all = | 
| 698 | (case prems0 of | |
| 699 | [prem] => is_catch_all_prem prem | |
| 700 | | _ => | |
| 701 | if exists is_catch_all_prem prems0 then error_at ctxt [concl] "Superfluous premises" | |
| 702 | else false); | |
| 53720 | 703 | val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default []; | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 704 | val prems = map (abstract_over_list fun_args) prems0; | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 705 | val actual_prems = | 
| 54901 
0b8871677e0b
use same name for feature internally as in user interface, to facilitate grepping
 blanchet parents: 
54900diff
changeset | 706 | (if catch_all orelse sequential then maps s_not_conj matchedss else []) @ | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 707 | (if catch_all then [] else prems); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 708 | |
| 53720 | 709 | val matchedsss' = AList.delete (op =) fun_name matchedsss | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 710 | |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [actual_prems]); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 711 | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 712 | val user_eqn = | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 713 | (actual_prems, concl) | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 714 | |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract_over_list fun_args | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 715 | |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies; | 
| 56152 | 716 | |
| 59599 | 717 | val _ = check_extra_frees ctxt fun_args fun_names user_eqn; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 718 | in | 
| 58393 | 719 |     (Disc {fun_name = fun_name, fun_T = fun_T, fun_args = fun_args, ctr = ctr, ctr_no = ctr_no,
 | 
| 720 | disc = disc, prems = actual_prems, auto_gen = catch_all, ctr_rhs_opt = ctr_rhs_opt, | |
| 721 | code_rhs_opt = code_rhs_opt, eqn_pos = eqn_pos, user_eqn = user_eqn}, | |
| 722 | matchedsss') | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 723 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 724 | |
| 59594 | 725 | fun dissect_coeqn_sel ctxt fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn_pos | 
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 726 | ctr_rhs_opt code_rhs_opt eqn0 of_spec_opt eqn = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 727 | let | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 728 | val (lhs, rhs) = HOLogic.dest_eq eqn | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 729 | handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eqn]; | 
| 59604 | 730 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 731 | val sel = head_of lhs; | 
| 53720 | 732 | val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free | 
| 59594 | 733 | handle TERM _ => error_at ctxt [eqn] "Ill-formed selector argument in left-hand side"; | 
| 59604 | 734 | val _ = check_fun_args ctxt eqn fun_args; | 
| 56152 | 735 | |
| 54160 | 736 | val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name) | 
| 59594 | 737 | handle Option.Option => error_at ctxt [eqn] "Ill-formed selector argument in left-hand side"; | 
| 54160 | 738 |     val {ctr, ...} =
 | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 739 | (case of_spec_opt of | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 740 | SOME of_spec => the (find_first (curry (op =) of_spec o #ctr) basic_ctr_specs) | 
| 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 741 | | NONE => filter (exists (curry (op =) sel) o #sels) basic_ctr_specs |> the_single | 
| 59594 | 742 | handle List.Empty => error_at ctxt [eqn] "Ambiguous selector (without \"of\")"); | 
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54978diff
changeset | 743 | val user_eqn = drop_all eqn0; | 
| 56152 | 744 | |
| 59599 | 745 | val _ = check_extra_frees ctxt fun_args fun_names user_eqn; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 746 | in | 
| 58393 | 747 |     Sel {fun_name = fun_name, fun_T = fun_T, fun_args = fun_args, ctr = ctr, sel = sel,
 | 
| 748 | rhs_term = rhs, ctr_rhs_opt = ctr_rhs_opt, code_rhs_opt = code_rhs_opt, eqn_pos = eqn_pos, | |
| 749 | user_eqn = user_eqn} | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 750 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 751 | |
| 59594 | 752 | fun dissect_coeqn_ctr ctxt fun_names sequentials (basic_ctr_specss : basic_corec_ctr_spec list list) | 
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 753 | eqn_pos eqn0 code_rhs_opt prems concl matchedsss = | 
| 53910 | 754 | let | 
| 54065 | 755 | val (lhs, rhs) = HOLogic.dest_eq concl; | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 756 | val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free; | 
| 56152 | 757 | |
| 59604 | 758 | val _ = check_fun_args ctxt concl fun_args; | 
| 59599 | 759 | val _ = check_extra_frees ctxt fun_args fun_names (drop_all eqn0); | 
| 56152 | 760 | |
| 58393 | 761 | val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name); | 
| 55343 | 762 | val (ctr, ctr_args) = strip_comb (unfold_lets_splits rhs); | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 763 |     val {disc, sels, ...} = the (find_first (curry (op =) ctr o #ctr) basic_ctr_specs)
 | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 764 | handle Option.Option => not_constructor_in_rhs ctxt [] ctr; | 
| 53341 | 765 | |
| 54065 | 766 | val disc_concl = betapply (disc, lhs); | 
| 54976 
b502f04c0442
repair 'exhaustive' feature for one-constructor types
 blanchet parents: 
54975diff
changeset | 767 | val (eqn_data_disc_opt, matchedsss') = | 
| 59042 | 768 | if null (tl basic_ctr_specs) andalso not (null sels) then | 
| 54976 
b502f04c0442
repair 'exhaustive' feature for one-constructor types
 blanchet parents: 
54975diff
changeset | 769 | (NONE, matchedsss) | 
| 
b502f04c0442
repair 'exhaustive' feature for one-constructor types
 blanchet parents: 
54975diff
changeset | 770 | else | 
| 59594 | 771 | apfst SOME (dissect_coeqn_disc ctxt fun_names sequentials basic_ctr_specss eqn_pos | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 772 | (SOME (abstract_over_list fun_args rhs)) code_rhs_opt prems disc_concl matchedsss); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 773 | |
| 54159 
eb5d58c99049
set stage for more flexible 'primrec' syntax for recursion through functions
 blanchet parents: 
54157diff
changeset | 774 | val sel_concls = sels ~~ ctr_args | 
| 56152 | 775 | |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg)) | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 776 | handle ListPair.UnequalLengths => partially_applied_ctr_in_rhs ctxt [rhs]; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 777 | |
| 54160 | 778 | val eqns_data_sel = | 
| 59594 | 779 | map (dissect_coeqn_sel ctxt fun_names basic_ctr_specss eqn_pos | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 780 | (SOME (abstract_over_list fun_args rhs)) code_rhs_opt eqn0 (SOME ctr)) | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 781 | sel_concls; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 782 | in | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 783 | (the_list eqn_data_disc_opt @ eqns_data_sel, matchedsss') | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 784 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 785 | |
| 59594 | 786 | fun dissect_coeqn_code ctxt has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss = | 
| 54065 | 787 | let | 
| 59594 | 788 | val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs ctxt has_call []); | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 789 | val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free; | 
| 56152 | 790 | |
| 59604 | 791 | val _ = check_fun_args ctxt concl fun_args; | 
| 59599 | 792 | val _ = check_extra_frees ctxt fun_args fun_names concl; | 
| 56152 | 793 | |
| 58393 | 794 | val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name); | 
| 54065 | 795 | |
| 59594 | 796 | val cond_ctrs = fold_rev_corec_code_rhs ctxt (fn cs => fn ctr => fn _ => | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 797 | if member (op = o apsnd #ctr) basic_ctr_specs ctr then cons (ctr, cs) | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 798 | else not_constructor_in_rhs ctxt [] ctr) [] rhs' [] | 
| 54065 | 799 | |> AList.group (op =); | 
| 800 | ||
| 54068 | 801 | val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs); | 
| 54065 | 802 | val ctr_concls = cond_ctrs |> map (fn (ctr, _) => | 
| 58393 | 803 | binder_types (fastype_of ctr) | 
| 59594 | 804 | |> map_index (fn (n, T) => massage_corec_code_rhs ctxt (fn _ => fn ctr' => fn args => | 
| 59612 
7ea413656b64
avoid needless 'if ... undefined' in generated theorems
 blanchet parents: 
59609diff
changeset | 805 | if ctr' = ctr then nth args n else Term.dummy_pattern T) [] rhs') | 
| 58393 | 806 | |> curry Term.list_comb ctr | 
| 807 | |> curry HOLogic.mk_eq lhs); | |
| 54902 
a9291e4d2366
internally allow different values for 'sequential' for different constructors
 blanchet parents: 
54901diff
changeset | 808 | |
| 59594 | 809 | val bads = maps (filter (Term.exists_subterm (is_free_in fun_names))) ctr_premss; | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 810 | val _ = null bads orelse unexpected_corec_call_in ctxt [eqn0] rhs; | 
| 59594 | 811 | |
| 54902 
a9291e4d2366
internally allow different values for 'sequential' for different constructors
 blanchet parents: 
54901diff
changeset | 812 | val sequentials = replicate (length fun_names) false; | 
| 54065 | 813 | in | 
| 59594 | 814 |     @{fold_map 2} (dissect_coeqn_ctr ctxt fun_names sequentials basic_ctr_specss eqn_pos eqn0
 | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 815 | (SOME (abstract_over_list fun_args rhs))) | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 816 | ctr_premss ctr_concls matchedsss | 
| 54065 | 817 | end; | 
| 818 | ||
| 59594 | 819 | fun dissect_coeqn ctxt has_call fun_names sequentials | 
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 820 | (basic_ctr_specss : basic_corec_ctr_spec list list) (eqn_pos, eqn0) of_spec_opt matchedsss = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 821 | let | 
| 54979 
d7593bfccf25
correctly extract code RHS, with loose bound variables
 blanchet parents: 
54978diff
changeset | 822 | val eqn = drop_all eqn0 | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 823 | handle TERM _ => ill_formed_formula ctxt eqn0; | 
| 54065 | 824 | val (prems, concl) = Logic.strip_horn eqn | 
| 57527 | 825 | |> map_prod (map HOLogic.dest_Trueprop) HOLogic.dest_Trueprop | 
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 826 | handle TERM _ => ill_formed_equation ctxt eqn; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 827 | |
| 54065 | 828 | val head = concl | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 829 | |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq)) | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 830 | |> head_of; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 831 | |
| 59601 | 832 | val rhs_opt = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd); | 
| 59605 | 833 | |
| 834 | fun check_num_args () = | |
| 835 | is_none rhs_opt orelse not (can dest_funT (fastype_of (the rhs_opt))) orelse | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 836 | missing_args_to_fun_on_lhs ctxt [eqn]; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 837 | |
| 54160 | 838 | val discs = maps (map #disc) basic_ctr_specss; | 
| 839 | val sels = maps (maps #sels) basic_ctr_specss; | |
| 840 | val ctrs = maps (map #ctr) basic_ctr_specss; | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 841 | in | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 842 | if member (op =) discs head orelse | 
| 58393 | 843 | (is_some rhs_opt andalso | 
| 844 | member (op =) (map SOME fun_names) (try (fst o dest_Free) head) andalso | |
| 845 | member (op =) (filter (null o binder_types o fastype_of) ctrs) (the rhs_opt)) then | |
| 59608 | 846 | (dissect_coeqn_disc ctxt fun_names sequentials basic_ctr_specss eqn_pos NONE NONE prems concl | 
| 59605 | 847 | matchedsss | 
| 848 | |>> single) | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 849 | else if member (op =) sels head then | 
| 59594 | 850 | (null prems orelse error_at ctxt [eqn] "Unexpected condition in selector formula"; | 
| 851 | ([dissect_coeqn_sel ctxt fun_names basic_ctr_specss eqn_pos NONE NONE eqn0 of_spec_opt | |
| 57550 | 852 | concl], matchedsss)) | 
| 59594 | 853 | else if is_some rhs_opt andalso is_Free head andalso is_free_in fun_names head then | 
| 55343 | 854 | if member (op =) ctrs (head_of (unfold_lets_splits (the rhs_opt))) then | 
| 59605 | 855 | (check_num_args (); | 
| 856 | dissect_coeqn_ctr ctxt fun_names sequentials basic_ctr_specss eqn_pos eqn0 | |
| 857 | (if null prems then | |
| 858 | SOME (snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_assums_concl eqn0)))) | |
| 859 | else | |
| 860 | NONE) | |
| 861 | prems concl matchedsss) | |
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 862 | else if null prems then | 
| 59605 | 863 | (check_num_args (); | 
| 864 | dissect_coeqn_code ctxt has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss | |
| 865 | |>> flat) | |
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 866 | else | 
| 59594 | 867 | error_at ctxt [eqn] "Cannot mix constructor and code views" | 
| 59600 | 868 | else if is_some rhs_opt then | 
| 59607 | 869 |       error_at ctxt [eqn] ("Ill-formed equation head: " ^ quote (Syntax.string_of_term ctxt head))
 | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 870 | else | 
| 59607 | 871 | error_at ctxt [eqn] "Expected equation or discriminator formula" | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 872 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 873 | |
| 54002 | 874 | fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list) | 
| 54153 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 875 |     ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
 | 
| 56858 | 876 | if is_none (#pred (nth ctr_specs ctr_no)) then | 
| 877 | I | |
| 878 | else | |
| 54068 | 879 | s_conjs prems | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 880 | |> curry subst_bounds (List.rev fun_args) | 
| 55969 | 881 | |> abs_tuple_balanced fun_args | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 882 | |> K |> nth_map (the (#pred (nth ctr_specs ctr_no))); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 883 | |
| 54153 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 884 | fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 885 | find_first (curry (op =) sel o #sel) sel_eqns | 
| 55969 | 886 |   |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple_balanced fun_args rhs_term)
 | 
| 53720 | 887 | |> the_default undef_const | 
| 53411 | 888 | |> K; | 
| 53360 | 889 | |
| 59594 | 890 | fun build_corec_args_mutual_call ctxt has_call (sel_eqns : coeqn_data_sel list) sel = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 891 | (case find_first (curry (op =) sel o #sel) sel_eqns of | 
| 54208 | 892 | NONE => (I, I, I) | 
| 893 |   | SOME {fun_args, rhs_term, ... } =>
 | |
| 53876 | 894 | let | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 895 | val bound_Ts = List.rev (map fastype_of fun_args); | 
| 59946 | 896 | |
| 54207 
9296ebf40db0
tuned names (to make them independent from temporary naming convention used in characteristic theorems)
 blanchet parents: 
54206diff
changeset | 897 |       fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
 | 
| 
9296ebf40db0
tuned names (to make them independent from temporary naming convention used in characteristic theorems)
 blanchet parents: 
54206diff
changeset | 898 | fun rewrite_end _ t = if has_call t then undef_const else t; | 
| 
9296ebf40db0
tuned names (to make them independent from temporary naming convention used in characteristic theorems)
 blanchet parents: 
54206diff
changeset | 899 | fun rewrite_cont bound_Ts t = | 
| 55969 | 900 | if has_call t then mk_tuple1_balanced bound_Ts (snd (strip_comb t)) else undef_const; | 
| 60001 | 901 | fun massage f _ = massage_let_if_case_corec ctxt has_call f bound_Ts rhs_term | 
| 55969 | 902 | |> abs_tuple_balanced fun_args; | 
| 53876 | 903 | in | 
| 54207 
9296ebf40db0
tuned names (to make them independent from temporary naming convention used in characteristic theorems)
 blanchet parents: 
54206diff
changeset | 904 | (massage rewrite_stop, massage rewrite_end, massage rewrite_cont) | 
| 54208 | 905 | end); | 
| 53360 | 906 | |
| 59594 | 907 | fun build_corec_arg_nested_call ctxt has_call (sel_eqns : coeqn_data_sel list) sel = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 908 | (case find_first (curry (op =) sel o #sel) sel_eqns of | 
| 54208 | 909 | NONE => I | 
| 910 |   | SOME {fun_args, rhs_term, ...} =>
 | |
| 53899 | 911 | let | 
| 59989 | 912 | fun massage_call bound_Ts U T t0 = | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 913 | let | 
| 59946 | 914 | val U2 = | 
| 915 | (case try dest_sumT U of | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 916 | SOME (U1, U2) => if U1 = T then U2 else invalid_map ctxt [] t0 | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 917 | | NONE => invalid_map ctxt [] t0); | 
| 59946 | 918 | |
| 59948 | 919 | fun rewrite bound_Ts (Abs (s, T', t')) = Abs (s, T', rewrite (T' :: bound_Ts) t') | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 920 | | rewrite bound_Ts (t as _ $ _) = | 
| 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 921 | let val (u, vs) = strip_comb t in | 
| 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 922 | if is_Free u andalso has_call u then | 
| 59946 | 923 | Inr_const T U2 $ mk_tuple1_balanced bound_Ts vs | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61348diff
changeset | 924 |                 else if try (fst o dest_Const) u = SOME @{const_name case_prod} then
 | 
| 55343 | 925 | map (rewrite bound_Ts) vs |> chop 1 | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61348diff
changeset | 926 | |>> HOLogic.mk_case_prod o the_single | 
| 55343 | 927 | |> Term.list_comb | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 928 | else | 
| 55343 | 929 | Term.list_comb (rewrite bound_Ts u, map (rewrite bound_Ts) vs) | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 930 | end | 
| 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 931 | | rewrite _ t = | 
| 59946 | 932 | if is_Free t andalso has_call t then Inr_const T U2 $ HOLogic.unit else t; | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 933 | in | 
| 59946 | 934 | rewrite bound_Ts t0 | 
| 55339 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 935 | end; | 
| 
f09037306f25
properly massage 'if's / 'case's etc. under lambdas
 blanchet parents: 
55100diff
changeset | 936 | |
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 937 | fun massage_noncall U T t = | 
| 64627 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 blanchet parents: 
63859diff
changeset | 938 | build_map ctxt [] [] (uncurry Inl_const o dest_sumT o snd) (T, U) $ t; | 
| 59947 | 939 | |
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 940 | val bound_Ts = List.rev (map fastype_of fun_args); | 
| 53899 | 941 | in | 
| 59989 | 942 | fn t => | 
| 943 | rhs_term | |
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 944 | |> massage_nested_corec_call ctxt has_call massage_call (K massage_noncall) bound_Ts | 
| 59989 | 945 | (range_type (fastype_of t)) (fastype_of1 (bound_Ts, rhs_term)) | 
| 946 | |> abs_tuple_balanced fun_args | |
| 54208 | 947 | end); | 
| 53360 | 948 | |
| 59594 | 949 | fun build_corec_args_sel ctxt has_call (all_sel_eqns : coeqn_data_sel list) | 
| 54002 | 950 | (ctr_spec : corec_ctr_spec) = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 951 | (case filter (curry (op =) (#ctr ctr_spec) o #ctr) all_sel_eqns of | 
| 54208 | 952 | [] => I | 
| 953 | | sel_eqns => | |
| 954 | let | |
| 955 | val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec; | |
| 956 | val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list; | |
| 957 | val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list; | |
| 958 | val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list; | |
| 959 | in | |
| 960 | I | |
| 961 | #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls' | |
| 962 | #> fold (fn (sel, (q, g, h)) => | |
| 59594 | 963 | let val (fq, fg, fh) = build_corec_args_mutual_call ctxt has_call sel_eqns sel in | 
| 54208 | 964 | nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls' | 
| 965 | #> fold (fn (sel, n) => nth_map n | |
| 59594 | 966 | (build_corec_arg_nested_call ctxt has_call sel_eqns sel)) nested_calls' | 
| 54208 | 967 | end); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 968 | |
| 59673 | 969 | fun build_defs ctxt bs mxs has_call arg_Tss (corec_specs : corec_spec list) | 
| 54153 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 970 | (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) = | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 971 | let | 
| 54160 | 972 | val corecs = map #corec corec_specs; | 
| 973 | val ctr_specss = map #ctr_specs corec_specs; | |
| 53360 | 974 | val corec_args = hd corecs | 
| 975 | |> fst o split_last o binder_types o fastype_of | |
| 59044 | 976 | |> map (fn T => | 
| 977 |           if range_type T = HOLogic.boolT then Abs (Name.uu_, domain_type T, @{term False})
 | |
| 54806 
a0f024caa04c
pass auto-proved exhaustiveness properties to tactic;
 panny parents: 
54628diff
changeset | 978 |           else Const (@{const_name undefined}, T))
 | 
| 53720 | 979 | |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss | 
| 59594 | 980 | |> fold2 (fold o build_corec_args_sel ctxt has_call) sel_eqnss ctr_specss; | 
| 59041 | 981 | |
| 59599 | 982 | val bad = fold (add_extra_frees ctxt [] []) corec_args []; | 
| 983 | val _ = null bad orelse | |
| 984 | (if exists has_call corec_args then nonprimitive_corec ctxt [] | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 985 | else extra_variable_in_rhs ctxt [] (hd bad)); | 
| 59599 | 986 | |
| 54835 | 987 | val excludess' = | 
| 53720 | 988 | disc_eqnss | 
| 53822 | 989 | |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x)) | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 990 | #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs []) | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 991 | #> maps (uncurry (map o pair) | 
| 53822 | 992 | #> map (fn ((fun_args, c, x, a), (_, c', y, a')) => | 
| 54068 | 993 | ((c, c', a orelse a'), (x, s_not (s_conjs y))) | 
| 57527 | 994 | ||> map_prod (map HOLogic.mk_Trueprop) HOLogic.mk_Trueprop | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 995 | ||> Logic.list_implies | 
| 55342 | 996 | ||> curry Logic.list_all (map dest_Free fun_args)))); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 997 | in | 
| 55343 | 998 | map (Term.list_comb o rpair corec_args) corecs | 
| 59873 | 999 | |> map2 abs_curried_balanced arg_Tss | 
| 59598 | 1000 | |> (fn ts => Syntax.check_terms ctxt ts | 
| 1001 | handle ERROR _ => nonprimitive_corec ctxt []) | |
| 61760 | 1002 |     |> @{map 3} (fn b => fn mx => fn t =>
 | 
| 1003 | ((b, mx), ((Binding.concealed (Thm.def_binding b), []), t))) bs mxs | |
| 54835 | 1004 | |> rpair excludess' | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1005 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1006 | |
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1007 | fun mk_actual_disc_eqns fun_binding arg_Ts exhaustive ({ctr_specs, ...} : corec_spec)
 | 
| 54153 
67487a607ce2
avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
 blanchet parents: 
54145diff
changeset | 1008 | (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) = | 
| 59603 | 1009 | let | 
| 1010 | val fun_name = Binding.name_of fun_binding; | |
| 1011 | val num_disc_eqns = length disc_eqns; | |
| 1012 | val num_ctrs = length ctr_specs; | |
| 1013 | in | |
| 1014 | if (exhaustive andalso num_disc_eqns <> 0) orelse num_disc_eqns <> num_ctrs - 1 then | |
| 1015 |       (num_disc_eqns > 0 orelse error ("Missing discriminator formula for " ^ quote fun_name);
 | |
| 1016 | disc_eqns) | |
| 54910 | 1017 | else | 
| 1018 | let | |
| 59603 | 1019 | val ctr_no = 0 upto length ctr_specs | 
| 55342 | 1020 | |> the o find_first (fn j => not (exists (curry (op =) j o #ctr_no) disc_eqns)); | 
| 59603 | 1021 |         val {ctr, disc, ...} = nth ctr_specs ctr_no;
 | 
| 58393 | 1022 | val sel_eqn_opt = find_first (equal ctr o #ctr) sel_eqns; | 
| 1023 | ||
| 1024 | val fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))); | |
| 54910 | 1025 | val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns) | 
| 1026 | |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options; | |
| 58393 | 1027 | val prems = maps (s_not_conj o #prems) disc_eqns; | 
| 1028 | val ctr_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE; | |
| 1029 | val code_rhs_opt = Option.map #code_rhs_opt sel_eqn_opt |> the_default NONE; | |
| 61760 | 1030 | val eqn_pos = Option.map (curry (op +) 1 o #eqn_pos) sel_eqn_opt | 
| 1031 | |> the_default 100000; (* FIXME *) | |
| 58393 | 1032 | |
| 1033 | val extra_disc_eqn = | |
| 59603 | 1034 |           {fun_name = fun_name, fun_T = fun_T, fun_args = fun_args, ctr = ctr, ctr_no = ctr_no,
 | 
| 58393 | 1035 | disc = disc, prems = prems, auto_gen = true, ctr_rhs_opt = ctr_rhs_opt, | 
| 1036 | code_rhs_opt = code_rhs_opt, eqn_pos = eqn_pos, user_eqn = undef_const}; | |
| 54910 | 1037 | in | 
| 59603 | 1038 | chop ctr_no disc_eqns ||> cons extra_disc_eqn |> op @ | 
| 54910 | 1039 | end | 
| 1040 | end; | |
| 53720 | 1041 | |
| 55100 | 1042 | fun find_corec_calls ctxt has_call (basic_ctr_specs : basic_corec_ctr_spec list) | 
| 1043 |     ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
 | |
| 54160 | 1044 | let | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1045 | val sel_no = find_first (curry (op =) ctr o #ctr) basic_ctr_specs | 
| 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1046 | |> find_index (curry (op =) sel) o #sels o the; | 
| 54160 | 1047 | in | 
| 55341 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 1048 | K (if has_call rhs_term then fold_rev_let_if_case ctxt (K cons) [] rhs_term [] else []) | 
| 
3d2c97392e25
adapted tactic to correctly handle 'if ... then ...' and 'case ...' under lambdas
 blanchet parents: 
55339diff
changeset | 1049 | |> nth_map sel_no |> AList.map_entry (op =) ctr | 
| 54160 | 1050 | end; | 
| 1051 | ||
| 54900 | 1052 | fun applied_fun_of fun_name fun_T fun_args = | 
| 55343 | 1053 | Term.list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0)); | 
| 54900 | 1054 | |
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1055 | fun is_trivial_implies thm = | 
| 54967 
78de75e3e52a
exhaustive rules like '(False ==> P) ==> P ==> P' are now filtered out as trivial
 blanchet parents: 
54959diff
changeset | 1056 | uncurry (member (op aconv)) (Logic.strip_horn (Thm.prop_of thm)); | 
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1057 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1058 | fun primcorec_ursive int auto opts fixes specs of_specs_opt lthy = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1059 | let | 
| 53352 | 1060 | val (bs, mxs) = map_split (apfst fst) fixes; | 
| 55969 | 1061 | val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> mk_tupleT_balanced) fixes |> split_list; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1062 | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 1063 | val _ = check_duplicate_const_names bs; | 
| 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 1064 | val _ = List.app (uncurry (check_top_sort lthy)) (bs ~~ arg_Ts); | 
| 54272 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 1065 | |
| 54902 
a9291e4d2366
internally allow different values for 'sequential' for different constructors
 blanchet parents: 
54901diff
changeset | 1066 | val actual_nn = length bs; | 
| 
a9291e4d2366
internally allow different values for 'sequential' for different constructors
 blanchet parents: 
54901diff
changeset | 1067 | |
| 59281 | 1068 | val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) | 
| 1069 | |> the_default Plugin_Name.default_filter; | |
| 1070 | val sequentials = replicate actual_nn (exists (can (fn Sequential_Option => ())) opts); | |
| 1071 | val exhaustives = replicate actual_nn (exists (can (fn Exhaustive_Option => ())) opts); | |
| 1072 | val transfers = replicate actual_nn (exists (can (fn Transfer_Option => ())) opts); | |
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1073 | |
| 54160 | 1074 | val fun_names = map Binding.name_of bs; | 
| 62497 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1075 | val qualifys = map (fold_rev (uncurry Binding.qualify o swap) o Binding.path_of) bs; | 
| 54160 | 1076 | val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 1077 | val frees = map (fst #>> Binding.name_of #> Free) fixes; | 
| 59594 | 1078 | val has_call = Term.exists_subterm (member (op =) frees); | 
| 54160 | 1079 | val eqns_data = | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1080 |       @{fold_map 2} (dissect_coeqn lthy has_call fun_names sequentials basic_ctr_specss)
 | 
| 55871 | 1081 | (tag_list 0 (map snd specs)) of_specs_opt [] | 
| 54160 | 1082 | |> flat o fst; | 
| 1083 | ||
| 59662 | 1084 | val missing = fun_names | 
| 1085 | |> filter (map (fn Disc x => #fun_name x | Sel x => #fun_name x) eqns_data | |
| 1086 | |> not oo member (op =)); | |
| 64705 
7596b0736ab9
more uniform errors in '(prim)(co)rec(ursive)' variants
 blanchet parents: 
64674diff
changeset | 1087 | val _ = null missing orelse missing_equations_for_const (hd missing); | 
| 56152 | 1088 | |
| 54160 | 1089 | val callssss = | 
| 1090 | map_filter (try (fn Sel x => x)) eqns_data | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1091 | |> partition_eq (op = o apply2 #fun_name) | 
| 54160 | 1092 |       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
 | 
| 54161 | 1093 | |> map (flat o snd) | 
| 54243 | 1094 | |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss | 
| 54160 | 1095 |       |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
 | 
| 1096 | (ctr, map (K []) sels))) basic_ctr_specss); | |
| 1097 | ||
| 59603 | 1098 | val (corec_specs0, _, coinduct_thm, coinduct_strong_thm, coinduct_thms, coinduct_strong_thms, | 
| 61334 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61301diff
changeset | 1099 | (coinduct_attrs, common_coinduct_attrs), n2m, lthy) = | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55571diff
changeset | 1100 | corec_specs_of bs arg_Ts res_Ts frees callssss lthy; | 
| 59603 | 1101 | val corec_specs = take actual_nn corec_specs0; | 
| 54178 | 1102 | val ctr_specss = map #ctr_specs corec_specs; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1103 | |
| 59603 | 1104 | val disc_eqnss0 = map_filter (try (fn Disc x => x)) eqns_data | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1105 | |> partition_eq (op = o apply2 #fun_name) | 
| 53720 | 1106 |       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1107 | |> map (sort (op < o apply2 #ctr_no |> make_ord) o flat o snd); | 
| 56152 | 1108 | |
| 59603 | 1109 | val _ = disc_eqnss0 |> map (fn x => | 
| 59602 | 1110 | let val dups = duplicates (op = o apply2 #ctr_no) x in | 
| 1111 | null dups orelse | |
| 59600 | 1112 | error_at lthy | 
| 59602 | 1113 | (maps (fn t => filter (curry (op =) (#ctr_no t) o #ctr_no) x) dups | 
| 59600 | 1114 |            |> map (fn {ctr_rhs_opt = SOME t, ...} => t | {user_eqn, ...} => user_eqn))
 | 
| 59606 | 1115 | "Overspecified case(s)" | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1116 | end); | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1117 | |
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1118 | val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1119 | |> partition_eq (op = o apply2 #fun_name) | 
| 53720 | 1120 |       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
 | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1121 | |> map (flat o snd); | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1122 | |
| 59602 | 1123 | val _ = sel_eqnss |> map (fn x => | 
| 1124 | let val dups = duplicates (op = o apply2 ctr_sel_of) x in | |
| 1125 | null dups orelse | |
| 1126 | error_at lthy | |
| 1127 | (maps (fn t => filter (curry (op =) (ctr_sel_of t) o ctr_sel_of) x) dups | |
| 1128 |            |> map (fn {ctr_rhs_opt = SOME t, ...} => t | {user_eqn, ...} => user_eqn))
 | |
| 59606 | 1129 | "Overspecified case(s)" | 
| 59602 | 1130 | end); | 
| 1131 | ||
| 53360 | 1132 | val arg_Tss = map (binder_types o snd o fst) fixes; | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1133 |     val disc_eqnss = @{map 6} mk_actual_disc_eqns bs arg_Tss exhaustives corec_specs sel_eqnss
 | 
| 59603 | 1134 | disc_eqnss0; | 
| 54835 | 1135 | val (defs, excludess') = | 
| 61334 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61301diff
changeset | 1136 | build_defs lthy bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss; | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1137 | |
| 55009 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1138 | val tac_opts = | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1139 |       map (fn {code_rhs_opt, ...} :: _ =>
 | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1140 | if auto orelse is_some code_rhs_opt then SOME (auto_tac o #context) else NONE) disc_eqnss; | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1141 | |
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1142 | fun exclude_tac tac_opt sequential (c, c', a) = | 
| 54901 
0b8871677e0b
use same name for feature internally as in user interface, to facilitate grepping
 blanchet parents: 
54900diff
changeset | 1143 | if a orelse c = c' orelse sequential then | 
| 61271 | 1144 |         SOME (fn {context = ctxt, prems = _} => HEADGOAL (mk_primcorec_assumption_tac ctxt []))
 | 
| 54901 
0b8871677e0b
use same name for feature internally as in user interface, to facilitate grepping
 blanchet parents: 
54900diff
changeset | 1145 | else | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1146 | tac_opt; | 
| 53822 | 1147 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1148 |     val excludess'' = @{map 3} (fn tac_opt => fn sequential => map (fn (j, goal) =>
 | 
| 59597 | 1149 | (j, (Option.map (Goal.prove (*no sorry*) lthy [] [] goal #> Thm.close_derivation) | 
| 1150 | (exclude_tac tac_opt sequential j), goal)))) | |
| 1151 | tac_opts sequentials excludess' | |
| 1152 | handle ERROR _ => use_primcorecursive (); | |
| 56152 | 1153 | |
| 54835 | 1154 | val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) excludess''; | 
| 55009 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1155 | val (goal_idxss, exclude_goalss) = excludess'' | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1156 | |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd)) | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1157 | |> split_list o map split_list; | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1158 | |
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1159 | fun list_all_fun_args extras = | 
| 54910 | 1160 | map2 (fn [] => I | 
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1161 |           | {fun_args, ...} :: _ => map (curry Logic.list_all (extras @ map dest_Free fun_args)))
 | 
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1162 | disc_eqnss; | 
| 54844 | 1163 | |
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1164 | val syntactic_exhaustives = | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1165 | map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns | 
| 54913 | 1166 | orelse exists #auto_gen disc_eqns) | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1167 | disc_eqnss; | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1168 | val de_facto_exhaustives = | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1169 | map2 (fn b => fn b' => b orelse b') exhaustives syntactic_exhaustives; | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1170 | |
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1171 | val nchotomy_goalss = | 
| 54904 | 1172 | map2 (fn false => K [] | true => single o HOLogic.mk_Trueprop o mk_dnf o map #prems) | 
| 1173 | de_facto_exhaustives disc_eqnss | |
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1174 | |> list_all_fun_args [] | 
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1175 | val nchotomy_taut_thmss = | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1176 |       @{map 5} (fn tac_opt => fn {exhaust_discs = res_exhaust_discs, ...} =>
 | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1177 |           fn {code_rhs_opt, ...} :: _ => fn [] => K []
 | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1178 | | [goal] => fn true => | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1179 | let | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1180 | val (_, _, arg_exhaust_discs, _, _) = | 
| 55400 | 1181 | case_thms_of_term lthy (the_default Term.dummy code_rhs_opt); | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1182 | in | 
| 59596 
c067eba942e7
no quick_and_dirty for goals that may fail + tuned messages
 blanchet parents: 
59595diff
changeset | 1183 |                 [Goal.prove (*no sorry*) lthy [] [] goal (fn {context = ctxt, ...} =>
 | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1184 | mk_primcorec_nchotomy_tac ctxt (res_exhaust_discs @ arg_exhaust_discs)) | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1185 | |> Thm.close_derivation] | 
| 59597 | 1186 | handle ERROR _ => use_primcorecursive () | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1187 | end | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1188 | | false => | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1189 | (case tac_opt of | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1190 | SOME tac => [Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation] | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1191 | | NONE => [])) | 
| 55400 | 1192 | tac_opts corec_specs disc_eqnss nchotomy_goalss syntactic_exhaustives; | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1193 | |
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1194 | val syntactic_exhaustives = | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1195 | map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1196 | orelse exists #auto_gen disc_eqns) | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1197 | disc_eqnss; | 
| 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1198 | |
| 55009 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1199 | val nchotomy_goalss = | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1200 | map2 (fn (NONE, false) => map (rpair []) | _ => K []) (tac_opts ~~ syntactic_exhaustives) | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1201 | nchotomy_goalss; | 
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1202 | |
| 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1203 | val goalss = nchotomy_goalss @ exclude_goalss; | 
| 54844 | 1204 | |
| 55462 | 1205 | fun prove thmss'' def_infos lthy = | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1206 | let | 
| 55462 | 1207 | val def_thms = map (snd o snd) def_infos; | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1208 | val ts = map fst def_infos; | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1209 | |
| 54972 
5747fdd4ad3b
fix 'primcorec' (as opposed to 'primcorecursive') with 'exhaustive')
 blanchet parents: 
54970diff
changeset | 1210 | val (nchotomy_thmss, exclude_thmss) = | 
| 55009 
d4b69107a86a
automatically solve proof obligations produced for code equations
 blanchet parents: 
55008diff
changeset | 1211 | (map2 append (take actual_nn thmss'') nchotomy_taut_thmss, drop actual_nn thmss''); | 
| 54613 | 1212 | |
| 54927 
a5a2598f0651
proper name generation to avoid clash with 'P' in user specification
 blanchet parents: 
54926diff
changeset | 1213 | val ps = | 
| 
a5a2598f0651
proper name generation to avoid clash with 'P' in user specification
 blanchet parents: 
54926diff
changeset | 1214 |           Variable.variant_frees lthy (maps (maps #fun_args) disc_eqnss) [("P", HOLogic.boolT)];
 | 
| 
a5a2598f0651
proper name generation to avoid clash with 'P' in user specification
 blanchet parents: 
54926diff
changeset | 1215 | |
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1216 | val exhaust_thmss = | 
| 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1217 | map2 (fn false => K [] | 
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1218 |               | true => fn disc_eqns as {fun_args, ...} :: _ =>
 | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1219 | let | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1220 | val p = Bound (length fun_args); | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1221 | fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1222 | in | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1223 | [mk_imp_p (map (mk_imp_p o map HOLogic.mk_Trueprop o #prems) disc_eqns)] | 
| 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1224 | end) | 
| 54904 | 1225 | de_facto_exhaustives disc_eqnss | 
| 54927 
a5a2598f0651
proper name generation to avoid clash with 'P' in user specification
 blanchet parents: 
54926diff
changeset | 1226 | |> list_all_fun_args ps | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1227 |           |> @{map 3} (fn disc_eqns as {fun_args, ...} :: _ => fn [] => K []
 | 
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1228 | | [nchotomy_thm] => fn [goal] => | 
| 61271 | 1229 | [Goal.prove_sorry lthy [] [] goal | 
| 1230 |                   (fn {context = ctxt, prems = _} =>
 | |
| 1231 | mk_primcorec_exhaust_tac ctxt | |
| 1232 |                       ("" (* for "P" *) :: map (fst o dest_Free) fun_args)
 | |
| 1233 | (length disc_eqns) nchotomy_thm) | |
| 54903 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1234 | |> Thm.close_derivation]) | 
| 
c664bd02bf94
internally allow different values for 'exhaustive' for different constructors
 blanchet parents: 
54902diff
changeset | 1235 | disc_eqnss nchotomy_thmss; | 
| 54921 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 blanchet parents: 
54917diff
changeset | 1236 | val nontriv_exhaust_thmss = map (filter_out is_trivial_implies) exhaust_thmss; | 
| 54844 | 1237 | |
| 54835 | 1238 | val excludess' = map (op ~~) (goal_idxss ~~ exclude_thmss); | 
| 1239 | fun mk_excludesss excludes n = | |
| 54973 | 1240 | fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm]))) | 
| 54974 
d1c76303244e
use correct default for exclude rules to avoid weird tactic failures
 blanchet parents: 
54973diff
changeset | 1241 | excludes (map (fn k => replicate k [asm_rl] @ replicate (n - k) []) (0 upto n - 1)); | 
| 54973 | 1242 | val excludessss = | 
| 1243 | map2 (fn excludes => mk_excludesss excludes o length o #ctr_specs) | |
| 1244 | (map2 append excludess' taut_thmss) corec_specs; | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1245 | |
| 54835 | 1246 |         fun prove_disc ({ctr_specs, ...} : corec_spec) excludesss
 | 
| 54948 | 1247 |             ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
 | 
| 54272 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 1248 |           if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
 | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 1249 | [] | 
| 
9d623cada37f
avoid subtle failure in the presence of top sort
 blanchet parents: 
54271diff
changeset | 1250 | else | 
| 53720 | 1251 | let | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1252 |               val {disc, corec_disc, ...} = nth ctr_specs ctr_no;
 | 
| 53720 | 1253 | val k = 1 + ctr_no; | 
| 1254 | val m = length prems; | |
| 54900 | 1255 | val goal = | 
| 1256 | applied_fun_of fun_name fun_T fun_args | |
| 1257 | |> curry betapply disc | |
| 53720 | 1258 | |> HOLogic.mk_Trueprop | 
| 1259 | |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) | |
| 1260 | |> curry Logic.list_all (map dest_Free fun_args); | |
| 1261 | in | |
| 54900 | 1262 |               if prems = [@{term False}] then
 | 
| 1263 | [] | |
| 1264 | else | |
| 61271 | 1265 | Goal.prove_sorry lthy [] [] goal | 
| 1266 |                   (fn {context = ctxt, prems = _} =>
 | |
| 1267 | mk_primcorec_disc_tac ctxt def_thms corec_disc k m excludesss) | |
| 54900 | 1268 | |> Thm.close_derivation | 
| 1269 | |> pair (#disc (nth ctr_specs ctr_no)) | |
| 54948 | 1270 | |> pair eqn_pos | 
| 54900 | 1271 | |> single | 
| 53720 | 1272 | end; | 
| 1273 | ||
| 57399 | 1274 |         fun prove_sel ({sel_defs, fp_nesting_maps, fp_nesting_map_ident0s, fp_nesting_map_comps,
 | 
| 57397 | 1275 | ctr_specs, ...} : corec_spec) (disc_eqns : coeqn_data_disc list) excludesss | 
| 56858 | 1276 |             ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, code_rhs_opt, eqn_pos, ...}
 | 
| 1277 | : coeqn_data_sel) = | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1278 | let | 
| 58393 | 1279 | val ctr_spec = the (find_first (curry (op =) ctr o #ctr) ctr_specs); | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1280 | val ctr_no = find_index (curry (op =) ctr o #ctr) ctr_specs; | 
| 54067 | 1281 | val prems = the_default (maps (s_not_conj o #prems) disc_eqns) | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1282 | (find_first (curry (op =) ctr_no o #ctr_no) disc_eqns |> Option.map #prems); | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1283 | val corec_sel = find_index (curry (op =) sel) (#sels ctr_spec) | 
| 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1284 | |> nth (#corec_sels ctr_spec); | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1285 | val k = 1 + ctr_no; | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1286 | val m = length prems; | 
| 54900 | 1287 | val goal = | 
| 1288 | applied_fun_of fun_name fun_T fun_args | |
| 53720 | 1289 | |> curry betapply sel | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 1290 | |> rpair (abstract_over_list fun_args rhs_term) | 
| 53720 | 1291 | |> HOLogic.mk_Trueprop o HOLogic.mk_eq | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1292 | |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) | 
| 53720 | 1293 | |> curry Logic.list_all (map dest_Free fun_args); | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1294 | val (distincts, _, _, split_sels, split_sel_asms) = case_thms_of_term lthy rhs_term; | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1295 | in | 
| 61271 | 1296 | Goal.prove_sorry lthy [] [] goal | 
| 1297 |               (fn {context = ctxt, prems = _} =>
 | |
| 1298 | mk_primcorec_sel_tac ctxt def_thms distincts split_sels split_sel_asms | |
| 61760 | 1299 | fp_nesting_maps fp_nesting_map_ident0s fp_nesting_map_comps corec_sel k m | 
| 1300 | excludesss) | |
| 54176 | 1301 | |> Thm.close_derivation | 
| 63222 
fe92356ade42
eliminated pointless alias (no warning for duplicates);
 wenzelm parents: 
63182diff
changeset | 1302 | |> `(is_some code_rhs_opt ? Local_Defs.fold lthy sel_defs) (*mildly too aggressive*) | 
| 53720 | 1303 | |> pair sel | 
| 54948 | 1304 | |> pair eqn_pos | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1305 | end; | 
| 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1306 | |
| 56858 | 1307 |         fun prove_ctr disc_alist sel_alist ({sel_defs, ...} : corec_spec)
 | 
| 1308 | (disc_eqns : coeqn_data_disc list) (sel_eqns : coeqn_data_sel list) | |
| 1309 |             ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
 | |
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1310 | (* don't try to prove theorems when some sel_eqns are missing *) | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1311 | if not (exists (curry (op =) ctr o #ctr) disc_eqns) | 
| 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1312 | andalso not (exists (curry (op =) ctr o #ctr) sel_eqns) | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1313 | orelse | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1314 | filter (curry (op =) ctr o #ctr) sel_eqns | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1315 | |> fst o finds (op = o apsnd #sel) sels | 
| 54951 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 1316 | |> exists (null o snd) then | 
| 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 1317 | [] | 
| 
e25b4d22082b
for code equations that coincide with ctr equations, make sure the usr's input is preserved for both
 blanchet parents: 
54948diff
changeset | 1318 | else | 
| 53720 | 1319 | let | 
| 56858 | 1320 | val (fun_name, fun_T, fun_args, prems, ctr_rhs_opt, code_rhs_opt, eqn_pos) = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1321 | (find_first (curry (op =) ctr o #ctr) disc_eqns, | 
| 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1322 | find_first (curry (op =) ctr o #ctr) sel_eqns) | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1323 | |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x, | 
| 56858 | 1324 | #ctr_rhs_opt x, #code_rhs_opt x, #eqn_pos x)) | 
| 1325 | ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], | |
| 1326 | #ctr_rhs_opt x, #code_rhs_opt x, #eqn_pos x)) | |
| 53722 
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
 panny parents: 
53720diff
changeset | 1327 | |> the o merge_options; | 
| 53720 | 1328 | val m = length prems; | 
| 54900 | 1329 | val goal = | 
| 56858 | 1330 | (case ctr_rhs_opt of | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1331 | SOME rhs => rhs | 
| 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1332 | | NONE => | 
| 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1333 | filter (curry (op =) ctr o #ctr) sel_eqns | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1334 | |> fst o finds (op = o apsnd #sel) sels | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 1335 | |> map (snd #> (fn [x] => (#fun_args x, #rhs_term x)) | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 1336 | #-> abstract_over_list) | 
| 55343 | 1337 | |> curry Term.list_comb ctr) | 
| 54900 | 1338 | |> curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args) | 
| 53720 | 1339 | |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) | 
| 1340 | |> curry Logic.list_all (map dest_Free fun_args); | |
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1341 | val disc_thm_opt = AList.lookup (op =) disc_alist disc; | 
| 56858 | 1342 | val sel_thms = map (snd o snd) (filter (member (op =) sels o fst) sel_alist); | 
| 53720 | 1343 | in | 
| 56858 | 1344 |               if prems = [@{term False}] then
 | 
| 1345 | [] | |
| 1346 | else | |
| 61271 | 1347 | Goal.prove_sorry lthy [] [] goal | 
| 1348 |                   (fn {context = ctxt, prems = _} =>
 | |
| 1349 | mk_primcorec_ctr_tac ctxt m collapse disc_thm_opt sel_thms) | |
| 63222 
fe92356ade42
eliminated pointless alias (no warning for duplicates);
 wenzelm parents: 
63182diff
changeset | 1350 | |> is_some code_rhs_opt ? Local_Defs.fold lthy sel_defs (*mildly too aggressive*) | 
| 54176 | 1351 | |> Thm.close_derivation | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1352 | |> pair ctr | 
| 54948 | 1353 | |> pair eqn_pos | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1354 | |> single | 
| 53876 | 1355 | end; | 
| 53720 | 1356 | |
| 55100 | 1357 | fun prove_code exhaustive (disc_eqns : coeqn_data_disc list) | 
| 1358 | (sel_eqns : coeqn_data_sel list) nchotomys ctr_alist ctr_specs = | |
| 54098 | 1359 | let | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1360 | val fun_data_opt = | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1361 | (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns, | 
| 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1362 | find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns) | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1363 | |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x)) | 
| 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1364 | ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x)) | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1365 | |> merge_options; | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1366 | in | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1367 | (case fun_data_opt of | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1368 | NONE => [] | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1369 | | SOME (fun_name, fun_T, fun_args, rhs_opt) => | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1370 | let | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1371 | val bound_Ts = List.rev (map fastype_of fun_args); | 
| 54173 | 1372 | |
| 54900 | 1373 | val lhs = applied_fun_of fun_name fun_T fun_args; | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1374 | val rhs_info_opt = | 
| 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1375 | (case rhs_opt of | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1376 | SOME rhs => | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1377 | let | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1378 | val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs; | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1379 | val cond_ctrs = | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1380 | fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs []; | 
| 54978 
afc156c7e4f7
cope gracefully with missing ctr equations by plugging in 'False ==> ...'
 blanchet parents: 
54976diff
changeset | 1381 | val ctr_thms = | 
| 
afc156c7e4f7
cope gracefully with missing ctr equations by plugging in 'False ==> ...'
 blanchet parents: 
54976diff
changeset | 1382 | map (the_default FalseE o AList.lookup (op =) ctr_alist o snd) cond_ctrs; | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1383 | in SOME (false, rhs, raw_rhs, ctr_thms) end | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1384 | | NONE => | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1385 | let | 
| 55100 | 1386 |                       fun prove_code_ctr ({ctr, sels, ...} : corec_ctr_spec) =
 | 
| 56858 | 1387 | if not (exists (curry (op =) ctr o fst) ctr_alist) then | 
| 1388 | NONE | |
| 1389 | else | |
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1390 | let | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1391 | val prems = find_first (curry (op =) ctr o #ctr) disc_eqns | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1392 | |> Option.map #prems |> the_default []; | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1393 | val t = | 
| 54923 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 blanchet parents: 
54921diff
changeset | 1394 | filter (curry (op =) ctr o #ctr) sel_eqns | 
| 55008 
b5b2e193ca33
use 'disc_exhausts' property both from types on which 'case's take place and on return type
 blanchet parents: 
55005diff
changeset | 1395 | |> fst o finds (op = o apsnd #sel) sels | 
| 60704 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 1396 | |> map (snd #> (fn [x] => (#fun_args x, #rhs_term x)) | 
| 
fdd965b35bcd
tuned ML signature (and rationalized code a bit)
 blanchet parents: 
60683diff
changeset | 1397 | #-> abstract_over_list) | 
| 55343 | 1398 | |> curry Term.list_comb ctr; | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1399 | in | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1400 | SOME (prems, t) | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1401 | end; | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1402 | val ctr_conds_argss_opt = map prove_code_ctr ctr_specs; | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1403 | val exhaustive_code = | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1404 | exhaustive | 
| 54948 | 1405 | orelse exists (is_some andf (null o fst o the)) ctr_conds_argss_opt | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1406 | orelse forall is_some ctr_conds_argss_opt | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1407 | andalso exists #auto_gen disc_eqns; | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1408 | val rhs = | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1409 | (if exhaustive_code then | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1410 | split_last (map_filter I ctr_conds_argss_opt) ||> snd | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1411 | else | 
| 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1412 |                            Const (@{const_name Code.abort}, @{typ String.literal} -->
 | 
| 55966 | 1413 | (HOLogic.unitT --> body_type fun_T) --> body_type fun_T) $ | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1414 | HOLogic.mk_literal fun_name $ | 
| 55966 | 1415 | absdummy HOLogic.unitT (incr_boundvars 1 lhs) | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1416 | |> pair (map_filter I ctr_conds_argss_opt)) | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1417 | |-> fold_rev (fn (prems, u) => mk_If (s_conjs prems) u) | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1418 | in | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1419 | SOME (exhaustive_code, rhs, rhs, map snd ctr_alist) | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1420 | end); | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1421 | in | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1422 | (case rhs_info_opt of | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1423 | NONE => [] | 
| 54905 
2fdec6c29eb7
don't generate any proof obligation for implicit (de facto) exclusiveness
 blanchet parents: 
54904diff
changeset | 1424 | | SOME (exhaustive_code, rhs, raw_rhs, ctr_thms) => | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1425 | let | 
| 59582 | 1426 | val ms = map (Logic.count_prems o Thm.prop_of) ctr_thms; | 
| 54900 | 1427 | val (raw_goal, goal) = (raw_rhs, rhs) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1428 | |> apply2 (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args) | 
| 63727 
2d21591967bc
generate proper goal when equation is entered programmatically
 traytel parents: 
63719diff
changeset | 1429 | #> abstract_over_list fun_args | 
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1430 | #> curry Logic.list_all (map dest_Free fun_args)); | 
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1431 | val (distincts, discIs, _, split_sels, split_sel_asms) = | 
| 55400 | 1432 | case_thms_of_term lthy raw_rhs; | 
| 54098 | 1433 | |
| 62698 | 1434 | val raw_code_thm = | 
| 61271 | 1435 | Goal.prove_sorry lthy [] [] raw_goal | 
| 1436 |                         (fn {context = ctxt, prems = _} =>
 | |
| 1437 | mk_primcorec_raw_code_tac ctxt distincts discIs split_sels split_sel_asms | |
| 61760 | 1438 | ms ctr_thms | 
| 1439 | (if exhaustive_code then try the_single nchotomys else NONE)) | |
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1440 | |> Thm.close_derivation; | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1441 | in | 
| 61271 | 1442 | Goal.prove_sorry lthy [] [] goal | 
| 1443 |                       (fn {context = ctxt, prems = _} =>
 | |
| 1444 | mk_primcorec_code_tac ctxt distincts split_sels raw_code_thm) | |
| 54591 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1445 | |> Thm.close_derivation | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1446 | |> single | 
| 
c822230fd22b
prevent exception when equations for a function are missing;
 panny parents: 
54279diff
changeset | 1447 | end) | 
| 54173 | 1448 | end) | 
| 1449 | end; | |
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1450 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1451 |         val disc_alistss = @{map 3} (map oo prove_disc) corec_specs excludessss disc_eqnss;
 | 
| 54948 | 1452 | val disc_alists = map (map snd o flat) disc_alistss; | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1453 |         val sel_alists = @{map 4} (map ooo prove_sel) corec_specs disc_eqnss excludessss sel_eqnss;
 | 
| 55870 | 1454 | val disc_thmss = map (map snd o sort_list_duplicates o flat) disc_alistss; | 
| 54948 | 1455 | val disc_thmsss' = map (map (map (snd o snd))) disc_alistss; | 
| 56858 | 1456 | val sel_thmss = map (map (fst o snd) o sort_list_duplicates) sel_alists; | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1457 | |
| 54959 
30ded82ff806
fixed 'disc_iff' tactic in the case where different equations use different variable names for the function arguments
 blanchet parents: 
54958diff
changeset | 1458 |         fun prove_disc_iff ({ctr_specs, ...} : corec_spec) exhaust_thms disc_thmss'
 | 
| 
30ded82ff806
fixed 'disc_iff' tactic in the case where different equations use different variable names for the function arguments
 blanchet parents: 
54958diff
changeset | 1459 |             (({fun_args = exhaust_fun_args, ...} : coeqn_data_disc) :: _) disc_thms
 | 
| 54948 | 1460 |             ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
 | 
| 55870 | 1461 | if null exhaust_thms orelse null disc_thms then | 
| 54900 | 1462 | [] | 
| 1463 | else | |
| 1464 | let | |
| 57983 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 blanchet parents: 
57895diff
changeset | 1465 |               val {disc, distinct_discss, ...} = nth ctr_specs ctr_no;
 | 
| 54900 | 1466 | val goal = | 
| 1467 | mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args |> curry betapply disc, | |
| 1468 | mk_conjs prems) | |
| 1469 | |> curry Logic.list_all (map dest_Free fun_args); | |
| 1470 | in | |
| 61271 | 1471 | Goal.prove_sorry lthy [] [] goal | 
| 1472 |                 (fn {context = ctxt, prems = _} =>
 | |
| 1473 | mk_primcorec_disc_iff_tac ctxt (map (fst o dest_Free) exhaust_fun_args) | |
| 1474 | (the_single exhaust_thms) disc_thms disc_thmss' (flat distinct_discss)) | |
| 54969 | 1475 | |> Thm.close_derivation | 
| 60362 | 1476 | |> fold (fn rule => perhaps (try (fn thm => Meson.first_order_resolve lthy thm rule))) | 
| 54969 | 1477 |                 @{thms eqTrueE eq_False[THEN iffD1] notnotD}
 | 
| 1478 | |> pair eqn_pos | |
| 1479 | |> single | |
| 54900 | 1480 | end; | 
| 1481 | ||
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1482 |         val disc_iff_thmss = @{map 6} (flat ooo map2 oooo prove_disc_iff) corec_specs exhaust_thmss
 | 
| 54959 
30ded82ff806
fixed 'disc_iff' tactic in the case where different equations use different variable names for the function arguments
 blanchet parents: 
54958diff
changeset | 1483 | disc_thmsss' disc_eqnss disc_thmsss' disc_eqnss | 
| 55870 | 1484 | |> map sort_list_duplicates; | 
| 54900 | 1485 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1486 |         val ctr_alists = @{map 6} (fn disc_alist => maps oooo prove_ctr disc_alist) disc_alists
 | 
| 56858 | 1487 | (map (map snd) sel_alists) corec_specs disc_eqnss sel_eqnss ctr_specss; | 
| 59043 | 1488 | val ctr_thmss0 = map (map snd) ctr_alists; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59044diff
changeset | 1489 | val ctr_thmss = map (map (snd o snd) o sort (int_ord o apply2 fst)) ctr_alists; | 
| 54097 
92c5bd3b342d
prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
 panny parents: 
54074diff
changeset | 1490 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58462diff
changeset | 1491 | val code_thmss = | 
| 59043 | 1492 |           @{map 6} prove_code exhaustives disc_eqnss sel_eqnss nchotomy_thmss ctr_thmss0 ctr_specss;
 | 
| 53791 | 1493 | |
| 54968 
baa2baf29eff
use 'disc_iff' as simp rules whenever possible + clean up '= True', '= False', etc.
 blanchet parents: 
54967diff
changeset | 1494 | val disc_iff_or_disc_thmss = | 
| 
baa2baf29eff
use 'disc_iff' as simp rules whenever possible + clean up '= True', '= False', etc.
 blanchet parents: 
54967diff
changeset | 1495 | map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss; | 
| 
baa2baf29eff
use 'disc_iff' as simp rules whenever possible + clean up '= True', '= False', etc.
 blanchet parents: 
54967diff
changeset | 1496 | val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss; | 
| 53795 | 1497 | |
| 53797 | 1498 | val common_name = mk_common_name fun_names; | 
| 62497 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1499 | val common_qualify = fold_rev I qualifys; | 
| 53797 | 1500 | |
| 63239 
d562c9948dee
explicit tagging of code equations de-baroquifies interface
 haftmann parents: 
63222diff
changeset | 1501 | val code_attrs = if plugins code_plugin then [Code.add_default_eqn_attrib Code.Equation] else []; | 
| 59283 | 1502 | |
| 55860 | 1503 | val anonymous_notes = | 
| 1504 | [(flat disc_iff_or_disc_thmss, simp_attrs)] | |
| 1505 | |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); | |
| 1506 | ||
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1507 | val common_notes = | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1508 | [(coinductN, if n2m then [coinduct_thm] else [], common_coinduct_attrs), | 
| 58286 | 1509 | (coinduct_strongN, if n2m then [coinduct_strong_thm] else [], common_coinduct_attrs)] | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1510 | |> filter_out (null o #2) | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1511 | |> map (fn (thmN, thms, attrs) => | 
| 62497 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1512 | ((common_qualify (Binding.qualify true common_name (Binding.name thmN)), attrs), | 
| 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1513 | [(thms, [])])); | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1514 | |
| 53791 | 1515 | val notes = | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1516 | [(coinductN, map (if n2m then single else K []) coinduct_thms, coinduct_attrs), | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1517 | (coinduct_strongN, map (if n2m then single else K []) coinduct_strong_thms, | 
| 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1518 | coinduct_attrs), | 
| 59283 | 1519 | (codeN, code_thmss, code_attrs @ nitpicksimp_attrs), | 
| 53797 | 1520 | (ctrN, ctr_thmss, []), | 
| 55860 | 1521 | (discN, disc_thmss, []), | 
| 54900 | 1522 | (disc_iffN, disc_iff_thmss, []), | 
| 54835 | 1523 | (excludeN, exclude_thmss, []), | 
| 54909 | 1524 | (exhaustN, nontriv_exhaust_thmss, []), | 
| 53795 | 1525 | (selN, sel_thmss, simp_attrs), | 
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1526 | (simpsN, simp_thmss, [])] | 
| 53791 | 1527 | |> maps (fn (thmN, thmss, attrs) => | 
| 62497 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1528 |             @{map 3} (fn fun_name => fn qualify => fn thms =>
 | 
| 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1529 | ((qualify (Binding.qualify true fun_name (Binding.name thmN)), attrs), | 
| 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1530 | [(thms, [])])) | 
| 
5b5b704f4811
respect qualification when noting theorems in prim(co)rec
 traytel parents: 
62318diff
changeset | 1531 | fun_names qualifys (take actual_nn thmss)) | 
| 53791 | 1532 | |> filter_out (null o fst o hd o snd); | 
| 60718 | 1533 | |
| 1534 | val fun_ts0 = map fst def_infos; | |
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1535 | in | 
| 55462 | 1536 | lthy | 
| 60718 | 1537 | |> Spec_Rules.add Spec_Rules.Equational (fun_ts0, flat sel_thmss) | 
| 1538 | |> Spec_Rules.add Spec_Rules.Equational (fun_ts0, flat ctr_thmss) | |
| 1539 | |> Spec_Rules.add Spec_Rules.Equational (fun_ts0, flat code_thmss) | |
| 58283 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 blanchet parents: 
58223diff
changeset | 1540 | |> Local_Theory.notes (anonymous_notes @ common_notes @ notes) | 
| 55462 | 1541 | |> snd | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1542 | |> (fn lthy => | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1543 | let | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1544 | val phi = Local_Theory.target_morphism lthy; | 
| 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1545 | val Ts = take actual_nn (map #T corec_specs); | 
| 59281 | 1546 | val fp_rec_sugar = | 
| 1547 |               {transfers = transfers, fun_names = fun_names, funs = map (Morphism.term phi) ts,
 | |
| 1548 | fun_defs = Morphism.fact phi def_thms, fpTs = Ts}; | |
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1549 | in | 
| 59281 | 1550 | interpret_gfp_rec_sugar plugins fp_rec_sugar lthy | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1551 | end) | 
| 53654 
8b9ea4420f81
prove simp theorems for newly generated definitions
 panny parents: 
53411diff
changeset | 1552 | end; | 
| 59281 | 1553 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1554 | fun after_qed thmss' = | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1555 | fold_map Local_Theory.define defs | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1556 | #> tap (uncurry (print_def_consts int)) | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1557 | #-> prove thmss'; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1558 | in | 
| 61334 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61301diff
changeset | 1559 | (goalss, after_qed, lthy) | 
| 53822 | 1560 | end; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1561 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1562 | fun primcorec_ursive_cmd int auto opts (raw_fixes, raw_specs_of) lthy = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1563 | let | 
| 54926 
28b621fce2f9
more SML-ish (less Haskell-ish) naming convention
 blanchet parents: 
54925diff
changeset | 1564 | val (raw_specs, of_specs_opt) = | 
| 56945 | 1565 | split_list raw_specs_of ||> map (Option.map (Syntax.read_term lthy)); | 
| 63064 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 wenzelm parents: 
62698diff
changeset | 1566 | val (fixes, specs) = | 
| 63182 | 1567 | fst (Specification.read_multi_specs raw_fixes (map (fn spec => (spec, [], [])) raw_specs) lthy); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1568 | in | 
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1569 | primcorec_ursive int auto opts fixes specs of_specs_opt lthy | 
| 59594 | 1570 | end; | 
| 53822 | 1571 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1572 | fun primcorecursive_cmd int = (fn (goalss, after_qed, lthy) => | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1573 | lthy | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1574 | |> Proof.theorem NONE after_qed goalss | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1575 | |> Proof.refine_singleton (Method.primitive_text (K I))) ooo | 
| 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1576 | primcorec_ursive_cmd int false; | 
| 54177 | 1577 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1578 | fun primcorec_cmd int = (fn (goalss, after_qed, lthy) => | 
| 59597 | 1579 | lthy |> after_qed (map (fn [] => [] | _ => use_primcorecursive ()) goalss)) ooo | 
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1580 | primcorec_ursive_cmd int true; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1581 | |
| 59281 | 1582 | val corec_option_parser = Parse.group (K "option") | 
| 1583 | (Plugin_Name.parse_filter >> Plugins_Option | |
| 1584 | || Parse.reserved "sequential" >> K Sequential_Option | |
| 59279 | 1585 | || Parse.reserved "exhaustive" >> K Exhaustive_Option | 
| 1586 | || Parse.reserved "transfer" >> K Transfer_Option); | |
| 55529 | 1587 | |
| 59794 | 1588 | val where_alt_props_of_parser = Parse.where_ |-- Parse.!!! (Parse.enum1 "|" | 
| 63352 | 1589 | ((Parse.prop >> pair Binding.empty_atts) -- Scan.option (Parse.reserved "of" |-- Parse.const))); | 
| 55529 | 1590 | |
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59873diff
changeset | 1591 | val _ = Outer_Syntax.local_theory_to_proof @{command_keyword primcorecursive}
 | 
| 55529 | 1592 | "define primitive corecursive functions" | 
| 1593 |   ((Scan.optional (@{keyword "("} |--
 | |
| 59281 | 1594 |       Parse.!!! (Parse.list1 corec_option_parser) --| @{keyword ")"}) []) --
 | 
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1595 | (Parse.vars -- where_alt_props_of_parser) >> uncurry (primcorecursive_cmd true)); | 
| 55529 | 1596 | |
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59873diff
changeset | 1597 | val _ = Outer_Syntax.local_theory @{command_keyword primcorec}
 | 
| 55529 | 1598 | "define primitive corecursive functions" | 
| 59281 | 1599 |   ((Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 corec_option_parser)
 | 
| 1600 |       --| @{keyword ")"}) []) --
 | |
| 64674 
ef0a5fd30f3b
print constants in 'primrec', 'primcorec(ursive)', 'corec(ursive)', like in 'definition' and 'fun(ction)'
 blanchet parents: 
64627diff
changeset | 1601 | (Parse.vars -- where_alt_props_of_parser) >> uncurry (primcorec_cmd true)); | 
| 55529 | 1602 | |
| 61348 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 blanchet parents: 
61334diff
changeset | 1603 | val _ = Theory.setup (gfp_rec_sugar_interpretation transfer_plugin | 
| 59281 | 1604 | gfp_rec_sugar_transfer_interpretation); | 
| 59275 
77cd4992edcd
Add plugin to generate transfer theorem for primrec and primcorec
 desharna parents: 
59058diff
changeset | 1605 | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 1606 | end; |