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