| author | haftmann | 
| Sun, 02 Jul 2017 20:13:38 +0200 | |
| changeset 66251 | cd935b7cb3fb | 
| parent 64638 | 235df39ade87 | 
| child 67700 | 0455834f7817 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_fp_def_sugar.ML  | 
| 49112 | 2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 57668 | 3  | 
Author: Martin Desharnais, TU Muenchen  | 
4  | 
Copyright 2012, 2013, 2014  | 
|
| 49112 | 5  | 
|
| 49389 | 6  | 
Sugared datatype and codatatype constructions.  | 
| 49112 | 7  | 
*)  | 
8  | 
||
| 49636 | 9  | 
signature BNF_FP_DEF_SUGAR =  | 
| 49112 | 10  | 
sig  | 
| 58460 | 11  | 
type fp_ctr_sugar =  | 
12  | 
    {ctrXs_Tss: typ list list,
 | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
13  | 
ctor_iff_dtor: thm,  | 
| 58460 | 14  | 
ctr_defs: thm list,  | 
| 58569 | 15  | 
ctr_sugar: Ctr_Sugar.ctr_sugar,  | 
| 58570 | 16  | 
ctr_transfers: thm list,  | 
| 58571 | 17  | 
case_transfers: thm list,  | 
| 58676 | 18  | 
disc_transfers: thm list,  | 
19  | 
sel_transfers: thm list}  | 
|
| 58458 | 20  | 
|
21  | 
type fp_bnf_sugar =  | 
|
| 58462 | 22  | 
    {map_thms: thm list,
 | 
| 58560 | 23  | 
map_disc_iffs: thm list,  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
24  | 
map_selss: thm list list,  | 
| 58462 | 25  | 
rel_injects: thm list,  | 
| 58562 | 26  | 
rel_distincts: thm list,  | 
| 58563 | 27  | 
rel_sels: thm list,  | 
| 58564 | 28  | 
rel_intros: thm list,  | 
| 58565 | 29  | 
rel_cases: thm list,  | 
| 62335 | 30  | 
pred_injects: thm list,  | 
| 58566 | 31  | 
set_thms: thm list,  | 
| 
58671
 
4cc24f1e52d5
preserve the structure of 'set_sel' theorem in ML
 
desharna 
parents: 
58659 
diff
changeset
 | 
32  | 
set_selssss: thm list list list list,  | 
| 
58673
 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 
desharna 
parents: 
58672 
diff
changeset
 | 
33  | 
set_introssss: thm list list list list,  | 
| 58568 | 34  | 
set_cases: thm list}  | 
| 58458 | 35  | 
|
| 58459 | 36  | 
type fp_co_induct_sugar =  | 
| 58461 | 37  | 
    {co_rec: term,
 | 
38  | 
common_co_inducts: thm list,  | 
|
39  | 
co_inducts: thm list,  | 
|
40  | 
co_rec_def: thm,  | 
|
| 58459 | 41  | 
co_rec_thms: thm list,  | 
42  | 
co_rec_discs: thm list,  | 
|
| 58572 | 43  | 
co_rec_disc_iffs: thm list,  | 
| 58573 | 44  | 
co_rec_selss: thm list list,  | 
45  | 
co_rec_codes: thm list,  | 
|
| 58574 | 46  | 
co_rec_transfers: thm list,  | 
| 58734 | 47  | 
co_rec_o_maps: thm list,  | 
| 58575 | 48  | 
common_rel_co_inducts: thm list,  | 
| 58576 | 49  | 
rel_co_inducts: thm list,  | 
| 58577 | 50  | 
common_set_inducts: thm list,  | 
51  | 
set_inducts: thm list}  | 
|
| 58457 | 52  | 
|
| 58180 | 53  | 
type fp_sugar =  | 
54  | 
    {T: typ,
 | 
|
55  | 
BT: typ,  | 
|
56  | 
X: typ,  | 
|
57  | 
fp: BNF_Util.fp_kind,  | 
|
58  | 
fp_res_index: int,  | 
|
59  | 
fp_res: BNF_FP_Util.fp_result,  | 
|
60  | 
pre_bnf: BNF_Def.bnf,  | 
|
| 58674 | 61  | 
fp_bnf: BNF_Def.bnf,  | 
| 58180 | 62  | 
absT_info: BNF_Comp.absT_info,  | 
63  | 
fp_nesting_bnfs: BNF_Def.bnf list,  | 
|
64  | 
live_nesting_bnfs: BNF_Def.bnf list,  | 
|
| 58457 | 65  | 
fp_ctr_sugar: fp_ctr_sugar,  | 
66  | 
fp_bnf_sugar: fp_bnf_sugar,  | 
|
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
67  | 
fp_co_induct_sugar: fp_co_induct_sugar option}  | 
| 58180 | 68  | 
|
| 59276 | 69  | 
val co_induct_of: 'a list -> 'a  | 
70  | 
val strong_co_induct_of: 'a list -> 'a  | 
|
71  | 
||
| 63803 | 72  | 
val morph_fp_bnf_sugar: morphism -> fp_bnf_sugar -> fp_bnf_sugar  | 
73  | 
val morph_fp_co_induct_sugar: morphism -> fp_co_induct_sugar -> fp_co_induct_sugar  | 
|
74  | 
val morph_fp_ctr_sugar: morphism -> fp_ctr_sugar -> fp_ctr_sugar  | 
|
| 58180 | 75  | 
val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar  | 
76  | 
val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar  | 
|
77  | 
val fp_sugar_of: Proof.context -> string -> fp_sugar option  | 
|
78  | 
val fp_sugar_of_global: theory -> string -> fp_sugar option  | 
|
79  | 
val fp_sugars_of: Proof.context -> fp_sugar list  | 
|
80  | 
val fp_sugars_of_global: theory -> fp_sugar list  | 
|
| 61760 | 81  | 
val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) ->  | 
82  | 
theory -> theory  | 
|
| 58188 | 83  | 
val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory  | 
| 58182 | 84  | 
val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory  | 
| 58188 | 85  | 
val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory  | 
| 51852 | 86  | 
|
| 63787 | 87  | 
val merge_type_args: BNF_Util.fp_kind -> ''a list * ''a list -> ''a list  | 
88  | 
  val type_args_named_constrained_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'a
 | 
|
89  | 
  val type_binding_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'b
 | 
|
90  | 
  val mixfix_of_spec: ((('a * 'b) * 'c) * 'd) * 'e -> 'b
 | 
|
91  | 
  val mixfixed_ctr_specs_of_spec: (('a * 'b) * 'c) * 'd -> 'b
 | 
|
92  | 
  val map_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'b
 | 
|
93  | 
  val rel_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'c
 | 
|
94  | 
  val pred_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'd
 | 
|
95  | 
val sel_default_eqs_of_spec: 'a * 'b -> 'b  | 
|
96  | 
||
| 59276 | 97  | 
val mk_parametricity_goal: Proof.context -> term list -> term -> term -> term  | 
| 53106 | 98  | 
|
| 52868 | 99  | 
val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list ->  | 
100  | 
'a list  | 
|
| 59818 | 101  | 
val mk_ctor: typ list -> term -> term  | 
102  | 
val mk_dtor: typ list -> term -> term  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
103  | 
val mk_bnf_sets: BNF_Def.bnf -> string * term list  | 
| 58230 | 104  | 
val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list  | 
| 57397 | 105  | 
val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list  | 
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
106  | 
|
| 63825 | 107  | 
val massage_simple_notes: string -> (bstring * 'a list * (int -> 'b)) list ->  | 
108  | 
    ((binding * 'c list) * ('a list * 'b) list) list
 | 
|
109  | 
val massage_multi_notes: string list -> typ list ->  | 
|
110  | 
(string * 'a list list * (string -> 'b)) list ->  | 
|
111  | 
    ((binding * 'b) * ('a list * 'c list) list) list
 | 
|
112  | 
||
| 63817 | 113  | 
val define_ctrs_dtrs_for_type: string -> typ -> term -> term -> thm -> thm -> int -> int list ->  | 
114  | 
term -> binding list -> mixfix list -> typ list list -> local_theory ->  | 
|
115  | 
(term list list * term list * thm * thm list) * local_theory  | 
|
| 63839 | 116  | 
val wrap_ctrs: (string -> bool) -> BNF_Util.fp_kind -> bool -> string -> thm -> int -> int list ->  | 
| 63817 | 117  | 
thm -> thm -> binding list -> binding list list -> term list -> term list -> thm -> thm list ->  | 
118  | 
local_theory -> Ctr_Sugar.ctr_sugar * local_theory  | 
|
119  | 
val derive_map_set_rel_pred_thms: (string -> bool) -> BNF_Util.fp_kind -> int -> typ list ->  | 
|
120  | 
typ list -> typ -> typ -> thm list -> thm list -> thm list -> thm list -> thm list ->  | 
|
| 64416 | 121  | 
thm list -> thm list -> thm list -> thm list -> string -> BNF_Def.bnf -> BNF_Def.bnf list ->  | 
122  | 
typ -> term -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm -> thm list ->  | 
|
123  | 
thm list -> thm list -> typ list list -> Ctr_Sugar.ctr_sugar -> local_theory ->  | 
|
| 63817 | 124  | 
(thm list * thm list * thm list list * thm list * thm list * thm list * thm list * thm list  | 
125  | 
* thm list * thm list * thm list list list list * thm list list list list * thm list  | 
|
126  | 
* thm list * thm list * thm list * thm list) * local_theory  | 
|
127  | 
||
| 58214 | 128  | 
type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list)  | 
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
129  | 
|
| 54256 | 130  | 
val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms  | 
| 58115 | 131  | 
val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms  | 
| 54256 | 132  | 
|
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
133  | 
type gfp_sugar_thms =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
134  | 
((thm list * thm) list * (Token.src list * Token.src list))  | 
| 57489 | 135  | 
* thm list list  | 
136  | 
* thm list list  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
137  | 
* (thm list list * Token.src list)  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
138  | 
* (thm list list list * Token.src list)  | 
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
139  | 
|
| 54256 | 140  | 
val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms  | 
| 58115 | 141  | 
val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms  | 
| 54256 | 142  | 
|
| 61551 | 143  | 
val mk_co_recs_prelims: Proof.context -> BNF_Util.fp_kind -> typ list list list -> typ list ->  | 
144  | 
typ list -> typ list -> typ list -> int list -> int list list -> term list ->  | 
|
145  | 
term list  | 
|
146  | 
* (typ list list * typ list list list list * term list list * term list list list list) option  | 
|
147  | 
* (string * term list * term list list  | 
|
148  | 
* (((term list list * term list list * term list list list list * term list list list list)  | 
|
149  | 
* term list list list) * typ list)) option  | 
|
| 
55772
 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 
blanchet 
parents: 
55701 
diff
changeset
 | 
150  | 
val repair_nullary_single_ctr: typ list list -> typ list list  | 
| 55867 | 151  | 
val mk_corec_p_pred_types: typ list -> int list -> typ list list  | 
152  | 
val mk_corec_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list ->  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
153  | 
int list list -> term ->  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
154  | 
typ list list  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
155  | 
* (typ list list list list * typ list list list * typ list list list list * typ list)  | 
| 58210 | 156  | 
val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory ->  | 
| 58209 | 157  | 
(term * thm) * local_theory  | 
| 55867 | 158  | 
val define_rec:  | 
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
159  | 
typ list list * typ list list list list * term list list * term list list list list ->  | 
| 
55862
 
b458558cbcc2
optimized simple non-recursive datatypes by reusing 'case' for 'rec' constant
 
blanchet 
parents: 
55861 
diff
changeset
 | 
160  | 
(string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context ->  | 
| 55864 | 161  | 
(term * thm) * Proof.context  | 
| 55867 | 162  | 
val define_corec: 'a * term list * term list list  | 
| 58448 | 163  | 
* (((term list list * term list list * term list list list list * term list list list list)  | 
164  | 
* term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list ->  | 
|
165  | 
term list -> term -> local_theory -> (term * thm) * local_theory  | 
|
| 
64626
 
f6d0578b46c9
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64624 
diff
changeset
 | 
166  | 
val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list ->  | 
| 
 
f6d0578b46c9
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64624 
diff
changeset
 | 
167  | 
(string * term list) list -> term -> term -> typ list -> typ list ->  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
168  | 
term list * ((term * (term * term)) list * (int * term)) list * term  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
169  | 
val finish_induct_prem: Proof.context -> int -> term list ->  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
170  | 
term list * ((term * (term * term)) list * (int * term)) list * term -> term  | 
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
171  | 
val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term ->  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
172  | 
term -> term -> int -> term list -> term list list -> term list -> term list list ->  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
173  | 
typ list list -> term  | 
| 58214 | 174  | 
val mk_induct_attrs: term list list -> Token.src list  | 
| 
58283
 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 
blanchet 
parents: 
58267 
diff
changeset
 | 
175  | 
val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list ->  | 
| 
 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 
blanchet 
parents: 
58267 
diff
changeset
 | 
176  | 
Token.src list * Token.src list  | 
| 
58335
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
177  | 
val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list ->  | 
| 55867 | 178  | 
     ('a * typ list list list list * term list list * 'b) option -> thm -> thm list ->
 | 
| 
55862
 
b458558cbcc2
optimized simple non-recursive datatypes by reusing 'case' for 'rec' constant
 
blanchet 
parents: 
55861 
diff
changeset
 | 
179  | 
BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list ->  | 
| 
 
b458558cbcc2
optimized simple non-recursive datatypes by reusing 'case' for 'rec' constant
 
blanchet 
parents: 
55861 
diff
changeset
 | 
180  | 
typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list ->  | 
| 55864 | 181  | 
term list -> thm list -> Proof.context -> lfp_sugar_thms  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
182  | 
val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list ->  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
183  | 
thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list ->  | 
| 62395 | 184  | 
thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list ->  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
185  | 
(thm list * thm) list  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
186  | 
val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list ->  | 
| 58448 | 187  | 
string * term list * term list list  | 
188  | 
* (((term list list * term list list * term list list list list * term list list list list)  | 
|
189  | 
* term list list list) * typ list) ->  | 
|
| 55867 | 190  | 
thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list ->  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
191  | 
typ list -> typ list list list -> int list list -> int list list -> int list -> thm list ->  | 
| 55864 | 192  | 
thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list ->  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
193  | 
thm list -> gfp_sugar_thms  | 
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
194  | 
|
| 
55575
 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 
blanchet 
parents: 
55573 
diff
changeset
 | 
195  | 
val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list ->  | 
| 62324 | 196  | 
binding list -> binding list list -> binding list -> (string * sort) list ->  | 
197  | 
typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory ->  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
198  | 
BNF_FP_Util.fp_result * local_theory) ->  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
199  | 
Ctr_Sugar.ctr_options  | 
| 
57206
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57205 
diff
changeset
 | 
200  | 
* ((((((binding option * (typ * sort)) list * binding) * mixfix)  | 
| 62324 | 201  | 
* ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) *  | 
202  | 
(binding * binding * binding))  | 
|
| 
57206
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57205 
diff
changeset
 | 
203  | 
* term list) list ->  | 
| 49297 | 204  | 
local_theory -> local_theory  | 
| 61301 | 205  | 
val co_datatype_cmd: BNF_Util.fp_kind ->  | 
| 62324 | 206  | 
(mixfix list -> binding list -> binding list -> binding list -> binding list list ->  | 
207  | 
binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->  | 
|
| 61301 | 208  | 
BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * Proof.context) ->  | 
209  | 
((Proof.context -> Plugin_Name.filter) * bool)  | 
|
210  | 
* ((((((binding option * (string * string option)) list * binding) * mixfix)  | 
|
211  | 
* ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list)  | 
|
| 62324 | 212  | 
* (binding * binding * binding)) * string list) list ->  | 
| 61301 | 213  | 
Proof.context -> local_theory  | 
| 63786 | 214  | 
|
215  | 
val parse_ctr_arg: (binding * string) parser  | 
|
216  | 
val parse_ctr_specs: ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list parser  | 
|
217  | 
val parse_spec: ((((((binding option * (string * string option)) list * binding) * mixfix)  | 
|
218  | 
* ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list)  | 
|
219  | 
* (binding * binding * binding)) * string list) parser  | 
|
220  | 
val parse_co_datatype: (Ctr_Sugar.ctr_options_cmd  | 
|
221  | 
* ((((((binding option * (string * string option)) list * binding) * mixfix)  | 
|
222  | 
* ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list)  | 
|
223  | 
* (binding * binding * binding)) * string list) list) parser  | 
|
224  | 
||
| 
55575
 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 
blanchet 
parents: 
55573 
diff
changeset
 | 
225  | 
val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list ->  | 
| 62324 | 226  | 
binding list -> binding list list -> binding list -> (string * sort) list ->  | 
227  | 
typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory ->  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
228  | 
BNF_FP_Util.fp_result * local_theory) ->  | 
| 
49308
 
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
 
blanchet 
parents: 
49302 
diff
changeset
 | 
229  | 
(local_theory -> local_theory) parser  | 
| 49112 | 230  | 
end;  | 
231  | 
||
| 49636 | 232  | 
structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR =  | 
| 49112 | 233  | 
struct  | 
234  | 
||
| 54006 | 235  | 
open Ctr_Sugar  | 
| 55571 | 236  | 
open BNF_FP_Rec_Sugar_Util  | 
| 49119 | 237  | 
open BNF_Util  | 
| 53222 | 238  | 
open BNF_Comp  | 
| 
49214
 
2a3cb4c71b87
construct the right iterator theorem in the recursive case
 
blanchet 
parents: 
49213 
diff
changeset
 | 
239  | 
open BNF_Def  | 
| 
51850
 
106afdf5806c
renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
 
blanchet 
parents: 
51847 
diff
changeset
 | 
240  | 
open BNF_FP_Util  | 
| 49636 | 241  | 
open BNF_FP_Def_Sugar_Tactics  | 
| 49119 | 242  | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
243  | 
val Eq_prefix = "Eq_";  | 
| 57489 | 244  | 
|
| 58093 | 245  | 
val case_transferN = "case_transfer";  | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
246  | 
val ctor_iff_dtorN = "ctor_iff_dtor";  | 
| 57999 | 247  | 
val ctr_transferN = "ctr_transfer";  | 
| 58095 | 248  | 
val disc_transferN = "disc_transfer";  | 
| 58676 | 249  | 
val sel_transferN = "sel_transfer";  | 
| 57489 | 250  | 
val corec_codeN = "corec_code";  | 
| 58448 | 251  | 
val corec_transferN = "corec_transfer";  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
252  | 
val map_disc_iffN = "map_disc_iff";  | 
| 58734 | 253  | 
val map_o_corecN = "map_o_corec";  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
254  | 
val map_selN = "map_sel";  | 
| 62335 | 255  | 
val pred_injectN = "pred_inject";  | 
| 58732 | 256  | 
val rec_o_mapN = "rec_o_map";  | 
| 58446 | 257  | 
val rec_transferN = "rec_transfer";  | 
| 63851 | 258  | 
val set0N = "set0";  | 
| 57893 | 259  | 
val set_casesN = "set_cases";  | 
| 57891 | 260  | 
val set_introsN = "set_intros";  | 
| 57700 | 261  | 
val set_inductN = "set_induct";  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
262  | 
val set_selN = "set_sel";  | 
| 
51777
 
48a0ae342ea0
generate proper attributes for coinduction rules
 
blanchet 
parents: 
51769 
diff
changeset
 | 
263  | 
|
| 58460 | 264  | 
type fp_ctr_sugar =  | 
265  | 
  {ctrXs_Tss: typ list list,
 | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
266  | 
ctor_iff_dtor: thm,  | 
| 58460 | 267  | 
ctr_defs: thm list,  | 
| 58569 | 268  | 
ctr_sugar: Ctr_Sugar.ctr_sugar,  | 
| 58570 | 269  | 
ctr_transfers: thm list,  | 
| 58571 | 270  | 
case_transfers: thm list,  | 
| 58676 | 271  | 
disc_transfers: thm list,  | 
272  | 
sel_transfers: thm list};  | 
|
| 58458 | 273  | 
|
274  | 
type fp_bnf_sugar =  | 
|
| 58462 | 275  | 
  {map_thms: thm list,
 | 
| 58560 | 276  | 
map_disc_iffs: thm list,  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
277  | 
map_selss: thm list list,  | 
| 58462 | 278  | 
rel_injects: thm list,  | 
| 58562 | 279  | 
rel_distincts: thm list,  | 
| 58563 | 280  | 
rel_sels: thm list,  | 
| 58564 | 281  | 
rel_intros: thm list,  | 
| 58565 | 282  | 
rel_cases: thm list,  | 
| 62335 | 283  | 
pred_injects: thm list,  | 
| 58566 | 284  | 
set_thms: thm list,  | 
| 
58671
 
4cc24f1e52d5
preserve the structure of 'set_sel' theorem in ML
 
desharna 
parents: 
58659 
diff
changeset
 | 
285  | 
set_selssss: thm list list list list,  | 
| 
58673
 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 
desharna 
parents: 
58672 
diff
changeset
 | 
286  | 
set_introssss: thm list list list list,  | 
| 58568 | 287  | 
set_cases: thm list};  | 
| 58458 | 288  | 
|
| 58459 | 289  | 
type fp_co_induct_sugar =  | 
| 58461 | 290  | 
  {co_rec: term,
 | 
291  | 
common_co_inducts: thm list,  | 
|
292  | 
co_inducts: thm list,  | 
|
293  | 
co_rec_def: thm,  | 
|
| 58459 | 294  | 
co_rec_thms: thm list,  | 
295  | 
co_rec_discs: thm list,  | 
|
| 58572 | 296  | 
co_rec_disc_iffs: thm list,  | 
| 58573 | 297  | 
co_rec_selss: thm list list,  | 
298  | 
co_rec_codes: thm list,  | 
|
| 58574 | 299  | 
co_rec_transfers: thm list,  | 
| 58734 | 300  | 
co_rec_o_maps: thm list,  | 
| 58575 | 301  | 
common_rel_co_inducts: thm list,  | 
| 58576 | 302  | 
rel_co_inducts: thm list,  | 
| 58577 | 303  | 
common_set_inducts: thm list,  | 
304  | 
set_inducts: thm list};  | 
|
| 58457 | 305  | 
|
| 58180 | 306  | 
type fp_sugar =  | 
307  | 
  {T: typ,
 | 
|
308  | 
BT: typ,  | 
|
309  | 
X: typ,  | 
|
310  | 
fp: fp_kind,  | 
|
311  | 
fp_res_index: int,  | 
|
312  | 
fp_res: fp_result,  | 
|
313  | 
pre_bnf: bnf,  | 
|
| 58674 | 314  | 
fp_bnf: bnf,  | 
| 58180 | 315  | 
absT_info: absT_info,  | 
316  | 
fp_nesting_bnfs: bnf list,  | 
|
317  | 
live_nesting_bnfs: bnf list,  | 
|
| 58457 | 318  | 
fp_ctr_sugar: fp_ctr_sugar,  | 
319  | 
fp_bnf_sugar: fp_bnf_sugar,  | 
|
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
320  | 
fp_co_induct_sugar: fp_co_induct_sugar option};  | 
| 58180 | 321  | 
|
| 59276 | 322  | 
fun co_induct_of (i :: _) = i;  | 
323  | 
fun strong_co_induct_of [_, s] = s;  | 
|
324  | 
||
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
325  | 
fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_selss, rel_injects, rel_distincts,
 | 
| 62335 | 326  | 
rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss,  | 
| 
58673
 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 
desharna 
parents: 
58672 
diff
changeset
 | 
327  | 
set_cases} : fp_bnf_sugar) =  | 
| 58462 | 328  | 
  {map_thms = map (Morphism.thm phi) map_thms,
 | 
| 58560 | 329  | 
map_disc_iffs = map (Morphism.thm phi) map_disc_iffs,  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
330  | 
map_selss = map (map (Morphism.thm phi)) map_selss,  | 
| 58462 | 331  | 
rel_injects = map (Morphism.thm phi) rel_injects,  | 
| 58562 | 332  | 
rel_distincts = map (Morphism.thm phi) rel_distincts,  | 
| 58563 | 333  | 
rel_sels = map (Morphism.thm phi) rel_sels,  | 
| 58564 | 334  | 
rel_intros = map (Morphism.thm phi) rel_intros,  | 
| 58565 | 335  | 
rel_cases = map (Morphism.thm phi) rel_cases,  | 
| 62335 | 336  | 
pred_injects = map (Morphism.thm phi) pred_injects,  | 
| 58566 | 337  | 
set_thms = map (Morphism.thm phi) set_thms,  | 
| 
58671
 
4cc24f1e52d5
preserve the structure of 'set_sel' theorem in ML
 
desharna 
parents: 
58659 
diff
changeset
 | 
338  | 
set_selssss = map (map (map (map (Morphism.thm phi)))) set_selssss,  | 
| 
58673
 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 
desharna 
parents: 
58672 
diff
changeset
 | 
339  | 
set_introssss = map (map (map (map (Morphism.thm phi)))) set_introssss,  | 
| 58568 | 340  | 
set_cases = map (Morphism.thm phi) set_cases};  | 
| 58458 | 341  | 
|
| 58461 | 342  | 
fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms,
 | 
| 58734 | 343  | 
co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, co_rec_o_maps,  | 
| 58577 | 344  | 
common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts} : fp_co_induct_sugar) =  | 
| 58461 | 345  | 
  {co_rec = Morphism.term phi co_rec,
 | 
346  | 
common_co_inducts = map (Morphism.thm phi) common_co_inducts,  | 
|
347  | 
co_inducts = map (Morphism.thm phi) co_inducts,  | 
|
348  | 
co_rec_def = Morphism.thm phi co_rec_def,  | 
|
| 58459 | 349  | 
co_rec_thms = map (Morphism.thm phi) co_rec_thms,  | 
350  | 
co_rec_discs = map (Morphism.thm phi) co_rec_discs,  | 
|
| 58572 | 351  | 
co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs,  | 
| 58573 | 352  | 
co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss,  | 
353  | 
co_rec_codes = map (Morphism.thm phi) co_rec_codes,  | 
|
| 58574 | 354  | 
co_rec_transfers = map (Morphism.thm phi) co_rec_transfers,  | 
| 58734 | 355  | 
co_rec_o_maps = map (Morphism.thm phi) co_rec_o_maps,  | 
| 58575 | 356  | 
common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts,  | 
| 58576 | 357  | 
rel_co_inducts = map (Morphism.thm phi) rel_co_inducts,  | 
| 58577 | 358  | 
common_set_inducts = map (Morphism.thm phi) common_set_inducts,  | 
359  | 
set_inducts = map (Morphism.thm phi) set_inducts};  | 
|
| 58459 | 360  | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
361  | 
fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctor_iff_dtor, ctr_defs, ctr_sugar, ctr_transfers,
 | 
| 
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
362  | 
case_transfers, disc_transfers, sel_transfers} : fp_ctr_sugar) =  | 
| 58460 | 363  | 
  {ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss,
 | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
364  | 
ctor_iff_dtor = Morphism.thm phi ctor_iff_dtor,  | 
| 58460 | 365  | 
ctr_defs = map (Morphism.thm phi) ctr_defs,  | 
| 58569 | 366  | 
ctr_sugar = morph_ctr_sugar phi ctr_sugar,  | 
| 58570 | 367  | 
ctr_transfers = map (Morphism.thm phi) ctr_transfers,  | 
| 58571 | 368  | 
case_transfers = map (Morphism.thm phi) case_transfers,  | 
| 58676 | 369  | 
disc_transfers = map (Morphism.thm phi) disc_transfers,  | 
370  | 
sel_transfers = map (Morphism.thm phi) sel_transfers};  | 
|
| 58460 | 371  | 
|
| 58674 | 372  | 
fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, fp_bnf, absT_info,
 | 
373  | 
fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar,  | 
|
| 58460 | 374  | 
fp_co_induct_sugar} : fp_sugar) =  | 
| 58180 | 375  | 
  {T = Morphism.typ phi T,
 | 
376  | 
BT = Morphism.typ phi BT,  | 
|
377  | 
X = Morphism.typ phi X,  | 
|
378  | 
fp = fp,  | 
|
379  | 
fp_res = morph_fp_result phi fp_res,  | 
|
380  | 
fp_res_index = fp_res_index,  | 
|
381  | 
pre_bnf = morph_bnf phi pre_bnf,  | 
|
| 58674 | 382  | 
fp_bnf = morph_bnf phi fp_bnf,  | 
| 58180 | 383  | 
absT_info = morph_absT_info phi absT_info,  | 
384  | 
fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs,  | 
|
385  | 
live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs,  | 
|
| 58460 | 386  | 
fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar,  | 
| 58458 | 387  | 
fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar,  | 
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
388  | 
fp_co_induct_sugar = Option.map (morph_fp_co_induct_sugar phi) fp_co_induct_sugar};  | 
| 58180 | 389  | 
|
390  | 
val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism;  | 
|
391  | 
||
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
392  | 
structure Data = Generic_Data  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
393  | 
(  | 
| 51840 | 394  | 
type T = fp_sugar Symtab.table;  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
395  | 
val empty = Symtab.empty;  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
396  | 
val extend = I;  | 
| 
55394
 
d5ebe055b599
more liberal merging of BNFs and constructor sugar
 
blanchet 
parents: 
55356 
diff
changeset
 | 
397  | 
fun merge data : T = Symtab.merge (K true) data;  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
398  | 
);  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
399  | 
|
| 58116 | 400  | 
fun fp_sugar_of_generic context =  | 
| 61760 | 401  | 
Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context);  | 
| 58116 | 402  | 
|
403  | 
fun fp_sugars_of_generic context =  | 
|
404  | 
Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) [];  | 
|
| 53907 | 405  | 
|
| 58116 | 406  | 
val fp_sugar_of = fp_sugar_of_generic o Context.Proof;  | 
407  | 
val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory;  | 
|
408  | 
||
409  | 
val fp_sugars_of = fp_sugars_of_generic o Context.Proof;  | 
|
410  | 
val fp_sugars_of_global = fp_sugars_of_generic o Context.Theory;  | 
|
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
411  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58634 
diff
changeset
 | 
412  | 
structure FP_Sugar_Plugin = Plugin(type T = fp_sugar list);  | 
| 
56347
 
6edbd4d20717
added new-style (co)datatype interpretation hook
 
blanchet 
parents: 
56254 
diff
changeset
 | 
413  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58634 
diff
changeset
 | 
414  | 
fun fp_sugars_interpretation name f =  | 
| 62322 | 415  | 
FP_Sugar_Plugin.interpretation name (fn fp_sugars => fn lthy =>  | 
416  | 
f (map (transfer_fp_sugar (Proof_Context.theory_of lthy)) fp_sugars) lthy);  | 
|
| 
56376
 
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
 
blanchet 
parents: 
56347 
diff
changeset
 | 
417  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58634 
diff
changeset
 | 
418  | 
val interpret_fp_sugars = FP_Sugar_Plugin.data;  | 
| 
56347
 
6edbd4d20717
added new-style (co)datatype interpretation hook
 
blanchet 
parents: 
56254 
diff
changeset
 | 
419  | 
|
| 62322 | 420  | 
val register_fp_sugars_raw =  | 
| 
56648
 
2ffa440b3074
manual merge + added 'rel_distincts' field to record for symmetry
 
blanchet 
parents: 
56641 
diff
changeset
 | 
421  | 
  fold (fn fp_sugar as {T = Type (s, _), ...} =>
 | 
| 62322 | 422  | 
    Local_Theory.declaration {syntax = false, pervasive = true}
 | 
423  | 
(fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar))));  | 
|
| 58182 | 424  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
425  | 
fun register_fp_sugars plugins fp_sugars =  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
426  | 
register_fp_sugars_raw fp_sugars #> interpret_fp_sugars plugins fp_sugars;  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
427  | 
|
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
428  | 
fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs  | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
429  | 
live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs  | 
| 
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
430  | 
map_thmss common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss  | 
| 62335 | 431  | 
rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss  | 
432  | 
set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss  | 
|
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
433  | 
sel_transferss co_rec_disc_iffss co_rec_codess co_rec_transferss common_rel_co_inducts  | 
| 
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
434  | 
rel_co_inductss common_set_inducts set_inductss co_rec_o_mapss noted =  | 
| 
56637
 
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
 
blanchet 
parents: 
56523 
diff
changeset
 | 
435  | 
let  | 
| 
 
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
 
blanchet 
parents: 
56523 
diff
changeset
 | 
436  | 
val fp_sugars =  | 
| 
 
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
 
blanchet 
parents: 
56523 
diff
changeset
 | 
437  | 
map_index (fn (kk, T) =>  | 
| 
56648
 
2ffa440b3074
manual merge + added 'rel_distincts' field to record for symmetry
 
blanchet 
parents: 
56641 
diff
changeset
 | 
438  | 
        {T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk,
 | 
| 57397 | 439  | 
pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk,  | 
| 58674 | 440  | 
fp_bnf = nth (#bnfs fp_res) kk,  | 
| 57397 | 441  | 
fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs,  | 
| 58460 | 442  | 
fp_ctr_sugar =  | 
443  | 
           {ctrXs_Tss = nth ctrXs_Tsss kk,
 | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
444  | 
ctor_iff_dtor = nth ctor_iff_dtors kk,  | 
| 58460 | 445  | 
ctr_defs = nth ctr_defss kk,  | 
| 58569 | 446  | 
ctr_sugar = nth ctr_sugars kk,  | 
| 58570 | 447  | 
ctr_transfers = nth ctr_transferss kk,  | 
| 58571 | 448  | 
case_transfers = nth case_transferss kk,  | 
| 58676 | 449  | 
disc_transfers = nth disc_transferss kk,  | 
450  | 
sel_transfers = nth sel_transferss kk},  | 
|
| 58459 | 451  | 
fp_bnf_sugar =  | 
| 58462 | 452  | 
           {map_thms = nth map_thmss kk,
 | 
| 58560 | 453  | 
map_disc_iffs = nth map_disc_iffss kk,  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
454  | 
map_selss = nth map_selsss kk,  | 
| 58462 | 455  | 
rel_injects = nth rel_injectss kk,  | 
| 58562 | 456  | 
rel_distincts = nth rel_distinctss kk,  | 
| 58563 | 457  | 
rel_sels = nth rel_selss kk,  | 
| 58564 | 458  | 
rel_intros = nth rel_intross kk,  | 
| 58565 | 459  | 
rel_cases = nth rel_casess kk,  | 
| 62335 | 460  | 
pred_injects = nth pred_injectss kk,  | 
| 58566 | 461  | 
set_thms = nth set_thmss kk,  | 
| 
58671
 
4cc24f1e52d5
preserve the structure of 'set_sel' theorem in ML
 
desharna 
parents: 
58659 
diff
changeset
 | 
462  | 
set_selssss = nth set_selsssss kk,  | 
| 
58673
 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 
desharna 
parents: 
58672 
diff
changeset
 | 
463  | 
set_introssss = nth set_introsssss kk,  | 
| 58568 | 464  | 
set_cases = nth set_casess kk},  | 
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
465  | 
fp_co_induct_sugar = SOME  | 
| 58461 | 466  | 
           {co_rec = nth co_recs kk,
 | 
467  | 
common_co_inducts = common_co_inducts,  | 
|
468  | 
co_inducts = nth co_inductss kk,  | 
|
469  | 
co_rec_def = nth co_rec_defs kk,  | 
|
| 58459 | 470  | 
co_rec_thms = nth co_rec_thmss kk,  | 
471  | 
co_rec_discs = nth co_rec_discss kk,  | 
|
| 58572 | 472  | 
co_rec_disc_iffs = nth co_rec_disc_iffss kk,  | 
| 58573 | 473  | 
co_rec_selss = nth co_rec_selsss kk,  | 
474  | 
co_rec_codes = nth co_rec_codess kk,  | 
|
| 58574 | 475  | 
co_rec_transfers = nth co_rec_transferss kk,  | 
| 58734 | 476  | 
co_rec_o_maps = nth co_rec_o_mapss kk,  | 
| 58575 | 477  | 
common_rel_co_inducts = common_rel_co_inducts,  | 
| 58576 | 478  | 
rel_co_inducts = nth rel_co_inductss kk,  | 
| 58577 | 479  | 
common_set_inducts = common_set_inducts,  | 
480  | 
set_inducts = nth set_inductss kk}}  | 
|
| 57633 | 481  | 
|> morph_fp_sugar (substitute_noted_thm noted)) Ts;  | 
| 
56637
 
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
 
blanchet 
parents: 
56523 
diff
changeset
 | 
482  | 
in  | 
| 58182 | 483  | 
register_fp_sugars_raw fp_sugars  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
484  | 
#> fold (interpret_bnf plugins) (#bnfs fp_res)  | 
| 
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
485  | 
#> interpret_fp_sugars plugins fp_sugars  | 
| 
56637
 
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
 
blanchet 
parents: 
56523 
diff
changeset
 | 
486  | 
end;  | 
| 51824 | 487  | 
|
| 49622 | 488  | 
fun quasi_unambiguous_case_names names =  | 
489  | 
let  | 
|
490  | 
val ps = map (`Long_Name.base_name) names;  | 
|
491  | 
val dups = Library.duplicates (op =) (map fst ps);  | 
|
492  | 
fun underscore s =  | 
|
| 59877 | 493  | 
let val ss = Long_Name.explode s  | 
494  | 
in space_implode "_" (drop (length ss - 2) ss) end;  | 
|
| 49622 | 495  | 
in  | 
496  | 
map (fn (base, full) => if member (op =) dups base then underscore full else base) ps  | 
|
| 58285 | 497  | 
|> Name.variant_list []  | 
| 49622 | 498  | 
end;  | 
499  | 
||
| 58180 | 500  | 
fun zipper_map f =  | 
501  | 
let  | 
|
502  | 
fun zed _ [] = []  | 
|
| 58686 | 503  | 
| zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys;  | 
| 58180 | 504  | 
in zed [] end;  | 
505  | 
||
| 63787 | 506  | 
fun cannot_merge_types fp =  | 
507  | 
  error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters");
 | 
|
508  | 
||
509  | 
fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp;  | 
|
510  | 
||
511  | 
fun merge_type_args fp (As, As') =  | 
|
512  | 
if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp;  | 
|
513  | 
||
514  | 
fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs;  | 
|
515  | 
fun type_binding_of_spec (((((_, b), _), _), _), _) = b;  | 
|
516  | 
fun mixfix_of_spec ((((_, mx), _), _), _) = mx;  | 
|
517  | 
fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs;  | 
|
518  | 
fun map_binding_of_spec ((_, (b, _, _)), _) = b;  | 
|
519  | 
fun rel_binding_of_spec ((_, (_, b, _)), _) = b;  | 
|
520  | 
fun pred_binding_of_spec ((_, (_, _, b)), _) = b;  | 
|
521  | 
fun sel_default_eqs_of_spec (_, ts) = ts;  | 
|
522  | 
||
| 58675 | 523  | 
fun ctr_sugar_kind_of_fp_kind Least_FP = Datatype  | 
| 58686 | 524  | 
| ctr_sugar_kind_of_fp_kind Greatest_FP = Codatatype;  | 
| 58675 | 525  | 
|
| 
58265
 
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
 
blanchet 
parents: 
58264 
diff
changeset
 | 
526  | 
fun uncurry_thm 0 thm = thm  | 
| 
 
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
 
blanchet 
parents: 
58264 
diff
changeset
 | 
527  | 
| uncurry_thm 1 thm = thm  | 
| 
 
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
 
blanchet 
parents: 
58264 
diff
changeset
 | 
528  | 
| uncurry_thm n thm = rotate_prems ~1 (uncurry_thm (n - 1) (rotate_prems 1 (conjI RS thm)));  | 
| 
 
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
 
blanchet 
parents: 
58264 
diff
changeset
 | 
529  | 
|
| 58180 | 530  | 
fun choose_binary_fun fs AB =  | 
531  | 
find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs;  | 
|
532  | 
fun build_binary_fun_app fs t u =  | 
|
533  | 
Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u));  | 
|
534  | 
||
| 62385 | 535  | 
fun build_the_rel ctxt Rs Ts A B =  | 
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
536  | 
build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B);  | 
| 58180 | 537  | 
fun build_rel_app ctxt Rs Ts t u =  | 
| 62385 | 538  | 
build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u;  | 
| 58180 | 539  | 
|
| 63851 | 540  | 
fun build_set_app ctxt A t = Term.betapply (build_set ctxt A (fastype_of t), t);  | 
541  | 
||
| 58180 | 542  | 
fun mk_parametricity_goal ctxt Rs t u =  | 
| 62385 | 543  | 
let val prem = build_the_rel ctxt Rs [] (fastype_of t) (fastype_of u) in  | 
| 58180 | 544  | 
HOLogic.mk_Trueprop (prem $ t $ u)  | 
545  | 
end;  | 
|
546  | 
||
| 
58284
 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 
blanchet 
parents: 
58283 
diff
changeset
 | 
547  | 
val name_of_set = name_of_const "set function" domain_type;  | 
| 58180 | 548  | 
|
| 
55397
 
c2506f61fd26
generate 'fundec_cong' attribute for new-style (co)datatypes
 
blanchet 
parents: 
55394 
diff
changeset
 | 
549  | 
val fundefcong_attrs = @{attributes [fundef_cong]};
 | 
| 
54145
 
297d1c603999
make sure that registered code equations are actually equations
 
blanchet 
parents: 
54107 
diff
changeset
 | 
550  | 
val nitpicksimp_attrs = @{attributes [nitpick_simp]};
 | 
| 49300 | 551  | 
val simp_attrs = @{attributes [simp]};
 | 
552  | 
||
| 49297 | 553  | 
val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));  | 
| 
49202
 
f493cd25737f
some work towards iterator and recursor properties
 
blanchet 
parents: 
49201 
diff
changeset
 | 
554  | 
|
| 55871 | 555  | 
fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss);  | 
| 51829 | 556  | 
|
| 55871 | 557  | 
fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss  | 
558  | 
| flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) =  | 
|
559  | 
p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss;  | 
|
| 51829 | 560  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
561  | 
fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) =  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
562  | 
  Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1));
 | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
563  | 
|
| 60784 | 564  | 
fun flip_rels ctxt n thm =  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
565  | 
let  | 
| 59582 | 566  | 
val Rs = Term.add_vars (Thm.prop_of thm) [];  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
567  | 
val Rs' = rev (drop (length Rs - n) Rs);  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
568  | 
in  | 
| 60784 | 569  | 
infer_instantiate ctxt (map (fn f => (fst f, Thm.cterm_of ctxt (mk_flip f))) Rs') thm  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
570  | 
end;  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
571  | 
|
| 49536 | 572  | 
fun mk_ctor_or_dtor get_T Ts t =  | 
573  | 
let val Type (_, Ts0) = get_T (fastype_of t) in  | 
|
574  | 
Term.subst_atomic_types (Ts0 ~~ Ts) t  | 
|
575  | 
end;  | 
|
576  | 
||
577  | 
val mk_ctor = mk_ctor_or_dtor range_type;  | 
|
578  | 
val mk_dtor = mk_ctor_or_dtor domain_type;  | 
|
579  | 
||
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
580  | 
fun mk_bnf_sets bnf =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
581  | 
let  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
582  | 
val Type (T_name, Us) = T_of_bnf bnf;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
583  | 
val lives = lives_of_bnf bnf;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
584  | 
val sets = sets_of_bnf bnf;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
585  | 
fun mk_set U =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
586  | 
(case find_index (curry (op =) U) lives of  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
587  | 
~1 => Term.dummy  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
588  | 
| i => nth sets i);  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
589  | 
in  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
590  | 
(T_name, map mk_set Us)  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
591  | 
end;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
592  | 
|
| 55867 | 593  | 
fun mk_xtor_co_recs thy fp fpTs Cs ts0 =  | 
| 
51911
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
594  | 
let  | 
| 
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
595  | 
val nn = length fpTs;  | 
| 
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
596  | 
val (fpTs0, Cs0) =  | 
| 
52207
 
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
 
blanchet 
parents: 
52197 
diff
changeset
 | 
597  | 
map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0  | 
| 
51911
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
598  | 
|> split_list;  | 
| 
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
599  | 
val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs);  | 
| 
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
600  | 
in  | 
| 
52170
 
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
 
blanchet 
parents: 
52169 
diff
changeset
 | 
601  | 
map (Term.subst_TVars rho) ts0  | 
| 
51911
 
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
 
blanchet 
parents: 
51910 
diff
changeset
 | 
602  | 
end;  | 
| 51827 | 603  | 
|
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
604  | 
fun liveness_of_fp_bnf n bnf =  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
605  | 
(case T_of_bnf bnf of  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
606  | 
Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
607  | 
| _ => replicate n false);  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
608  | 
|
| 57397 | 609  | 
fun add_nesting_bnf_names Us =  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
610  | 
let  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
611  | 
fun add (Type (s, Ts)) ss =  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
612  | 
let val (needs, ss') = fold_map add Ts ss in  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
613  | 
if exists I needs then (true, insert (op =) s ss') else (false, ss')  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
614  | 
end  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
615  | 
| add T ss = (member (op =) Us T, ss);  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
616  | 
in snd oo add end;  | 
| 
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
617  | 
|
| 57397 | 618  | 
fun nesting_bnfs ctxt ctr_Tsss Us =  | 
619  | 
map_filter (bnf_of ctxt) (fold (fold (fold (add_nesting_bnf_names Us))) ctr_Tsss []);  | 
|
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
620  | 
|
| 
54237
 
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
 
blanchet 
parents: 
54236 
diff
changeset
 | 
621  | 
fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
622  | 
|
| 58346 | 623  | 
fun massage_simple_notes base =  | 
624  | 
filter_out (null o #2)  | 
|
625  | 
#> map (fn (thmN, thms, f_attrs) =>  | 
|
626  | 
((Binding.qualify true base (Binding.name thmN), []),  | 
|
627  | 
map_index (fn (i, thm) => ([thm], f_attrs i)) thms));  | 
|
628  | 
||
629  | 
fun massage_multi_notes b_names Ts =  | 
|
630  | 
maps (fn (thmN, thmss, attrs) =>  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
631  | 
    @{map 3} (fn b_name => fn Type (T_name, _) => fn thms =>
 | 
| 58346 | 632  | 
((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])]))  | 
633  | 
b_names Ts thmss)  | 
|
634  | 
#> filter_out (null o fst o hd o snd);  | 
|
635  | 
||
| 63816 | 636  | 
fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings  | 
637  | 
ctr_mixfixes ctr_Tss lthy =  | 
|
638  | 
let  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
639  | 
val ctor_absT = domain_type (fastype_of ctor);  | 
| 63816 | 640  | 
|
641  | 
val (((w, xss), u'), _) = lthy  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
642  | 
|> yield_singleton (mk_Frees "w") ctor_absT  | 
| 63816 | 643  | 
||>> mk_Freess "x" ctr_Tss  | 
644  | 
||>> yield_singleton Variable.variant_fixes fp_b_name;  | 
|
645  | 
||
646  | 
val u = Free (u', fpT);  | 
|
647  | 
||
648  | 
val ctor_iff_dtor_thm =  | 
|
649  | 
let  | 
|
650  | 
val goal =  | 
|
651  | 
fold_rev Logic.all [w, u]  | 
|
652  | 
(mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));  | 
|
653  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
654  | 
in  | 
|
655  | 
        Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} =>
 | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
656  | 
mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT])  | 
| 63816 | 657  | 
(Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor)  | 
658  | 
|> Thm.close_derivation  | 
|
659  | 
end;  | 
|
660  | 
||
661  | 
val ctr_rhss =  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
662  | 
map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs))  | 
| 63816 | 663  | 
ks xss;  | 
664  | 
||
665  | 
val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy  | 
|
666  | 
|> Local_Theory.open_target |> snd  | 
|
667  | 
      |> apfst split_list o @{fold_map 3} (fn b => fn mx => fn rhs =>
 | 
|
668  | 
Local_Theory.define ((b, mx),  | 
|
669  | 
((Thm.make_def_binding (Config.get lthy bnf_internals) b, []), rhs))  | 
|
670  | 
#>> apsnd snd) ctr_bindings ctr_mixfixes ctr_rhss  | 
|
671  | 
||> `Local_Theory.close_target;  | 
|
672  | 
||
673  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
|
674  | 
||
675  | 
val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;  | 
|
676  | 
val ctrs0 = map (Morphism.term phi) raw_ctrs;  | 
|
677  | 
in  | 
|
678  | 
((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy)  | 
|
679  | 
end;  | 
|
680  | 
||
| 63839 | 681  | 
fun wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition  | 
| 63816 | 682  | 
disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs lthy =  | 
683  | 
let  | 
|
684  | 
    val sumEN_thm' = unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms);
 | 
|
685  | 
||
686  | 
    fun exhaust_tac {context = ctxt, prems = _} =
 | 
|
687  | 
mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm';  | 
|
688  | 
||
689  | 
val inject_tacss =  | 
|
690  | 
      map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} =>
 | 
|
691  | 
mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms;  | 
|
692  | 
||
693  | 
val half_distinct_tacss =  | 
|
694  | 
      map (map (fn (def, def') => fn {context = ctxt, ...} =>
 | 
|
695  | 
mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def']))  | 
|
696  | 
(mk_half_pairss (`I ctr_defs));  | 
|
697  | 
||
698  | 
val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss;  | 
|
699  | 
||
700  | 
fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs);  | 
|
701  | 
    val ctr_specs = @{map 3} ctr_spec_of disc_bindings ctrs0 sel_bindingss;
 | 
|
702  | 
||
703  | 
    val (ctr_sugar as {case_cong, ...}, lthy) =
 | 
|
704  | 
free_constructors (ctr_sugar_kind_of_fp_kind fp) tacss  | 
|
| 63853 | 705  | 
((((plugins, discs_sels), standard_binding), ctr_specs), sel_default_eqs) lthy;  | 
| 63816 | 706  | 
|
707  | 
val anonymous_notes =  | 
|
708  | 
[([case_cong], fundefcong_attrs)]  | 
|
709  | 
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));  | 
|
710  | 
||
711  | 
val notes =  | 
|
712  | 
if Config.get lthy bnf_internals then  | 
|
713  | 
[(ctor_iff_dtorN, [ctor_iff_dtor_thm], K [])]  | 
|
714  | 
|> massage_simple_notes fp_b_name  | 
|
715  | 
else  | 
|
716  | 
[];  | 
|
717  | 
in  | 
|
718  | 
(ctr_sugar, lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)  | 
|
719  | 
end;  | 
|
720  | 
||
| 63814 | 721  | 
fun derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps  | 
| 64416 | 722  | 
fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs  | 
723  | 
live_nesting_rel_eq_onps fp_nested_rel_eq_onps fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor  | 
|
724  | 
dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm  | 
|
| 63851 | 725  | 
extra_unfolds_map extra_unfolds_set extra_unfolds_rel ctr_Tss  | 
| 58676 | 726  | 
    ({casex, case_thms, discs, selss, sel_defs, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss,
 | 
727  | 
injects, distincts, distinct_discsss, ...} : ctr_sugar)  | 
|
| 58346 | 728  | 
lthy =  | 
| 61344 | 729  | 
let  | 
730  | 
val n = length ctr_Tss;  | 
|
731  | 
val ms = map length ctr_Tss;  | 
|
732  | 
||
733  | 
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);  | 
|
734  | 
||
735  | 
val fpBT = B_ify_T fpT;  | 
|
736  | 
val live_AsBs = filter (op <>) (As ~~ Bs);  | 
|
| 62335 | 737  | 
val live_As = map fst live_AsBs;  | 
| 61344 | 738  | 
val fTs = map (op -->) live_AsBs;  | 
739  | 
||
| 62335 | 740  | 
val ((((((((xss, yss), fs), Ps), Rs), ta), tb), thesis), names_lthy) = lthy  | 
| 61344 | 741  | 
|> fold (fold Variable.declare_typ) [As, Bs]  | 
742  | 
|> mk_Freess "x" ctr_Tss  | 
|
743  | 
||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss)  | 
|
744  | 
||>> mk_Frees "f" fTs  | 
|
| 62335 | 745  | 
||>> mk_Frees "P" (map mk_pred1T live_As)  | 
| 61344 | 746  | 
||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs)  | 
747  | 
||>> yield_singleton (mk_Frees "a") fpT  | 
|
748  | 
||>> yield_singleton (mk_Frees "b") fpBT  | 
|
749  | 
||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT;  | 
|
750  | 
||
751  | 
val ctrAs = map (mk_ctr As) ctrs;  | 
|
752  | 
val ctrBs = map (mk_ctr Bs) ctrs;  | 
|
753  | 
||
| 63814 | 754  | 
val ctr_defs' = map2 (fn m => fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) ms ctr_defs;  | 
755  | 
||
| 63840 | 756  | 
val ABfs = live_AsBs ~~ fs;  | 
757  | 
||
| 62335 | 758  | 
fun derive_rel_case relAsBs rel_inject_thms rel_distinct_thms =  | 
| 61344 | 759  | 
let  | 
760  | 
val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb;  | 
|
761  | 
||
762  | 
fun mk_assms ctrA ctrB ctxt =  | 
|
763  | 
let  | 
|
764  | 
val argA_Ts = binder_types (fastype_of ctrA);  | 
|
765  | 
val argB_Ts = binder_types (fastype_of ctrB);  | 
|
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
766  | 
val ((argAs, argBs), names_ctxt) = ctxt  | 
| 61344 | 767  | 
|> mk_Frees "x" argA_Ts  | 
768  | 
||>> mk_Frees "y" argB_Ts;  | 
|
769  | 
val ctrA_args = list_comb (ctrA, argAs);  | 
|
770  | 
val ctrB_args = list_comb (ctrB, argBs);  | 
|
771  | 
in  | 
|
772  | 
(fold_rev Logic.all (argAs @ argBs) (Logic.list_implies  | 
|
773  | 
(mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) ::  | 
|
774  | 
map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs,  | 
|
775  | 
thesis)),  | 
|
776  | 
names_ctxt)  | 
|
777  | 
end;  | 
|
778  | 
||
779  | 
        val (assms, names_lthy) = @{fold_map 2} mk_assms ctrAs ctrBs names_lthy;
 | 
|
780  | 
val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis);  | 
|
781  | 
in  | 
|
782  | 
        Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
 | 
|
| 62335 | 783  | 
mk_rel_case_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects  | 
| 61344 | 784  | 
rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs)  | 
785  | 
|> singleton (Proof_Context.export names_lthy lthy)  | 
|
786  | 
|> Thm.close_derivation  | 
|
787  | 
end;  | 
|
788  | 
||
| 62335 | 789  | 
fun derive_case_transfer rel_case_thm =  | 
| 61344 | 790  | 
let  | 
791  | 
val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy;  | 
|
792  | 
val caseA = mk_case As C casex;  | 
|
793  | 
val caseB = mk_case Bs E casex;  | 
|
794  | 
val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB;  | 
|
795  | 
in  | 
|
796  | 
        Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
 | 
|
| 62335 | 797  | 
mk_case_transfer_tac ctxt rel_case_thm case_thms)  | 
| 61344 | 798  | 
|> singleton (Proof_Context.export names_lthy lthy)  | 
799  | 
|> Thm.close_derivation  | 
|
800  | 
end;  | 
|
801  | 
in  | 
|
802  | 
if live = 0 then  | 
|
| 
61348
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
803  | 
if plugins transfer_plugin then  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
804  | 
let  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
805  | 
val relAsBs = HOLogic.eq_const fpT;  | 
| 62335 | 806  | 
val rel_case_thm = derive_rel_case relAsBs [] [];  | 
807  | 
||
808  | 
val case_transfer_thm = derive_case_transfer rel_case_thm;  | 
|
| 
61348
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
809  | 
|
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
810  | 
val notes =  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
811  | 
[(case_transferN, [case_transfer_thm], K [])]  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
812  | 
|> massage_simple_notes fp_b_name;  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
813  | 
|
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
814  | 
val (noted, lthy') = lthy  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
815  | 
|> Local_Theory.notes notes;  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
816  | 
|
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
817  | 
val subst = Morphism.thm (substitute_noted_thm noted);  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
818  | 
in  | 
| 62335 | 819  | 
(([], [], [], [], [], [], [], [], [], [], [], [], [], [], [subst case_transfer_thm], [],  | 
820  | 
[]), lthy')  | 
|
| 
61348
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
821  | 
end  | 
| 
 
d7215449be83
disable generation of 'case_transfer' for 'nibble', due to quadratic proof -- to make 'HOL-Proofs' happier
 
blanchet 
parents: 
61347 
diff
changeset
 | 
822  | 
else  | 
| 62335 | 823  | 
(([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy)  | 
| 61344 | 824  | 
else  | 
825  | 
let  | 
|
826  | 
val mapx = mk_map live As Bs (map_of_bnf fp_bnf);  | 
|
827  | 
val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf);  | 
|
828  | 
val setAs = map (mk_set As) (sets_of_bnf fp_bnf);  | 
|
829  | 
val discAs = map (mk_disc_or_sel As) discs;  | 
|
830  | 
val discBs = map (mk_disc_or_sel Bs) discs;  | 
|
831  | 
val selAss = map (map (mk_disc_or_sel As)) selss;  | 
|
832  | 
val selBss = map (map (mk_disc_or_sel Bs)) selss;  | 
|
833  | 
||
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
834  | 
val map_ctor_thm =  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
835  | 
if fp = Least_FP then  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
836  | 
fp_map_thm  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
837  | 
else  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
838  | 
let  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
839  | 
val ctorA = mk_ctor As ctor;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
840  | 
val ctorB = mk_ctor Bs ctor;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
841  | 
|
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
842  | 
val y_T = domain_type (fastype_of ctorA);  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
843  | 
val (y as Free (y_s, _), _) = lthy  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
844  | 
|> yield_singleton (mk_Frees "y") y_T;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
845  | 
|
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
846  | 
val ctor_cong =  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
847  | 
infer_instantiate' lthy [NONE, NONE, SOME (Thm.cterm_of lthy ctorB)] arg_cong;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
848  | 
val fp_map_thm' = fp_map_thm  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
849  | 
|> infer_instantiate' lthy (replicate live NONE @  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
850  | 
[SOME (Thm.cterm_of lthy (ctorA $ y))])  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
851  | 
|> unfold_thms lthy [dtor_ctor];  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
852  | 
|
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
853  | 
val map_ctor_thm0 = fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans);  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
854  | 
in  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
855  | 
Thm.generalize ([], [y_s]) (Thm.maxidx_of map_ctor_thm0 + 1) map_ctor_thm0  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
856  | 
end;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
857  | 
|
| 63840 | 858  | 
val map_thms =  | 
859  | 
let  | 
|
860  | 
fun mk_goal ctrA ctrB xs ys =  | 
|
861  | 
let  | 
|
862  | 
val fmap = list_comb (mapx, fs);  | 
|
863  | 
||
864  | 
fun mk_arg (x as Free (_, T)) (Free (_, U)) =  | 
|
865  | 
if T = U then x  | 
|
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
866  | 
else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x;  | 
| 63840 | 867  | 
|
868  | 
val xs' = map2 mk_arg xs ys;  | 
|
869  | 
in  | 
|
870  | 
mk_Trueprop_eq (fmap $ list_comb (ctrA, xs), list_comb (ctrB, xs'))  | 
|
871  | 
end;  | 
|
872  | 
||
873  | 
            val goals = @{map 4} mk_goal ctrAs ctrBs xss yss;
 | 
|
874  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
875  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
876  | 
in  | 
|
877  | 
            Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
|
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
878  | 
mk_map_tac ctxt abs_inverses pre_map_def map_ctor_thm live_nesting_map_id0s ctr_defs'  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
879  | 
extra_unfolds_map)  | 
| 63840 | 880  | 
|> Thm.close_derivation  | 
881  | 
|> Conjunction.elim_balanced (length goals)  | 
|
882  | 
end;  | 
|
| 61344 | 883  | 
|
| 63851 | 884  | 
val set0_thms =  | 
885  | 
let  | 
|
886  | 
fun mk_goal A setA ctrA xs =  | 
|
887  | 
let  | 
|
888  | 
val sets = map (build_set_app lthy A)  | 
|
889  | 
(filter (exists_subtype_in [A] o fastype_of) xs);  | 
|
890  | 
in  | 
|
891  | 
mk_Trueprop_eq (setA $ list_comb (ctrA, xs),  | 
|
| 63862 | 892  | 
(if null sets then HOLogic.mk_set A [] else Library.foldl1 mk_union sets))  | 
| 63851 | 893  | 
end;  | 
894  | 
||
895  | 
val goals =  | 
|
896  | 
              @{map 2} (fn live_A => fn setA => map2 (mk_goal live_A setA) ctrAs xss) live_As setAs
 | 
|
897  | 
|> flat;  | 
|
898  | 
in  | 
|
899  | 
if null goals then  | 
|
900  | 
[]  | 
|
901  | 
else  | 
|
902  | 
let  | 
|
903  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
904  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
905  | 
in  | 
|
906  | 
                Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
|
907  | 
mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_set_thms  | 
|
908  | 
fp_nesting_set_maps live_nesting_set_maps ctr_defs' extra_unfolds_set)  | 
|
909  | 
|> Thm.close_derivation  | 
|
910  | 
|> Conjunction.elim_balanced (length goals)  | 
|
911  | 
end  | 
|
912  | 
end;  | 
|
913  | 
val set_thms = set0_thms  | 
|
914  | 
          |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left});
 | 
|
915  | 
||
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
916  | 
val rel_ctor_thm =  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
917  | 
if fp = Least_FP then  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
918  | 
fp_rel_thm  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
919  | 
else  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
920  | 
let  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
921  | 
val ctorA = mk_ctor As ctor;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
922  | 
val ctorB = mk_ctor Bs ctor;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
923  | 
|
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
924  | 
val y_T = domain_type (fastype_of ctorA);  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
925  | 
val z_T = domain_type (fastype_of ctorB);  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
926  | 
val ((y as Free (y_s, _), z as Free (z_s, _)), _) = lthy  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
927  | 
|> yield_singleton (mk_Frees "y") y_T  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
928  | 
||>> yield_singleton (mk_Frees "z") z_T;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
929  | 
|
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
930  | 
val rel_ctor_thm0 = fp_rel_thm  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
931  | 
|> infer_instantiate' lthy (replicate live NONE @  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
932  | 
[SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))])  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
933  | 
|> unfold_thms lthy [dtor_ctor];  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
934  | 
in  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
935  | 
Thm.generalize ([], [y_s, z_s]) (Thm.maxidx_of rel_ctor_thm0 + 1) rel_ctor_thm0  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
936  | 
end;  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
937  | 
|
| 63841 | 938  | 
val rel_inject_thms =  | 
939  | 
let  | 
|
940  | 
fun mk_goal ctrA ctrB xs ys =  | 
|
941  | 
let  | 
|
| 63851 | 942  | 
val lhs = list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys);  | 
| 63841 | 943  | 
val conjuncts = map2 (build_rel_app lthy Rs []) xs ys;  | 
944  | 
in  | 
|
945  | 
HOLogic.mk_Trueprop  | 
|
946  | 
(if null conjuncts then lhs  | 
|
947  | 
else HOLogic.mk_eq (lhs, Library.foldr1 HOLogic.mk_conj conjuncts))  | 
|
948  | 
end;  | 
|
949  | 
||
950  | 
            val goals = @{map 4} mk_goal ctrAs ctrBs xss yss;
 | 
|
951  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
952  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
953  | 
in  | 
|
954  | 
            Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
|
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
955  | 
mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs'  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
956  | 
extra_unfolds_rel)  | 
| 63841 | 957  | 
|> Thm.close_derivation  | 
958  | 
|> Conjunction.elim_balanced (length goals)  | 
|
959  | 
end;  | 
|
960  | 
||
961  | 
val half_rel_distinct_thmss =  | 
|
962  | 
let  | 
|
963  | 
fun mk_goal ((ctrA, xs), (ctrB, ys)) =  | 
|
| 63851 | 964  | 
HOLogic.mk_Trueprop (HOLogic.mk_not  | 
965  | 
(list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys)));  | 
|
| 63841 | 966  | 
|
967  | 
val rel_infos = (ctrAs ~~ xss, ctrBs ~~ yss);  | 
|
968  | 
||
969  | 
val goalss = map (map mk_goal) (mk_half_pairss rel_infos);  | 
|
970  | 
val goals = flat goalss;  | 
|
971  | 
in  | 
|
972  | 
unflat goalss  | 
|
| 63851 | 973  | 
(if null goals then  | 
974  | 
[]  | 
|
| 63841 | 975  | 
else  | 
976  | 
let  | 
|
977  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
978  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
979  | 
in  | 
|
980  | 
                   Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
|
| 
64415
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
981  | 
mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs  | 
| 
 
7ca48c274553
more uniform treatment of codatatype vs. datatype map and rel theorem generation (towards nonuniform codatatypes)
 
blanchet 
parents: 
64413 
diff
changeset
 | 
982  | 
ctr_defs' extra_unfolds_rel)  | 
| 63841 | 983  | 
|> Thm.close_derivation  | 
984  | 
|> Conjunction.elim_balanced (length goals)  | 
|
985  | 
end)  | 
|
986  | 
end;  | 
|
| 63840 | 987  | 
|
| 61344 | 988  | 
val rel_flip = rel_flip_of_bnf fp_bnf;  | 
989  | 
||
990  | 
fun mk_other_half_rel_distinct_thm thm =  | 
|
991  | 
          flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2);
 | 
|
992  | 
||
993  | 
val other_half_rel_distinct_thmss =  | 
|
994  | 
map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss;  | 
|
995  | 
val (rel_distinct_thms, _) =  | 
|
996  | 
join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;  | 
|
997  | 
||
| 63841 | 998  | 
fun mk_rel_intro_thm m thm =  | 
999  | 
uncurry_thm m (thm RS iffD2) handle THM _ => thm;  | 
|
1000  | 
||
1001  | 
val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms;  | 
|
1002  | 
||
| 61344 | 1003  | 
val rel_code_thms =  | 
1004  | 
          map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @
 | 
|
1005  | 
          map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms;
 | 
|
1006  | 
||
1007  | 
val ctr_transfer_thms =  | 
|
1008  | 
let  | 
|
1009  | 
val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs;  | 
|
1010  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1011  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1012  | 
in  | 
|
1013  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1014  | 
              (fn {context = ctxt, prems = _} =>
 | 
|
1015  | 
mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs)  | 
|
1016  | 
|> Thm.close_derivation  | 
|
1017  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1018  | 
end;  | 
|
1019  | 
||
1020  | 
val (set_cases_thms, set_cases_attrss) =  | 
|
1021  | 
let  | 
|
1022  | 
fun mk_prems assms elem t ctxt =  | 
|
1023  | 
(case fastype_of t of  | 
|
1024  | 
Type (type_name, xs) =>  | 
|
1025  | 
(case bnf_of ctxt type_name of  | 
|
1026  | 
NONE => ([], ctxt)  | 
|
1027  | 
| SOME bnf =>  | 
|
1028  | 
apfst flat (fold_map (fn set => fn ctxt =>  | 
|
1029  | 
let  | 
|
1030  | 
val T = HOLogic.dest_setT (range_type (fastype_of set));  | 
|
1031  | 
val new_var = not (T = fastype_of elem);  | 
|
1032  | 
val (x, ctxt') =  | 
|
1033  | 
if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt);  | 
|
1034  | 
in  | 
|
1035  | 
mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt'  | 
|
1036  | 
|>> map (new_var ? Logic.all x)  | 
|
1037  | 
end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt))  | 
|
1038  | 
| T => rpair ctxt  | 
|
| 63851 | 1039  | 
(if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else []));  | 
| 61344 | 1040  | 
in  | 
1041  | 
split_list (map (fn set =>  | 
|
1042  | 
let  | 
|
1043  | 
val A = HOLogic.dest_setT (range_type (fastype_of set));  | 
|
1044  | 
val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy;  | 
|
1045  | 
val premss =  | 
|
1046  | 
map (fn ctr =>  | 
|
1047  | 
let  | 
|
1048  | 
val (args, names_lthy) =  | 
|
1049  | 
mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy;  | 
|
1050  | 
in  | 
|
1051  | 
flat (zipper_map (fn (prev_args, arg, next_args) =>  | 
|
1052  | 
let  | 
|
1053  | 
val (args_with_elem, args_without_elem) =  | 
|
1054  | 
if fastype_of arg = A then  | 
|
1055  | 
(prev_args @ [elem] @ next_args, prev_args @ next_args)  | 
|
1056  | 
else  | 
|
1057  | 
`I (prev_args @ [arg] @ next_args);  | 
|
1058  | 
in  | 
|
1059  | 
mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))]  | 
|
1060  | 
elem arg names_lthy  | 
|
1061  | 
|> fst  | 
|
1062  | 
|> map (fold_rev Logic.all args_without_elem)  | 
|
1063  | 
end) args)  | 
|
1064  | 
end) ctrAs;  | 
|
1065  | 
val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis);  | 
|
1066  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1067  | 
val thm =  | 
|
1068  | 
                  Goal.prove_sorry lthy vars (flat premss) goal (fn {context = ctxt, prems} =>
 | 
|
1069  | 
mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms)  | 
|
1070  | 
|> Thm.close_derivation  | 
|
1071  | 
|> rotate_prems ~1;  | 
|
1072  | 
||
1073  | 
val cases_set_attr =  | 
|
1074  | 
Attrib.internal (K (Induct.cases_pred (name_of_set set)));  | 
|
1075  | 
||
1076  | 
val ctr_names = quasi_unambiguous_case_names (flat  | 
|
1077  | 
(map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs)));  | 
|
1078  | 
in  | 
|
1079  | 
                (* TODO: @{attributes [elim?]} *)
 | 
|
| 63019 | 1080  | 
(thm, [Attrib.consumes 1, cases_set_attr, Attrib.case_names ctr_names])  | 
| 61344 | 1081  | 
end) setAs)  | 
1082  | 
end;  | 
|
1083  | 
||
1084  | 
val (set_intros_thmssss, set_intros_thms) =  | 
|
1085  | 
let  | 
|
1086  | 
fun mk_goals A setA ctr_args t ctxt =  | 
|
1087  | 
(case fastype_of t of  | 
|
1088  | 
Type (type_name, innerTs) =>  | 
|
1089  | 
(case bnf_of ctxt type_name of  | 
|
1090  | 
NONE => ([], ctxt)  | 
|
1091  | 
| SOME bnf =>  | 
|
1092  | 
apfst flat (fold_map (fn set => fn ctxt =>  | 
|
1093  | 
let  | 
|
1094  | 
val T = HOLogic.dest_setT (range_type (fastype_of set));  | 
|
| 63838 | 1095  | 
val (y, ctxt') = yield_singleton (mk_Frees "y") T ctxt;  | 
1096  | 
val assm = mk_Trueprop_mem (y, set $ t);  | 
|
| 61344 | 1097  | 
in  | 
| 63838 | 1098  | 
apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args y ctxt')  | 
| 61344 | 1099  | 
end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt))  | 
1100  | 
| T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt));  | 
|
1101  | 
||
1102  | 
val (goalssss, _) =  | 
|
1103  | 
fold_map (fn set =>  | 
|
1104  | 
let val A = HOLogic.dest_setT (range_type (fastype_of set)) in  | 
|
| 63838 | 1105  | 
                  @{fold_map 2} (fn ctr => fn xs =>
 | 
1106  | 
fold_map (mk_goals A set (Term.list_comb (ctr, xs))) xs)  | 
|
1107  | 
ctrAs xss  | 
|
| 61344 | 1108  | 
end) setAs lthy;  | 
1109  | 
val goals = flat (flat (flat goalssss));  | 
|
1110  | 
in  | 
|
| 63797 | 1111  | 
`(unflattt goalssss)  | 
| 63851 | 1112  | 
(if null goals then  | 
1113  | 
[]  | 
|
| 61344 | 1114  | 
else  | 
1115  | 
let  | 
|
1116  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1117  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1118  | 
in  | 
|
1119  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1120  | 
                     (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms)
 | 
|
1121  | 
|> Thm.close_derivation  | 
|
1122  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1123  | 
end)  | 
|
1124  | 
end;  | 
|
1125  | 
||
1126  | 
val rel_sel_thms =  | 
|
1127  | 
let  | 
|
1128  | 
val n = length discAs;  | 
|
1129  | 
fun mk_conjunct n k discA selAs discB selBs =  | 
|
1130  | 
(if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @  | 
|
| 63851 | 1131  | 
(if null selAs then  | 
1132  | 
[]  | 
|
| 61344 | 1133  | 
else  | 
1134  | 
[Library.foldr HOLogic.mk_imp  | 
|
1135  | 
(if n = 1 then [] else [discA $ ta, discB $ tb],  | 
|
1136  | 
Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs [])  | 
|
1137  | 
(map (rapp ta) selAs) (map (rapp tb) selBs)))]);  | 
|
1138  | 
||
1139  | 
val goals =  | 
|
| 62335 | 1140  | 
if n = 0 then  | 
1141  | 
[]  | 
|
| 61344 | 1142  | 
else  | 
1143  | 
[mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb,  | 
|
1144  | 
                   (case flat (@{map 5} (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of
 | 
|
1145  | 
                     [] => @{term True}
 | 
|
1146  | 
| conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))];  | 
|
1147  | 
||
1148  | 
fun prove goal =  | 
|
1149  | 
Variable.add_free_names lthy goal []  | 
|
1150  | 
              |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} =>
 | 
|
1151  | 
mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust  | 
|
1152  | 
(flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms  | 
|
1153  | 
live_nesting_rel_eqs))  | 
|
1154  | 
|> Thm.close_derivation;  | 
|
1155  | 
in  | 
|
1156  | 
map prove goals  | 
|
1157  | 
end;  | 
|
1158  | 
||
| 62335 | 1159  | 
val (rel_case_thm, rel_case_attrs) =  | 
| 61344 | 1160  | 
let  | 
| 62335 | 1161  | 
val thm = derive_rel_case relAsBs rel_inject_thms rel_distinct_thms;  | 
| 61344 | 1162  | 
val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs);  | 
1163  | 
in  | 
|
| 63019 | 1164  | 
            (thm, [Attrib.case_names ctr_names, Attrib.consumes 1] @ @{attributes [cases pred]})
 | 
| 61344 | 1165  | 
end;  | 
1166  | 
||
| 62335 | 1167  | 
val case_transfer_thm = derive_case_transfer rel_case_thm;  | 
| 61344 | 1168  | 
|
1169  | 
val sel_transfer_thms =  | 
|
| 62335 | 1170  | 
if null selAss then  | 
1171  | 
[]  | 
|
| 61344 | 1172  | 
else  | 
| 58346 | 1173  | 
let  | 
| 61344 | 1174  | 
val shared_sels = foldl1 (uncurry (inter (op =))) (map (op ~~) (selAss ~~ selBss));  | 
1175  | 
val goals = map (uncurry (mk_parametricity_goal names_lthy Rs)) shared_sels;  | 
|
| 58346 | 1176  | 
in  | 
| 62335 | 1177  | 
if null goals then  | 
1178  | 
[]  | 
|
| 61344 | 1179  | 
else  | 
1180  | 
let  | 
|
1181  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1182  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1183  | 
in  | 
|
1184  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1185  | 
                    (fn {context = ctxt, prems = _} =>
 | 
|
1186  | 
mk_sel_transfer_tac ctxt n sel_defs case_transfer_thm)  | 
|
1187  | 
|> Thm.close_derivation  | 
|
1188  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1189  | 
end  | 
|
1190  | 
end;  | 
|
1191  | 
||
1192  | 
val disc_transfer_thms =  | 
|
1193  | 
let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in  | 
|
1194  | 
if null goals then  | 
|
1195  | 
[]  | 
|
| 58346 | 1196  | 
else  | 
| 61344 | 1197  | 
let  | 
1198  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1199  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1200  | 
in  | 
|
1201  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1202  | 
                  (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt
 | 
|
1203  | 
(the_single rel_sel_thms) (the_single exhaust_discs)  | 
|
1204  | 
(flat (flat distinct_discsss)))  | 
|
1205  | 
|> Thm.close_derivation  | 
|
1206  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1207  | 
end  | 
|
1208  | 
end;  | 
|
1209  | 
||
1210  | 
val map_disc_iff_thms =  | 
|
| 58676 | 1211  | 
let  | 
| 61344 | 1212  | 
val discsB = map (mk_disc_or_sel Bs) discs;  | 
1213  | 
val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs;  | 
|
1214  | 
||
1215  | 
fun mk_goal (discA_t, discB) =  | 
|
1216  | 
if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then  | 
|
1217  | 
NONE  | 
|
1218  | 
else  | 
|
1219  | 
SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t));  | 
|
1220  | 
||
1221  | 
val goals = map_filter mk_goal (discsA_t ~~ discsB);  | 
|
| 58676 | 1222  | 
in  | 
| 61344 | 1223  | 
if null goals then  | 
1224  | 
[]  | 
|
| 58676 | 1225  | 
else  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1226  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1227  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1228  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1229  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1230  | 
Goal.prove_sorry lthy vars [] goal  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1231  | 
                  (fn {context = ctxt, prems = _} =>
 | 
| 61344 | 1232  | 
mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss)  | 
1233  | 
map_thms)  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1234  | 
|> Thm.close_derivation  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1235  | 
|> Conjunction.elim_balanced (length goals)  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1236  | 
end  | 
| 58676 | 1237  | 
end;  | 
1238  | 
||
| 61344 | 1239  | 
val (map_sel_thmss, map_sel_thms) =  | 
1240  | 
let  | 
|
1241  | 
fun mk_goal discA selA selB =  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1242  | 
let  | 
| 61344 | 1243  | 
val prem = Term.betapply (discA, ta);  | 
1244  | 
val lhs = selB $ (Term.list_comb (mapx, fs) $ ta);  | 
|
1245  | 
val lhsT = fastype_of lhs;  | 
|
1246  | 
val map_rhsT =  | 
|
1247  | 
map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT;  | 
|
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
1248  | 
val map_rhs = build_map lthy [] []  | 
| 61344 | 1249  | 
(the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT);  | 
1250  | 
val rhs = (case map_rhs of  | 
|
1251  | 
                    Const (@{const_name id}, _) => selA $ ta
 | 
|
1252  | 
| _ => map_rhs $ (selA $ ta));  | 
|
1253  | 
val concl = mk_Trueprop_eq (lhs, rhs);  | 
|
1254  | 
in  | 
|
1255  | 
if is_refl_bool prem then concl  | 
|
1256  | 
else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl)  | 
|
1257  | 
end;  | 
|
1258  | 
||
1259  | 
            val goalss = @{map 3} (map2 o mk_goal) discAs selAss selBss;
 | 
|
1260  | 
val goals = flat goalss;  | 
|
1261  | 
in  | 
|
| 63797 | 1262  | 
`(unflat goalss)  | 
| 63851 | 1263  | 
(if null goals then  | 
1264  | 
[]  | 
|
1265  | 
else  | 
|
1266  | 
let  | 
|
1267  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1268  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1269  | 
in  | 
|
1270  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1271  | 
                     (fn {context = ctxt, prems = _} =>
 | 
|
1272  | 
mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss)  | 
|
1273  | 
map_thms (flat sel_thmss) live_nesting_map_id0s)  | 
|
1274  | 
|> Thm.close_derivation  | 
|
1275  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1276  | 
end)  | 
|
| 61344 | 1277  | 
end;  | 
1278  | 
||
1279  | 
val (set_sel_thmssss, set_sel_thms) =  | 
|
1280  | 
let  | 
|
1281  | 
fun mk_goal setA discA selA ctxt =  | 
|
1282  | 
let  | 
|
1283  | 
val prem = Term.betapply (discA, ta);  | 
|
1284  | 
val sel_rangeT = range_type (fastype_of selA);  | 
|
1285  | 
val A = HOLogic.dest_setT (range_type (fastype_of setA));  | 
|
1286  | 
||
1287  | 
fun travese_nested_types t ctxt =  | 
|
1288  | 
(case fastype_of t of  | 
|
1289  | 
Type (type_name, innerTs) =>  | 
|
1290  | 
(case bnf_of ctxt type_name of  | 
|
1291  | 
NONE => ([], ctxt)  | 
|
1292  | 
| SOME bnf =>  | 
|
1293  | 
let  | 
|
1294  | 
fun seq_assm a set ctxt =  | 
|
1295  | 
let  | 
|
1296  | 
val T = HOLogic.dest_setT (range_type (fastype_of set));  | 
|
1297  | 
val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt;  | 
|
1298  | 
val assm = mk_Trueprop_mem (x, set $ a);  | 
|
1299  | 
in  | 
|
1300  | 
travese_nested_types x ctxt'  | 
|
1301  | 
|>> map (Logic.mk_implies o pair assm)  | 
|
1302  | 
end;  | 
|
1303  | 
in  | 
|
1304  | 
fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt  | 
|
1305  | 
|>> flat  | 
|
1306  | 
end)  | 
|
1307  | 
| T =>  | 
|
1308  | 
if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt));  | 
|
1309  | 
||
1310  | 
val (concls, ctxt') =  | 
|
1311  | 
if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt)  | 
|
1312  | 
else travese_nested_types (selA $ ta) ctxt;  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1313  | 
in  | 
| 61344 | 1314  | 
if exists_subtype_in [A] sel_rangeT then  | 
1315  | 
if is_refl_bool prem then (concls, ctxt')  | 
|
1316  | 
else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt')  | 
|
1317  | 
else  | 
|
1318  | 
([], ctxt)  | 
|
1319  | 
end;  | 
|
1320  | 
||
1321  | 
val (goalssss, _) =  | 
|
1322  | 
              fold_map (fn set => @{fold_map 2} (fold_map o mk_goal set) discAs selAss)
 | 
|
1323  | 
setAs names_lthy;  | 
|
1324  | 
val goals = flat (flat (flat goalssss));  | 
|
1325  | 
in  | 
|
| 63797 | 1326  | 
`(unflattt goalssss)  | 
| 63851 | 1327  | 
(if null goals then  | 
1328  | 
[]  | 
|
1329  | 
else  | 
|
1330  | 
let  | 
|
1331  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
|
1332  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
1333  | 
in  | 
|
1334  | 
Goal.prove_sorry lthy vars [] goal  | 
|
1335  | 
                     (fn {context = ctxt, prems = _} =>
 | 
|
1336  | 
mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss)  | 
|
1337  | 
(flat sel_thmss) set0_thms)  | 
|
1338  | 
|> Thm.close_derivation  | 
|
1339  | 
|> Conjunction.elim_balanced (length goals)  | 
|
1340  | 
end)  | 
|
| 61344 | 1341  | 
end;  | 
1342  | 
||
| 62335 | 1343  | 
val pred_injects =  | 
1344  | 
let  | 
|
1345  | 
fun top_sweep_rewr_conv rewrs =  | 
|
1346  | 
              Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) @{context};
 | 
|
1347  | 
||
1348  | 
val rel_eq_onp_with_tops_of = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv  | 
|
1349  | 
              (top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]})));
 | 
|
1350  | 
||
| 
63845
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1351  | 
val eq_onps = map rel_eq_onp_with_tops_of  | 
| 
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1352  | 
(map rel_eq_onp_of_bnf fp_bnfs @ fp_nesting_rel_eq_onps @ live_nesting_rel_eq_onps @  | 
| 
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1353  | 
fp_nested_rel_eq_onps);  | 
| 62335 | 1354  | 
val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) live_As);  | 
1355  | 
val cts = map (SOME o Thm.cterm_of lthy) (map mk_eq_onp Ps);  | 
|
1356  | 
||
1357  | 
val get_rhs = Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> snd;  | 
|
1358  | 
||
1359  | 
val pred_eq_onp_conj =  | 
|
1360  | 
              List.foldr (fn (_, thm) => thm RS @{thm eq_onp_live_step}) @{thm refl[of True]};
 | 
|
1361  | 
||
1362  | 
fun predify_rel_inject rel_inject =  | 
|
1363  | 
let  | 
|
1364  | 
val conjuncts = try (get_rhs #> HOLogic.dest_conj) rel_inject |> the_default [];  | 
|
1365  | 
||
1366  | 
fun postproc thm =  | 
|
| 
63845
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1367  | 
if null conjuncts then  | 
| 
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1368  | 
                    thm RS (@{thm eq_onp_same_args} RS iffD1)
 | 
| 
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1369  | 
else  | 
| 62335 | 1370  | 
                    @{thm box_equals} OF [thm, @{thm eq_onp_same_args},
 | 
| 
63845
 
61a03e429cbd
generalized code towards nonuniform (co)datatypes
 
blanchet 
parents: 
63842 
diff
changeset
 | 
1371  | 
                      pred_eq_onp_conj conjuncts |> unfold_thms lthy @{thms simp_thms(21)}];
 | 
| 62335 | 1372  | 
in  | 
1373  | 
rel_inject  | 
|
1374  | 
|> Thm.instantiate' cTs cts  | 
|
1375  | 
|> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv  | 
|
1376  | 
(Raw_Simplifier.rewrite lthy false  | 
|
1377  | 
                     @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]})))
 | 
|
1378  | 
|> unfold_thms lthy eq_onps  | 
|
1379  | 
|> postproc  | 
|
1380  | 
                |> unfold_thms lthy @{thms top_conj}
 | 
|
1381  | 
end;  | 
|
1382  | 
in  | 
|
1383  | 
rel_inject_thms  | 
|
1384  | 
            |> map (unfold_thms lthy [@{thm conj_assoc}])
 | 
|
1385  | 
|> map predify_rel_inject  | 
|
1386  | 
|> Proof_Context.export names_lthy lthy  | 
|
1387  | 
end;  | 
|
1388  | 
||
| 61344 | 1389  | 
val anonymous_notes =  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
1390  | 
[(rel_code_thms, nitpicksimp_attrs)]  | 
| 61344 | 1391  | 
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));  | 
1392  | 
||
1393  | 
val notes =  | 
|
| 63851 | 1394  | 
(if Config.get lthy bnf_internals then  | 
1395  | 
[(set0N, set0_thms, K [])]  | 
|
1396  | 
else  | 
|
1397  | 
[]) @  | 
|
| 61344 | 1398  | 
[(case_transferN, [case_transfer_thm], K []),  | 
1399  | 
(ctr_transferN, ctr_transfer_thms, K []),  | 
|
1400  | 
(disc_transferN, disc_transfer_thms, K []),  | 
|
1401  | 
(sel_transferN, sel_transfer_thms, K []),  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
1402  | 
(mapN, map_thms, K (nitpicksimp_attrs @ simp_attrs)),  | 
| 61344 | 1403  | 
(map_disc_iffN, map_disc_iff_thms, K simp_attrs),  | 
1404  | 
(map_selN, map_sel_thms, K []),  | 
|
| 62335 | 1405  | 
(pred_injectN, pred_injects, K simp_attrs),  | 
1406  | 
(rel_casesN, [rel_case_thm], K rel_case_attrs),  | 
|
| 61344 | 1407  | 
(rel_distinctN, rel_distinct_thms, K simp_attrs),  | 
1408  | 
(rel_injectN, rel_inject_thms, K simp_attrs),  | 
|
1409  | 
(rel_introsN, rel_intro_thms, K []),  | 
|
1410  | 
(rel_selN, rel_sel_thms, K []),  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
1411  | 
(setN, set_thms, K (case_fp fp nitpicksimp_attrs [] @ simp_attrs)),  | 
| 61344 | 1412  | 
(set_casesN, set_cases_thms, nth set_cases_attrss),  | 
1413  | 
(set_introsN, set_intros_thms, K []),  | 
|
1414  | 
(set_selN, set_sel_thms, K [])]  | 
|
1415  | 
|> massage_simple_notes fp_b_name;  | 
|
1416  | 
||
1417  | 
val (noted, lthy') = lthy  | 
|
1418  | 
|> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) map_thms)  | 
|
1419  | 
|> fp = Least_FP  | 
|
1420  | 
? Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) rel_code_thms)  | 
|
1421  | 
|> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) set0_thms)  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
1422  | 
|> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (rel_code_thms @ map_thms @ set_thms))  | 
| 61344 | 1423  | 
|> Local_Theory.notes (anonymous_notes @ notes);  | 
1424  | 
||
1425  | 
val subst = Morphism.thm (substitute_noted_thm noted);  | 
|
1426  | 
in  | 
|
1427  | 
((map subst map_thms,  | 
|
1428  | 
map subst map_disc_iff_thms,  | 
|
1429  | 
map (map subst) map_sel_thmss,  | 
|
1430  | 
map subst rel_inject_thms,  | 
|
1431  | 
map subst rel_distinct_thms,  | 
|
1432  | 
map subst rel_sel_thms,  | 
|
1433  | 
map subst rel_intro_thms,  | 
|
| 62335 | 1434  | 
[subst rel_case_thm],  | 
1435  | 
map subst pred_injects,  | 
|
| 61344 | 1436  | 
map subst set_thms,  | 
1437  | 
map (map (map (map subst))) set_sel_thmssss,  | 
|
1438  | 
map (map (map (map subst))) set_intros_thmssss,  | 
|
1439  | 
map subst set_cases_thms,  | 
|
1440  | 
map subst ctr_transfer_thms,  | 
|
1441  | 
[subst case_transfer_thm],  | 
|
1442  | 
map subst disc_transfer_thms,  | 
|
1443  | 
map subst sel_transfer_thms), lthy')  | 
|
1444  | 
end  | 
|
1445  | 
end;  | 
|
| 58346 | 1446  | 
|
| 58214 | 1447  | 
type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list);  | 
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
1448  | 
|
| 55867 | 1449  | 
fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) =  | 
| 54256 | 1450  | 
((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),  | 
| 58291 | 1451  | 
(map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms;  | 
| 54256 | 1452  | 
|
| 58115 | 1453  | 
val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism;  | 
| 54256 | 1454  | 
|
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
1455  | 
type gfp_sugar_thms =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
1456  | 
((thm list * thm) list * (Token.src list * Token.src list))  | 
| 57489 | 1457  | 
* thm list list  | 
1458  | 
* thm list list  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
1459  | 
* (thm list list * Token.src list)  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57999 
diff
changeset
 | 
1460  | 
* (thm list list list * Token.src list);  | 
| 
53746
 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 
blanchet 
parents: 
53741 
diff
changeset
 | 
1461  | 
|
| 
57260
 
8747af0d1012
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
 
blanchet 
parents: 
57206 
diff
changeset
 | 
1462  | 
fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
1463  | 
corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
1464  | 
(corec_selsss, corec_sel_attrs)) =  | 
| 54256 | 1465  | 
((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,  | 
| 
57260
 
8747af0d1012
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
 
blanchet 
parents: 
57206 
diff
changeset
 | 
1466  | 
coinduct_attrs_pair),  | 
| 57489 | 1467  | 
map (map (Morphism.thm phi)) corecss,  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
1468  | 
map (map (Morphism.thm phi)) corec_discss,  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
1469  | 
(map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
1470  | 
(map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms;  | 
| 54256 | 1471  | 
|
| 58115 | 1472  | 
val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism;  | 
| 54256 | 1473  | 
|
| 
58352
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1474  | 
fun unzip_recT (Type (@{type_name prod}, [_, TFree x]))
 | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1475  | 
      (T as Type (@{type_name prod}, Ts as [_, TFree y])) =
 | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1476  | 
if x = y then [T] else Ts  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1477  | 
  | unzip_recT _ (Type (@{type_name prod}, Ts as [_, TFree _])) = Ts
 | 
| 58159 | 1478  | 
| unzip_recT _ T = [T];  | 
1479  | 
||
| 61551 | 1480  | 
fun mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts =  | 
| 51832 | 1481  | 
let  | 
| 
52214
 
4cc5a80bba80
generalized recursors, effectively reverting inductive half of c7a034d01936
 
blanchet 
parents: 
52213 
diff
changeset
 | 
1482  | 
val Css = map2 replicate ns Cs;  | 
| 55869 | 1483  | 
val x_Tssss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1484  | 
      @{map 6} (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_rec_fun_T =>
 | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
1485  | 
map2 (map2 unzip_recT)  | 
| 55867 | 1486  | 
ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T)))  | 
1487  | 
absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts;  | 
|
| 51832 | 1488  | 
|
| 55869 | 1489  | 
val x_Tsss' = map (map flat_rec_arg_args) x_Tssss;  | 
1490  | 
val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css;  | 
|
| 51832 | 1491  | 
|
| 61551 | 1492  | 
val ((fss, xssss), _) = ctxt  | 
| 55869 | 1493  | 
|> mk_Freess "f" f_Tss  | 
1494  | 
||>> mk_Freessss "x" x_Tssss;  | 
|
| 51832 | 1495  | 
in  | 
| 61551 | 1496  | 
(f_Tss, x_Tssss, fss, xssss)  | 
| 51832 | 1497  | 
end;  | 
1498  | 
||
| 58159 | 1499  | 
fun unzip_corecT (Type (@{type_name sum}, _)) T = [T]
 | 
1500  | 
  | unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts
 | 
|
1501  | 
| unzip_corecT _ T = [T];  | 
|
1502  | 
||
| 55867 | 1503  | 
(*avoid "'a itself" arguments in corecursors*)  | 
| 55966 | 1504  | 
fun repair_nullary_single_ctr [[]] = [[HOLogic.unitT]]  | 
| 
55772
 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 
blanchet 
parents: 
55701 
diff
changeset
 | 
1505  | 
| repair_nullary_single_ctr Tss = Tss;  | 
| 
 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 
blanchet 
parents: 
55701 
diff
changeset
 | 
1506  | 
|
| 55867 | 1507  | 
fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts =  | 
| 51829 | 1508  | 
let  | 
| 
55772
 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 
blanchet 
parents: 
55701 
diff
changeset
 | 
1509  | 
val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss;  | 
| 55869 | 1510  | 
val g_absTs = map range_type fun_Ts;  | 
| 61760 | 1511  | 
val g_Tsss =  | 
1512  | 
      map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs);
 | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1513  | 
    val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
 | 
| 55869 | 1514  | 
Cs ctr_Tsss' g_Tsss;  | 
1515  | 
val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss;  | 
|
| 52870 | 1516  | 
in  | 
| 55869 | 1517  | 
(q_Tssss, g_Tsss, g_Tssss, g_absTs)  | 
| 52870 | 1518  | 
end;  | 
| 51829 | 1519  | 
|
| 55867 | 1520  | 
fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;  | 
| 52898 | 1521  | 
|
| 55867 | 1522  | 
fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec =  | 
1523  | 
(mk_corec_p_pred_types Cs ns,  | 
|
1524  | 
mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss  | 
|
1525  | 
(binder_fun_types (fastype_of dtor_corec)));  | 
|
| 51829 | 1526  | 
|
| 61551 | 1527  | 
fun mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts =  | 
| 55867 | 1528  | 
let  | 
1529  | 
val p_Tss = mk_corec_p_pred_types Cs ns;  | 
|
| 51829 | 1530  | 
|
| 55869 | 1531  | 
val (q_Tssss, g_Tsss, g_Tssss, corec_types) =  | 
| 55867 | 1532  | 
mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts;  | 
| 51831 | 1533  | 
|
| 61551 | 1534  | 
val (((((Free (x, _), cs), pss), qssss), gssss), _) = ctxt  | 
| 55869 | 1535  | 
|> yield_singleton (mk_Frees "x") dummyT  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
1536  | 
||>> mk_Frees "a" Cs  | 
| 51829 | 1537  | 
||>> mk_Freess "p" p_Tss  | 
| 55869 | 1538  | 
||>> mk_Freessss "q" q_Tssss  | 
1539  | 
||>> mk_Freessss "g" g_Tssss;  | 
|
| 51831 | 1540  | 
|
1541  | 
val cpss = map2 (map o rapp) cs pss;  | 
|
| 51829 | 1542  | 
|
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
1543  | 
fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd);  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
1544  | 
|
| 55871 | 1545  | 
fun build_dtor_corec_arg _ [] [cg] = cg  | 
1546  | 
| build_dtor_corec_arg T [cq] [cg, cg'] =  | 
|
1547  | 
mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg)  | 
|
1548  | 
(build_sum_inj Inr_const (fastype_of cg', T) $ cg');  | 
|
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
1549  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1550  | 
    val pgss = @{map 3} flat_corec_preds_predsss_gettersss pss qssss gssss;
 | 
| 55869 | 1551  | 
val cqssss = map2 (map o map o map o rapp) cs qssss;  | 
1552  | 
val cgssss = map2 (map o map o map o rapp) cs gssss;  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1553  | 
    val cqgsss = @{map 3} (@{map 3} (@{map 3} build_dtor_corec_arg)) g_Tsss cqssss cgssss;
 | 
| 51831 | 1554  | 
in  | 
| 61551 | 1555  | 
(x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types))  | 
| 51831 | 1556  | 
end;  | 
| 51829 | 1557  | 
|
| 61551 | 1558  | 
fun mk_co_recs_prelims ctxt fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 =  | 
| 51903 | 1559  | 
let  | 
| 61551 | 1560  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 52169 | 1561  | 
|
| 55867 | 1562  | 
val (xtor_co_rec_fun_Ts, xtor_co_recs) =  | 
1563  | 
mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd);  | 
|
| 51903 | 1564  | 
|
| 61551 | 1565  | 
val (recs_args_types, corecs_args_types) =  | 
| 
52207
 
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
 
blanchet 
parents: 
52197 
diff
changeset
 | 
1566  | 
if fp = Least_FP then  | 
| 61551 | 1567  | 
mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts  | 
1568  | 
|> (rpair NONE o SOME)  | 
|
| 51903 | 1569  | 
else  | 
| 61551 | 1570  | 
mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts  | 
1571  | 
|> (pair NONE o SOME);  | 
|
| 51903 | 1572  | 
in  | 
| 61551 | 1573  | 
(xtor_co_recs, recs_args_types, corecs_args_types)  | 
| 51903 | 1574  | 
end;  | 
1575  | 
||
| 55869 | 1576  | 
fun mk_preds_getterss_join c cps absT abs cqgss =  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
1577  | 
let  | 
| 55869 | 1578  | 
val n = length cqgss;  | 
1579  | 
val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqgss;  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
1580  | 
in  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55772 
diff
changeset
 | 
1581  | 
Term.lambda c (mk_IfN absT cps ts)  | 
| 51900 | 1582  | 
end;  | 
| 51886 | 1583  | 
|
| 58210 | 1584  | 
fun define_co_rec_as fp Cs fpT b rhs lthy0 =  | 
| 51897 | 1585  | 
let  | 
| 
52170
 
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
 
blanchet 
parents: 
52169 
diff
changeset
 | 
1586  | 
val thy = Proof_Context.theory_of lthy0;  | 
| 
 
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
 
blanchet 
parents: 
52169 
diff
changeset
 | 
1587  | 
|
| 55864 | 1588  | 
val ((cst, (_, def)), (lthy', lthy)) = lthy0  | 
| 
61101
 
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
 
traytel 
parents: 
60801 
diff
changeset
 | 
1589  | 
|> Local_Theory.open_target |> snd  | 
| 62093 | 1590  | 
|> Local_Theory.define  | 
1591  | 
((b, NoSyn), ((Thm.make_def_binding (Config.get lthy0 bnf_internals) b, []), rhs))  | 
|
| 
61101
 
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
 
traytel 
parents: 
60801 
diff
changeset
 | 
1592  | 
||> `Local_Theory.close_target;  | 
| 52327 | 1593  | 
|
1594  | 
val phi = Proof_Context.export_morphism lthy lthy';  | 
|
1595  | 
||
| 
58211
 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 
blanchet 
parents: 
58210 
diff
changeset
 | 
1596  | 
val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst);  | 
| 55864 | 1597  | 
val def' = Morphism.thm phi def;  | 
| 52327 | 1598  | 
in  | 
| 55864 | 1599  | 
((cst', def'), lthy')  | 
| 52327 | 1600  | 
end;  | 
1601  | 
||
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1602  | 
fun define_rec (_, _, fss, xssss) mk_binding fpTs Cs reps ctor_rec =  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1603  | 
let  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1604  | 
val nn = length fpTs;  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1605  | 
val (ctor_rec_absTs, fpT) = strip_typeN nn (fastype_of ctor_rec)  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1606  | 
|>> map domain_type ||> domain_type;  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1607  | 
in  | 
| 58210 | 1608  | 
define_co_rec_as Least_FP Cs fpT (mk_binding recN)  | 
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1609  | 
(fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_rec,  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1610  | 
         @{map 4} (fn ctor_rec_absT => fn rep => fn fs => fn xsss =>
 | 
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1611  | 
mk_case_absumprod ctor_rec_absT rep fs (map (map HOLogic.mk_tuple) xsss)  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1612  | 
(map flat_rec_arg_args xsss))  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1613  | 
ctor_rec_absTs reps fss xssss)))  | 
| 
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1614  | 
end;  | 
| 51897 | 1615  | 
|
| 61760 | 1616  | 
fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss  | 
1617  | 
dtor_corec =  | 
|
| 51897 | 1618  | 
let  | 
| 51899 | 1619  | 
val nn = length fpTs;  | 
| 55867 | 1620  | 
val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec)));  | 
| 51897 | 1621  | 
in  | 
| 58210 | 1622  | 
define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN)  | 
| 55869 | 1623  | 
(fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec,  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1624  | 
         @{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss)))
 | 
| 52320 | 1625  | 
end;  | 
| 51897 | 1626  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1627  | 
fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0)))  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1628  | 
(Type (_, Xs_Ts0)) =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1629  | 
(case AList.lookup (op =) setss_fp_nesting T_name of  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1630  | 
NONE => []  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1631  | 
| SOME raw_sets0 =>  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1632  | 
let  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1633  | 
val (Xs_Ts, (Ts, raw_sets)) =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1634  | 
filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0))  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1635  | 
|> split_list ||> split_list;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1636  | 
val sets = map (mk_set Ts0) raw_sets;  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1637  | 
val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts;  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1638  | 
val xysets = map (pair x) (ys ~~ sets);  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1639  | 
val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts;  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1640  | 
in  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1641  | 
flat (map2 (map o apfst o cons) xysets ppremss)  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1642  | 
end)  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1643  | 
| mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1644  | 
[([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))]  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1645  | 
| mk_induct_raw_prem_prems _ _ _ _ _ = [];  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1646  | 
|
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1647  | 
fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts =  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1648  | 
let  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1649  | 
val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts;  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1650  | 
val pprems =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1651  | 
flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts);  | 
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1652  | 
val y = Term.list_comb (ctr, map alter_x xs);  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1653  | 
val p' = enforce_type names_ctxt domain_type (fastype_of y) p;  | 
| 64607 | 1654  | 
in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end;  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1655  | 
|
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1656  | 
fun close_induct_prem_prem nn ps xs t =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1657  | 
fold_rev Logic.all (map Free (drop (nn + length xs)  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1658  | 
(rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1659  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1660  | 
fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) =  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1661  | 
let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in  | 
| 64607 | 1662  | 
close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) =>  | 
1663  | 
mk_Trueprop_mem (y, set $ x')) xysets,  | 
|
1664  | 
HOLogic.mk_Trueprop (p' $ x)))  | 
|
1665  | 
end;  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1666  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1667  | 
fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) =  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1668  | 
fold_rev Logic.all xs (Logic.list_implies  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1669  | 
(map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl));  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1670  | 
|
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1671  | 
fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1672  | 
let  | 
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1673  | 
fun build_the_rel T Xs_T =  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1674  | 
build_rel [] ctxt [] [] (fn (T, X) =>  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1675  | 
nth rs' (find_index (fn Xs => member (op =) Xs X) Xss)  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1676  | 
|> enforce_type ctxt domain_type T)  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1677  | 
(T, Xs_T)  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1678  | 
|> Term.subst_atomic_types (flat Xss ~~ flat fpTss);  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1679  | 
fun build_rel_app usel vsel Xs_T =  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1680  | 
fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T);  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1681  | 
in  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1682  | 
(if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1683  | 
(if null usels then  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1684  | 
[]  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1685  | 
else  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1686  | 
[Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],  | 
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1687  | 
          Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))])
 | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1688  | 
end;  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1689  | 
|
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1690  | 
fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss =  | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1691  | 
  @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n)
 | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1692  | 
(1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1693  | 
|> flat |> Library.foldr1 HOLogic.mk_conj  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1694  | 
  handle List.Empty => @{term True};
 | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1695  | 
|
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1696  | 
fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1697  | 
fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,  | 
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
1698  | 
HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1699  | 
ctrXs_Tss)));  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1700  | 
|
| 
62158
 
c25c62055180
generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
 
blanchet 
parents: 
62093 
diff
changeset
 | 
1701  | 
fun postproc_co_induct ctxt nn prop prop_conj =  | 
| 57471 | 1702  | 
Drule.zero_var_indexes  | 
1703  | 
#> `(conj_dests nn)  | 
|
| 58472 | 1704  | 
#>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop))  | 
1705  | 
##> (fn thm => Thm.permute_prems 0 (~ nn)  | 
|
| 57471 | 1706  | 
(if nn = 1 then thm RS prop  | 
| 
62158
 
c25c62055180
generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
 
blanchet 
parents: 
62093 
diff
changeset
 | 
1707  | 
     else funpow nn (fn thm => unfold_thms ctxt @{thms conj_assoc} (thm RS prop_conj)) thm));
 | 
| 57471 | 1708  | 
|
1709  | 
fun mk_induct_attrs ctrss =  | 
|
| 63019 | 1710  | 
let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss);  | 
1711  | 
in [Attrib.case_names induct_cases] end;  | 
|
| 57471 | 1712  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1713  | 
fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct  | 
| 57535 | 1714  | 
ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs =  | 
| 57471 | 1715  | 
let  | 
| 58446 | 1716  | 
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);  | 
| 63839 | 1717  | 
val B_ify = Term.map_types B_ify_T;  | 
| 58446 | 1718  | 
|
1719  | 
val fpB_Ts = map B_ify_T fpA_Ts;  | 
|
1720  | 
val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss;  | 
|
1721  | 
val ctrBss = map (map B_ify) ctrAss;  | 
|
| 57471 | 1722  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1723  | 
val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt  | 
| 57471 | 1724  | 
|> mk_Frees "R" (map2 mk_pred2T As Bs)  | 
1725  | 
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts)  | 
|
1726  | 
||>> mk_Freesss "a" ctrAs_Tsss  | 
|
1727  | 
||>> mk_Freesss "b" ctrBs_Tsss;  | 
|
1728  | 
||
| 58291 | 1729  | 
val prems =  | 
| 57471 | 1730  | 
let  | 
1731  | 
fun mk_prem ctrA ctrB argAs argBs =  | 
|
1732  | 
fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies)  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1733  | 
(map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs)  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1734  | 
(HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts  | 
| 57565 | 1735  | 
(Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs)))));  | 
| 57471 | 1736  | 
in  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1737  | 
        flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss)
 | 
| 57471 | 1738  | 
end;  | 
1739  | 
||
| 57525 | 1740  | 
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1741  | 
(map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs));  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1742  | 
val vars = Variable.add_free_names ctxt goal [];  | 
| 57471 | 1743  | 
|
| 
57898
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
1744  | 
val rel_induct0_thm =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1745  | 
      Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} =>
 | 
| 61760 | 1746  | 
mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts  | 
1747  | 
ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs)  | 
|
| 57471 | 1748  | 
|> Thm.close_derivation;  | 
1749  | 
in  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1750  | 
    (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm,
 | 
| 57471 | 1751  | 
mk_induct_attrs ctrAss)  | 
1752  | 
end;  | 
|
1753  | 
||
| 
58335
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
1754  | 
fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms  | 
| 64624 | 1755  | 
live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1756  | 
abs_inverses ctrss ctr_defss recs rec_defs ctxt =  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1757  | 
let  | 
| 51827 | 1758  | 
val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss;  | 
1759  | 
||
| 51815 | 1760  | 
val nn = length pre_bnfs;  | 
| 51827 | 1761  | 
val ns = map length ctr_Tsss;  | 
1762  | 
val mss = map (map length) ctr_Tsss;  | 
|
| 51815 | 1763  | 
|
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1764  | 
val pre_map_defs = map map_def_of_bnf pre_bnfs;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1765  | 
val pre_set_defss = map set_defs_of_bnf pre_bnfs;  | 
| 57399 | 1766  | 
val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;  | 
1767  | 
val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs;  | 
|
| 57397 | 1768  | 
val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1769  | 
|
| 51816 | 1770  | 
val fp_b_names = map base_name_of_typ fpTs;  | 
| 51811 | 1771  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1772  | 
val (((ps, xsss), us'), names_ctxt) = ctxt  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1773  | 
|> mk_Frees "P" (map mk_pred1T fpTs)  | 
| 51827 | 1774  | 
||>> mk_Freesss "x" ctr_Tsss  | 
| 51816 | 1775  | 
||>> Variable.variant_fixes fp_b_names;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1776  | 
|
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1777  | 
val us = map2 (curry Free) us' fpTs;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1778  | 
|
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1779  | 
val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1780  | 
|
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1781  | 
val (induct_thms, induct_thm) =  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1782  | 
let  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1783  | 
        val raw_premss = @{map 4} (@{map 3}
 | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1784  | 
o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting)  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1785  | 
ps ctrss ctr_Tsss ctrXs_Tsss;  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1786  | 
val concl =  | 
| 
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1787  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us));  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1788  | 
val goal =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1789  | 
Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps)))  | 
| 
64576
 
ce8802dc3145
refactored induction principle generation code, for reuse for nonuniform datatypes
 
blanchet 
parents: 
64416 
diff
changeset
 | 
1790  | 
(raw_premss, concl);  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1791  | 
val vars = Variable.add_free_names ctxt goal [];  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1792  | 
val kksss = map (map (map (fst o snd) o #2)) raw_premss;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1793  | 
|
| 64607 | 1794  | 
val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss);  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1795  | 
|
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1796  | 
val thm =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1797  | 
          Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} =>
 | 
| 64624 | 1798  | 
mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses  | 
| 57397 | 1799  | 
abs_inverses fp_nesting_set_maps pre_set_defss)  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1800  | 
|> Thm.close_derivation;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1801  | 
in  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1802  | 
`(conj_dests nn) thm  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1803  | 
end;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1804  | 
|
| 52305 | 1805  | 
val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;  | 
1806  | 
||
| 55867 | 1807  | 
fun mk_rec_thmss (_, x_Tssss, fss, _) recs rec_defs ctor_rec_thms =  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1808  | 
let  | 
| 55867 | 1809  | 
val frecs = map (lists_bmoc fss) recs;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1810  | 
|
| 55871 | 1811  | 
fun mk_goal frec xctr f xs fxs =  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1812  | 
fold_rev (fold_rev Logic.all) (xs :: fss)  | 
| 55867 | 1813  | 
(mk_Trueprop_eq (frec $ xctr, Term.list_comb (f, fxs)));  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1814  | 
|
| 52302 | 1815  | 
fun maybe_tick (T, U) u f =  | 
1816  | 
if try (fst o HOLogic.dest_prodT) U = SOME T then  | 
|
1817  | 
Term.lambda u (HOLogic.mk_prod (u, f $ u))  | 
|
1818  | 
else  | 
|
1819  | 
f;  | 
|
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1820  | 
|
| 55867 | 1821  | 
fun build_rec (x as Free (_, T)) U =  | 
| 52368 | 1822  | 
if T = U then  | 
1823  | 
x  | 
|
1824  | 
else  | 
|
| 
58352
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1825  | 
let  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1826  | 
val build_simple =  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1827  | 
indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1828  | 
(fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk));  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1829  | 
in  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1830  | 
build_map ctxt [] [] build_simple (T, U) $ x  | 
| 
58352
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
1831  | 
end;  | 
| 
52301
 
7935e82a4ae4
simpler, more robust iterator goal construction code
 
blanchet 
parents: 
52300 
diff
changeset
 | 
1832  | 
|
| 55867 | 1833  | 
val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss;  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1834  | 
        val goalss = @{map 5} (@{map 4} o mk_goal) frecs xctrss fss xsss fxsss;
 | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1835  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1836  | 
        val tacss = @{map 4} (map ooo
 | 
| 57399 | 1837  | 
mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs)  | 
| 64624 | 1838  | 
ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss;  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1839  | 
|
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1840  | 
fun prove goal tac =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1841  | 
Goal.prove_sorry ctxt [] [] goal (tac o #context)  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1842  | 
|> Thm.close_derivation;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1843  | 
in  | 
| 52305 | 1844  | 
map2 (map2 prove) goalss tacss  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1845  | 
end;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1846  | 
|
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
1847  | 
val rec_thmss = mk_rec_thmss (the rec_args_typess) recs rec_defs ctor_rec_thms;  | 
| 
58335
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
1848  | 
|
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1849  | 
in  | 
| 57471 | 1850  | 
((induct_thms, induct_thm, mk_induct_attrs ctrss),  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
1851  | 
(rec_thmss, nitpicksimp_attrs @ simp_attrs))  | 
| 
51808
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1852  | 
end;  | 
| 
 
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
 
blanchet 
parents: 
51805 
diff
changeset
 | 
1853  | 
|
| 
58283
 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 
blanchet 
parents: 
58267 
diff
changeset
 | 
1854  | 
fun mk_coinduct_attrs fpTs ctrss discss mss =  | 
| 57302 | 1855  | 
let  | 
1856  | 
val fp_b_names = map base_name_of_typ fpTs;  | 
|
| 57535 | 1857  | 
|
| 57302 | 1858  | 
fun mk_coinduct_concls ms discs ctrs =  | 
1859  | 
let  | 
|
1860  | 
fun mk_disc_concl disc = [name_of_disc disc];  | 
|
1861  | 
fun mk_ctr_concl 0 _ = []  | 
|
| 
58284
 
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
 
blanchet 
parents: 
58283 
diff
changeset
 | 
1862  | 
| mk_ctr_concl _ ctr = [name_of_ctr ctr];  | 
| 57302 | 1863  | 
val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]];  | 
1864  | 
val ctr_concls = map2 mk_ctr_concl ms ctrs;  | 
|
1865  | 
in  | 
|
1866  | 
flat (map2 append disc_concls ctr_concls)  | 
|
1867  | 
end;  | 
|
| 57535 | 1868  | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
1869  | 
val coinduct_cases = quasi_unambiguous_case_names (map (prefix Eq_prefix) fp_b_names);  | 
| 57302 | 1870  | 
val coinduct_conclss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1871  | 
      @{map 3} (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss;
 | 
| 57302 | 1872  | 
|
| 63019 | 1873  | 
val coinduct_case_names_attr = Attrib.case_names coinduct_cases;  | 
| 57302 | 1874  | 
val coinduct_case_concl_attrs =  | 
| 63019 | 1875  | 
map2 (fn casex => fn concls => Attrib.case_conclusion (casex, concls))  | 
| 57302 | 1876  | 
coinduct_cases coinduct_conclss;  | 
1877  | 
||
| 
62158
 
c25c62055180
generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
 
blanchet 
parents: 
62093 
diff
changeset
 | 
1878  | 
val common_coinduct_attrs = coinduct_case_names_attr :: coinduct_case_concl_attrs;  | 
| 63019 | 1879  | 
val coinduct_attrs = Attrib.consumes 1 :: coinduct_case_names_attr :: coinduct_case_concl_attrs;  | 
| 57302 | 1880  | 
in  | 
1881  | 
(coinduct_attrs, common_coinduct_attrs)  | 
|
1882  | 
end;  | 
|
1883  | 
||
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1884  | 
fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list)  | 
| 57535 | 1885  | 
abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct  | 
1886  | 
live_nesting_rel_eqs =  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1887  | 
let  | 
| 58446 | 1888  | 
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);  | 
1889  | 
val fpB_Ts = map B_ify_T fpA_Ts;  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1890  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
1891  | 
val (Rs, IRs, fpAs, fpBs, _) =  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1892  | 
let  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1893  | 
val fp_names = map base_name_of_typ fpA_Ts;  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1894  | 
val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1895  | 
|> mk_Frees "R" (map2 mk_pred2T As Bs)  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1896  | 
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts)  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1897  | 
||>> Variable.variant_fixes fp_names  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1898  | 
||>> Variable.variant_fixes (map (suffix "'") fp_names);  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1899  | 
in  | 
| 58291 | 1900  | 
(Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts,  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1901  | 
names_ctxt)  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1902  | 
end;  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1903  | 
|
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1904  | 
val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) =  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1905  | 
let  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1906  | 
val discss = map #discs ctr_sugars;  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1907  | 
val selsss = map #selss ctr_sugars;  | 
| 58291 | 1908  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1909  | 
fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss);  | 
| 58291 | 1910  | 
fun mk_selsss ts Ts =  | 
1911  | 
map2 (map o map o rapp) ts (map (map (map (mk_disc_or_sel Ts))) selsss);  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1912  | 
in  | 
| 57302 | 1913  | 
((mk_discss fpAs As, mk_selsss fpAs As),  | 
1914  | 
(mk_discss fpBs Bs, mk_selsss fpBs Bs))  | 
|
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1915  | 
end;  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1916  | 
|
| 58291 | 1917  | 
val prems =  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1918  | 
let  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1919  | 
fun mk_prem_ctr_concls n k discA_t selA_ts discB_t selB_ts =  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1920  | 
(if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1921  | 
(case (selA_ts, selB_ts) of  | 
| 57303 | 1922  | 
([], []) => []  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1923  | 
| (_ :: _, _ :: _) =>  | 
| 57565 | 1924  | 
[Library.foldr HOLogic.mk_imp  | 
1925  | 
(if n = 1 then [] else [discA_t, discB_t],  | 
|
1926  | 
Library.foldr1 HOLogic.mk_conj  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1927  | 
(map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]);  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1928  | 
|
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1929  | 
fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1930  | 
          Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n)
 | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1931  | 
(1 upto n) discA_ts selA_tss discB_ts selB_tss))  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1932  | 
          handle List.Empty => @{term True};
 | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1933  | 
|
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1934  | 
fun mk_prem IR tA tB n discA_ts selA_tss discB_ts selB_tss =  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1935  | 
fold_rev Logic.all [tA, tB] (Logic.mk_implies (HOLogic.mk_Trueprop (IR $ tA $ tB),  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1936  | 
HOLogic.mk_Trueprop (mk_prem_concl n discA_ts selA_tss discB_ts selB_tss)));  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1937  | 
in  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
1938  | 
        @{map 8} mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss
 | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1939  | 
end;  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1940  | 
|
| 57525 | 1941  | 
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1942  | 
IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts)));  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1943  | 
val vars = Variable.add_free_names ctxt goal [];  | 
| 57302 | 1944  | 
|
| 
57898
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
1945  | 
val rel_coinduct0_thm =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1946  | 
      Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} =>
 | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
1947  | 
mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems  | 
| 
57898
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
1948  | 
(map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars)  | 
| 
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
1949  | 
(map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects  | 
| 
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
1950  | 
rel_pre_defs abs_inverses live_nesting_rel_eqs)  | 
| 57302 | 1951  | 
|> Thm.close_derivation;  | 
1952  | 
in  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1953  | 
    (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm,
 | 
| 
58283
 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 
blanchet 
parents: 
58267 
diff
changeset
 | 
1954  | 
mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss)  | 
| 
57301
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1955  | 
end;  | 
| 
 
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
 
desharna 
parents: 
57279 
diff
changeset
 | 
1956  | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
1957  | 
fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs  | 
| 63814 | 1958  | 
ctor_defs dtor_ctors Abs_pre_inverses =  | 
| 57700 | 1959  | 
let  | 
1960  | 
fun mk_prems A Ps ctr_args t ctxt =  | 
|
1961  | 
(case fastype_of t of  | 
|
| 57891 | 1962  | 
Type (type_name, innerTs) =>  | 
| 57700 | 1963  | 
(case bnf_of ctxt type_name of  | 
1964  | 
NONE => ([], ctxt)  | 
|
1965  | 
| SOME bnf =>  | 
|
1966  | 
let  | 
|
1967  | 
fun seq_assm a set ctxt =  | 
|
1968  | 
let  | 
|
1969  | 
val X = HOLogic.dest_setT (range_type (fastype_of set));  | 
|
1970  | 
val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt;  | 
|
1971  | 
val assm = mk_Trueprop_mem (x, set $ a);  | 
|
1972  | 
in  | 
|
1973  | 
(case build_binary_fun_app Ps x a of  | 
|
1974  | 
NONE =>  | 
|
1975  | 
mk_prems A Ps ctr_args x ctxt'  | 
|
1976  | 
|>> map (Logic.all x o Logic.mk_implies o pair assm)  | 
|
1977  | 
| SOME f =>  | 
|
1978  | 
([Logic.all x  | 
|
1979  | 
(Logic.mk_implies (assm,  | 
|
1980  | 
Logic.mk_implies (HOLogic.mk_Trueprop f,  | 
|
1981  | 
HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))],  | 
|
1982  | 
ctxt'))  | 
|
1983  | 
end;  | 
|
1984  | 
in  | 
|
| 57891 | 1985  | 
fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt  | 
| 57700 | 1986  | 
|>> flat  | 
1987  | 
end)  | 
|
1988  | 
| T =>  | 
|
1989  | 
if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt)  | 
|
1990  | 
else ([], ctxt));  | 
|
1991  | 
||
1992  | 
fun mk_prems_for_ctr A Ps ctr ctxt =  | 
|
1993  | 
let  | 
|
| 57893 | 1994  | 
val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt;  | 
| 57700 | 1995  | 
in  | 
1996  | 
fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt'  | 
|
1997  | 
|>> map (fold_rev Logic.all args) o flat  | 
|
1998  | 
|>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr)))  | 
|
1999  | 
end;  | 
|
2000  | 
||
2001  | 
fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt =  | 
|
2002  | 
let  | 
|
2003  | 
val ((x, fp), ctxt') = ctxt  | 
|
2004  | 
|> yield_singleton (mk_Frees "x") A  | 
|
2005  | 
||>> yield_singleton (mk_Frees "a") fpT;  | 
|
2006  | 
val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x)  | 
|
2007  | 
(the (build_binary_fun_app Ps x fp)));  | 
|
2008  | 
in  | 
|
2009  | 
fold_map (mk_prems_for_ctr A Ps) ctrs ctxt'  | 
|
2010  | 
|>> split_list  | 
|
2011  | 
|>> map_prod flat flat  | 
|
2012  | 
|>> apfst (rpair concl)  | 
|
2013  | 
end;  | 
|
2014  | 
||
2015  | 
fun mk_thm ctxt fpTs ctrss sets =  | 
|
2016  | 
let  | 
|
2017  | 
val A = HOLogic.dest_setT (range_type (fastype_of (hd sets)));  | 
|
2018  | 
val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt;  | 
|
| 58291 | 2019  | 
val (((prems, concl), case_names), ctxt'') =  | 
2020  | 
fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt'  | 
|
| 57700 | 2021  | 
|>> apfst split_list o split_list  | 
2022  | 
|>> apfst (apfst flat)  | 
|
2023  | 
|>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj))  | 
|
2024  | 
|>> apsnd flat;  | 
|
2025  | 
||
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2026  | 
val vars = fold (Variable.add_free_names ctxt) (concl :: prems) [];  | 
| 
57898
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
2027  | 
val thm =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2028  | 
Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl)  | 
| 
57898
 
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
 
blanchet 
parents: 
57893 
diff
changeset
 | 
2029  | 
            (fn {context = ctxt, prems} =>
 | 
| 61760 | 2030  | 
mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts  | 
2031  | 
exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses)  | 
|
| 57700 | 2032  | 
|> Thm.close_derivation;  | 
2033  | 
||
| 63019 | 2034  | 
val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names);  | 
| 57700 | 2035  | 
val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets;  | 
2036  | 
in  | 
|
2037  | 
(thm, case_names_attr :: induct_set_attrs)  | 
|
2038  | 
end  | 
|
| 63019 | 2039  | 
val consumes_attr = Attrib.consumes 1;  | 
| 57700 | 2040  | 
in  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2041  | 
map (mk_thm ctxt fpTs ctrss  | 
| 58472 | 2042  | 
#> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr))  | 
2043  | 
(transpose setss)  | 
|
| 57700 | 2044  | 
end;  | 
2045  | 
||
| 60737 | 2046  | 
fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt =  | 
2047  | 
let  | 
|
2048  | 
val n = Thm.nprems_of coind;  | 
|
2049  | 
val m = Thm.nprems_of (hd rel_monos) - n;  | 
|
2050  | 
fun mk_inst phi =  | 
|
2051  | 
(phi, Thm.cterm_of ctxt (mk_union (Var phi, HOLogic.eq_const (fst (dest_pred2T (#2 phi))))));  | 
|
2052  | 
val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map mk_inst;  | 
|
2053  | 
fun mk_unfold rel_eq rel_mono =  | 
|
2054  | 
let  | 
|
2055  | 
        val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl];
 | 
|
2056  | 
        val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset});
 | 
|
2057  | 
      in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS eqTrueI end;
 | 
|
2058  | 
    val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def
 | 
|
2059  | 
imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)};  | 
|
2060  | 
in  | 
|
2061  | 
Thm.instantiate ([], insts) coind  | 
|
2062  | 
|> unfold_thms ctxt unfolds  | 
|
2063  | 
end;  | 
|
2064  | 
||
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2065  | 
fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors  | 
| 64624 | 2066  | 
live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2067  | 
(ctr_sugars : ctr_sugar list) =  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2068  | 
let  | 
| 51815 | 2069  | 
val nn = length pre_bnfs;  | 
2070  | 
||
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2071  | 
val pre_rel_defs = map rel_def_of_bnf pre_bnfs;  | 
| 57397 | 2072  | 
val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2073  | 
|
| 51816 | 2074  | 
val fp_b_names = map base_name_of_typ fpTs;  | 
| 51811 | 2075  | 
|
| 
53210
 
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
 
blanchet 
parents: 
53203 
diff
changeset
 | 
2076  | 
val discss = map #discs ctr_sugars;  | 
| 
 
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
 
blanchet 
parents: 
53203 
diff
changeset
 | 
2077  | 
val selsss = map #selss ctr_sugars;  | 
| 51840 | 2078  | 
val exhausts = map #exhaust ctr_sugars;  | 
2079  | 
val disc_thmsss = map #disc_thmss ctr_sugars;  | 
|
2080  | 
val sel_thmsss = map #sel_thmss ctr_sugars;  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2081  | 
|
| 62395 | 2082  | 
val (((rs, us'), vs'), _) = ctxt  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2083  | 
|> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs)  | 
| 51816 | 2084  | 
||>> Variable.variant_fixes fp_b_names  | 
2085  | 
||>> Variable.variant_fixes (map (suffix "'") fp_b_names);  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2086  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2087  | 
val us = map2 (curry Free) us' fpTs;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2088  | 
val udiscss = map2 (map o rapp) us discss;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2089  | 
val uselsss = map2 (map o map o rapp) us selsss;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2090  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2091  | 
val vs = map2 (curry Free) vs' fpTs;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2092  | 
val vdiscss = map2 (map o rapp) vs discss;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2093  | 
val vselsss = map2 (map o map o rapp) vs selsss;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2094  | 
|
| 62395 | 2095  | 
    val uvrs = @{map 3} (fn r => fn u => fn v => r $ u $ v) rs us vs;
 | 
2096  | 
val uv_eqs = map2 (curry HOLogic.mk_eq) us vs;  | 
|
2097  | 
val strong_rs =  | 
|
2098  | 
      @{map 4} (fn u => fn v => fn uvr => fn uv_eq =>
 | 
|
2099  | 
fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;  | 
|
2100  | 
||
2101  | 
val concl =  | 
|
2102  | 
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
|
2103  | 
        (@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v)))
 | 
|
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2104  | 
uvrs us vs))  | 
| 62395 | 2105  | 
|
2106  | 
fun mk_goal rs0' =  | 
|
| 
64638
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
2107  | 
      Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs)
 | 
| 
 
235df39ade87
generalized generation of coinduction goal (towards nonuniform codatatypes)
 
blanchet 
parents: 
64637 
diff
changeset
 | 
2108  | 
(map alter_r rs0'))  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2109  | 
uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss,  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2110  | 
concl);  | 
| 62395 | 2111  | 
|
2112  | 
val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else []));  | 
|
2113  | 
||
2114  | 
fun prove dtor_coinduct' goal =  | 
|
2115  | 
Variable.add_free_names ctxt goal []  | 
|
2116  | 
      |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} =>
 | 
|
| 64624 | 2117  | 
mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses  | 
| 62395 | 2118  | 
abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss))  | 
2119  | 
|> Thm.close_derivation;  | 
|
2120  | 
||
2121  | 
val rel_eqs = map rel_eq_of_bnf pre_bnfs;  | 
|
2122  | 
val rel_monos = map rel_mono_of_bnf pre_bnfs;  | 
|
2123  | 
val dtor_coinducts =  | 
|
2124  | 
[dtor_coinduct] @  | 
|
2125  | 
(if strong then [mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p ctxt]  | 
|
2126  | 
else []);  | 
|
2127  | 
in  | 
|
2128  | 
    map2 (postproc_co_induct ctxt nn mp @{thm conj_commute[THEN iffD1]} oo prove)
 | 
|
2129  | 
dtor_coinducts goals  | 
|
2130  | 
end;  | 
|
2131  | 
||
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2132  | 
fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2133  | 
(x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2134  | 
dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses  | 
| 
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2135  | 
mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs =  | 
| 62395 | 2136  | 
let  | 
2137  | 
fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec =  | 
|
2138  | 
iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]];  | 
|
2139  | 
||
2140  | 
val ctor_dtor_corec_thms =  | 
|
2141  | 
      @{map 3} mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms;
 | 
|
2142  | 
||
2143  | 
val pre_map_defs = map map_def_of_bnf pre_bnfs;  | 
|
2144  | 
val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;  | 
|
2145  | 
||
2146  | 
val fp_b_names = map base_name_of_typ fpTs;  | 
|
2147  | 
||
2148  | 
val ctrss = map #ctrs ctr_sugars;  | 
|
2149  | 
val discss = map #discs ctr_sugars;  | 
|
2150  | 
val selsss = map #selss ctr_sugars;  | 
|
2151  | 
val disc_thmsss = map #disc_thmss ctr_sugars;  | 
|
2152  | 
val discIss = map #discIs ctr_sugars;  | 
|
2153  | 
val sel_thmsss = map #sel_thmss ctr_sugars;  | 
|
2154  | 
||
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2155  | 
val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct  | 
| 64624 | 2156  | 
dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2157  | 
ctr_defss ctr_sugars;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2158  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2159  | 
fun mk_maybe_not pos = not pos ? HOLogic.mk_not;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2160  | 
|
| 55869 | 2161  | 
val gcorecs = map (lists_bmoc pgss) corecs;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2162  | 
|
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
2163  | 
val corec_thmss =  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2164  | 
let  | 
| 62395 | 2165  | 
val (us', _) = ctxt  | 
2166  | 
|> Variable.variant_fixes fp_b_names;  | 
|
2167  | 
||
2168  | 
val us = map2 (curry Free) us' fpTs;  | 
|
2169  | 
||
| 55869 | 2170  | 
fun mk_goal c cps gcorec n k ctr m cfs' =  | 
2171  | 
fold_rev (fold_rev Logic.all) ([c] :: pgss)  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2172  | 
(Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,  | 
| 55869 | 2173  | 
mk_Trueprop_eq (gcorec $ c, Term.list_comb (ctr, take m cfs'))));  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2174  | 
|
| 55869 | 2175  | 
val mk_U = typ_subst_nonatomic (map2 (fn C => fn fpT => (mk_sumT (fpT, C), fpT)) Cs fpTs);  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
2176  | 
|
| 55869 | 2177  | 
fun tack (c, u) f =  | 
2178  | 
let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in  | 
|
2179  | 
Term.lambda x' (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ x')  | 
|
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
2180  | 
end;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2181  | 
|
| 55869 | 2182  | 
fun build_corec cqg =  | 
2183  | 
let val T = fastype_of cqg in  | 
|
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
2184  | 
if exists_subtype_in Cs T then  | 
| 
58352
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2185  | 
let  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2186  | 
val U = mk_U T;  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2187  | 
val build_simple =  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2188  | 
indexify fst (map2 (curry mk_sumT) fpTs Cs)  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2189  | 
(fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk));  | 
| 
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2190  | 
in  | 
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2191  | 
build_map ctxt [] [] build_simple (T, U) $ cqg  | 
| 52368 | 2192  | 
end  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
2193  | 
else  | 
| 55869 | 2194  | 
cqg  | 
| 
52347
 
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
 
blanchet 
parents: 
52346 
diff
changeset
 | 
2195  | 
end;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2196  | 
|
| 55869 | 2197  | 
val cqgsss' = map (map (map build_corec)) cqgsss;  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2198  | 
        val goalss = @{map 8} (@{map 4} oooo mk_goal) cs cpss gcorecs ns kss ctrss mss cqgsss';
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2199  | 
|
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
2200  | 
val tacss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2201  | 
          @{map 4} (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s)
 | 
| 55867 | 2202  | 
ctor_dtor_corec_thms pre_map_defs abs_inverses ctr_defss;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2203  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2204  | 
fun prove goal tac =  | 
| 62395 | 2205  | 
Goal.prove_sorry ctxt [] [] goal (tac o #context)  | 
| 51815 | 2206  | 
|> Thm.close_derivation;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2207  | 
in  | 
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
2208  | 
map2 (map2 prove) goalss tacss  | 
| 62395 | 2209  | 
        |> map (map (unfold_thms ctxt @{thms case_sum_if}))
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2210  | 
end;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2211  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2212  | 
val corec_disc_iff_thmss =  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2213  | 
let  | 
| 55869 | 2214  | 
fun mk_goal c cps gcorec n k disc =  | 
2215  | 
mk_Trueprop_eq (disc $ (gcorec $ c),  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2216  | 
            if n = 1 then @{const True}
 | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2217  | 
else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2218  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2219  | 
        val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss;
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2220  | 
|
| 62395 | 2221  | 
        fun mk_case_split' cp = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt cp)] @{thm case_split};
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2222  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2223  | 
val case_splitss' = map (map mk_case_split') cpss;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2224  | 
|
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2225  | 
        val tacss = @{map 3} (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss;
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2226  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2227  | 
fun prove goal tac =  | 
| 62395 | 2228  | 
Variable.add_free_names ctxt goal []  | 
2229  | 
|> (fn vars => Goal.prove_sorry ctxt vars [] goal (tac o #context))  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2230  | 
|> Thm.close_derivation;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2231  | 
|
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2232  | 
fun proves [_] [_] = []  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2233  | 
| proves goals tacs = map2 prove goals tacs;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2234  | 
in  | 
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
2235  | 
map2 proves goalss tacss  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2236  | 
end;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2237  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2238  | 
fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs);  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2239  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2240  | 
val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2241  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2242  | 
fun mk_corec_sel_thm corec_thm sel sel_thm =  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2243  | 
let  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2244  | 
val (domT, ranT) = dest_funT (fastype_of sel);  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2245  | 
val arg_cong' =  | 
| 62395 | 2246  | 
Thm.instantiate' (map (SOME o Thm.ctyp_of ctxt) [domT, ranT])  | 
2247  | 
[NONE, NONE, SOME (Thm.cterm_of ctxt sel)] arg_cong  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2248  | 
|> Thm.varifyT_global;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2249  | 
val sel_thm' = sel_thm RSN (2, trans);  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2250  | 
in  | 
| 55867 | 2251  | 
corec_thm RS arg_cong' RS sel_thm'  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2252  | 
end;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2253  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2254  | 
fun mk_corec_sel_thms corec_thmss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2255  | 
      @{map 3} (@{map 3} (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss;
 | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2256  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2257  | 
val corec_sel_thmsss = mk_corec_sel_thms corec_thmss;  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2258  | 
in  | 
| 57535 | 2259  | 
((coinduct_thms_pairs,  | 
| 
58283
 
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
 
blanchet 
parents: 
58267 
diff
changeset
 | 
2260  | 
mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss),  | 
| 57489 | 2261  | 
corec_thmss,  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2262  | 
corec_disc_thmss,  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2263  | 
(corec_disc_iff_thmss, simp_attrs),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2264  | 
(corec_sel_thmsss, simp_attrs))  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2265  | 
end;  | 
| 
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2266  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58634 
diff
changeset
 | 
2267  | 
fun define_co_datatypes prepare_plugins prepare_constraint prepare_typ prepare_term fp construct_fp  | 
| 63818 | 2268  | 
((raw_plugins, discs_sels0), specs) lthy =  | 
| 49112 | 2269  | 
let  | 
| 63818 | 2270  | 
val plugins = prepare_plugins lthy raw_plugins;  | 
| 
57094
 
589ec121ce1a
don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
 
blanchet 
parents: 
57091 
diff
changeset
 | 
2271  | 
val discs_sels = discs_sels0 orelse fp = Greatest_FP;  | 
| 49278 | 2272  | 
|
| 49367 | 2273  | 
val nn = length specs;  | 
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2274  | 
val fp_bs = map type_binding_of_spec specs;  | 
| 49498 | 2275  | 
val fp_b_names = map Binding.name_of fp_bs;  | 
2276  | 
val fp_common_name = mk_common_name fp_b_names;  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2277  | 
val map_bs = map map_binding_of_spec specs;  | 
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2278  | 
val rel_bs = map rel_binding_of_spec specs;  | 
| 62324 | 2279  | 
val pred_bs = map pred_binding_of_spec specs;  | 
| 
49298
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2280  | 
|
| 51758 | 2281  | 
fun prepare_type_arg (_, (ty, c)) =  | 
| 63818 | 2282  | 
let val TFree (s, _) = prepare_typ lthy ty in  | 
2283  | 
TFree (s, prepare_constraint lthy c)  | 
|
| 
49298
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2284  | 
end;  | 
| 
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2285  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2286  | 
val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs;  | 
| 58234 | 2287  | 
    val unsorted_Ass0 = map (map (resort_tfree_or_tvar @{sort type})) Ass0;
 | 
| 55700 | 2288  | 
val unsorted_As = Library.foldr1 (merge_type_args fp) unsorted_Ass0;  | 
| 
53266
 
89f7f1cc9ae3
cleaner handling of bootstrapping "fake" context, with fewer (no?) obscure bugs
 
blanchet 
parents: 
53264 
diff
changeset
 | 
2289  | 
val num_As = length unsorted_As;  | 
| 
55699
 
1f9cc432ecd4
reuse same parser for '(co)datatype(_new)' as for 'bnf_decl'
 
blanchet 
parents: 
55575 
diff
changeset
 | 
2290  | 
|
| 
 
1f9cc432ecd4
reuse same parser for '(co)datatype(_new)' as for 'bnf_decl'
 
blanchet 
parents: 
55575 
diff
changeset
 | 
2291  | 
val set_boss = map (map fst o type_args_named_constrained_of_spec) specs;  | 
| 
 
1f9cc432ecd4
reuse same parser for '(co)datatype(_new)' as for 'bnf_decl'
 
blanchet 
parents: 
55575 
diff
changeset
 | 
2292  | 
val set_bss = map (map (the_default Binding.empty)) set_boss;  | 
| 49119 | 2293  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2294  | 
fun add_fake_type spec =  | 
| 61259 | 2295  | 
      Typedecl.basic_typedecl {final = true}
 | 
| 
63048
 
1836456b7d82
avoid duplicate mixfix messages in '(co)datatype' type name
 
blanchet 
parents: 
63019 
diff
changeset
 | 
2296  | 
(type_binding_of_spec spec, num_As, Mixfix.reset_pos (mixfix_of_spec spec));  | 
| 51768 | 2297  | 
|
| 63818 | 2298  | 
val (fake_T_names, fake_lthy) = fold_map add_fake_type specs lthy;  | 
| 49119 | 2299  | 
|
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2300  | 
val qsoty = quote o Syntax.string_of_typ fake_lthy;  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2301  | 
|
| 54253 | 2302  | 
val _ = (case Library.duplicates (op =) unsorted_As of [] => ()  | 
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2303  | 
      | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^
 | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2304  | 
"datatype specification"));  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2305  | 
|
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2306  | 
val bad_args =  | 
| 63818 | 2307  | 
map (Logic.type_map (singleton (Variable.polymorphic lthy))) unsorted_As  | 
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2308  | 
|> filter_out Term.is_TVar;  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2309  | 
val _ = null bad_args orelse  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2310  | 
      error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^
 | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2311  | 
"datatype specification");  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2312  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2313  | 
val mixfixes = map mixfix_of_spec specs;  | 
| 49119 | 2314  | 
|
| 54253 | 2315  | 
val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => ()  | 
| 49119 | 2316  | 
      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
 | 
2317  | 
||
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2318  | 
val mx_ctr_specss = map mixfixed_ctr_specs_of_spec specs;  | 
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2319  | 
val ctr_specss = map (map fst) mx_ctr_specss;  | 
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2320  | 
val ctr_mixfixess = map (map snd) mx_ctr_specss;  | 
| 49119 | 2321  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2322  | 
val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss;  | 
| 49336 | 2323  | 
val ctr_bindingss =  | 
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2324  | 
map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of_ctr_spec)) fp_b_names  | 
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2325  | 
ctr_specss;  | 
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2326  | 
val ctr_argsss = map (map args_of_ctr_spec) ctr_specss;  | 
| 49119 | 2327  | 
|
| 49336 | 2328  | 
val sel_bindingsss = map (map (map fst)) ctr_argsss;  | 
| 
49298
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2329  | 
val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;  | 
| 
57200
 
aab87ffa60cc
use 'where' clause for selector default value syntax
 
blanchet 
parents: 
57199 
diff
changeset
 | 
2330  | 
val raw_sel_default_eqss = map sel_default_eqs_of_spec specs;  | 
| 49286 | 2331  | 
|
| 
49308
 
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
 
blanchet 
parents: 
49302 
diff
changeset
 | 
2332  | 
val (As :: _) :: fake_ctr_Tsss =  | 
| 
49298
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2333  | 
burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);  | 
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2334  | 
val As' = map dest_TFree As;  | 
| 
49298
 
36e551d3af3b
support for sort constraints in new (co)data commands
 
blanchet 
parents: 
49297 
diff
changeset
 | 
2335  | 
|
| 
49183
 
0cc46e2dee7e
careful about constructor types w.r.t. fake context (third step of localization)
 
blanchet 
parents: 
49182 
diff
changeset
 | 
2336  | 
val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];  | 
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2337  | 
val _ = (case subtract (op =) As' rhs_As' of [] => ()  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2338  | 
      | extras => error ("Extra type variables on right-hand side: " ^
 | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2339  | 
commas (map (qsoty o TFree) extras)));  | 
| 49165 | 2340  | 
|
| 
53268
 
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
 
blanchet 
parents: 
53266 
diff
changeset
 | 
2341  | 
val fake_Ts = map (fn s => Type (s, As)) fake_T_names;  | 
| 
 
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
 
blanchet 
parents: 
53266 
diff
changeset
 | 
2342  | 
|
| 63818 | 2343  | 
val ((((Bs0, Cs), Es), Xs), _) = lthy  | 
| 63796 | 2344  | 
|> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As  | 
2345  | 
|> mk_TFrees num_As  | 
|
2346  | 
||>> mk_TFrees nn  | 
|
2347  | 
||>> mk_TFrees nn  | 
|
2348  | 
||>> variant_tfrees fp_b_names;  | 
|
2349  | 
||
| 
53262
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2350  | 
fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) =  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2351  | 
s = s' andalso (Ts = Ts' orelse  | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2352  | 
          error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^
 | 
| 
 
23927b18dce2
rationalize message generation + added a warning
 
blanchet 
parents: 
53260 
diff
changeset
 | 
2353  | 
" (expected " ^ qsoty T' ^ ")"))  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
2354  | 
| eq_fpT_check _ _ = false;  | 
| 49146 | 2355  | 
|
| 53222 | 2356  | 
fun freeze_fp (T as Type (s, Ts)) =  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
2357  | 
(case find_index (eq_fpT_check T) fake_Ts of  | 
| 53222 | 2358  | 
~1 => Type (s, map freeze_fp Ts)  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49590 
diff
changeset
 | 
2359  | 
| kk => nth Xs kk)  | 
| 49204 | 2360  | 
| freeze_fp T = T;  | 
| 49121 | 2361  | 
|
| 53222 | 2362  | 
val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts);  | 
2363  | 
||
| 
52310
 
28063e412793
support induction principles with multiple occurrences of the same type in "fpTs" and (hopefully) with loss of recursion (e.g. primrec definition of is_nil, where the IH can be dropped)
 
blanchet 
parents: 
52309 
diff
changeset
 | 
2364  | 
val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss;  | 
| 55966 | 2365  | 
val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss;  | 
| 49119 | 2366  | 
|
| 58731 | 2367  | 
val _ =  | 
2368  | 
let  | 
|
| 62687 | 2369  | 
fun add_deps i =  | 
2370  | 
fold (fn T => fold_index (fn (j, X) =>  | 
|
| 63851 | 2371  | 
(i <> j andalso exists_subtype_in [X] T) ? insert (op =) (i, j)) Xs);  | 
| 62687 | 2372  | 
|
2373  | 
val add_missing_nodes = fold (AList.default (op =) o rpair []) (0 upto nn - 1);  | 
|
2374  | 
||
2375  | 
val deps = fold_index (uncurry (fold o add_deps)) ctrXs_Tsss []  | 
|
2376  | 
|> AList.group (op =)  | 
|
2377  | 
|> add_missing_nodes;  | 
|
2378  | 
||
2379  | 
val G = Int_Graph.make (map (apfst (rpair ())) deps);  | 
|
2380  | 
val sccs = map (sort int_ord) (Int_Graph.strong_conn G);  | 
|
| 58731 | 2381  | 
|
| 63694 | 2382  | 
val str_of_scc = prefix (co_prefix fp ^ "datatype ") o  | 
| 58731 | 2383  | 
space_implode " and " o map (suffix " = \<dots>" o Long_Name.base_name);  | 
2384  | 
||
2385  | 
fun warn [_] = ()  | 
|
2386  | 
| warn sccs =  | 
|
2387  | 
            warning ("Defined types not fully mutually " ^ co_prefix fp ^ "recursive\n\
 | 
|
2388  | 
\Alternative specification:\n" ^  | 
|
2389  | 
cat_lines (map (prefix " " o str_of_scc o map (nth fp_b_names)) sccs));  | 
|
| 62687 | 2390  | 
in  | 
2391  | 
warn (order_strong_conn (op =) Int_Graph.make Int_Graph.topological_order deps sccs)  | 
|
2392  | 
end;  | 
|
| 58731 | 2393  | 
|
| 55701 | 2394  | 
val killed_As =  | 
2395  | 
map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE)  | 
|
| 63551 | 2396  | 
(As ~~ transpose set_boss);  | 
| 55701 | 2397  | 
|
| 62720 | 2398  | 
    val (((pre_bnfs, absT_infos), _), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0,
 | 
| 63813 | 2399  | 
dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, ctor_dtors,  | 
2400  | 
ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels, xtor_co_rec_thms,  | 
|
2401  | 
xtor_rel_co_induct, dtor_set_inducts, xtor_co_rec_transfers, xtor_co_rec_o_maps, ...},  | 
|
| 
53203
 
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
 
traytel 
parents: 
53143 
diff
changeset
 | 
2402  | 
lthy)) =  | 
| 63802 | 2403  | 
fixpoint_bnf false I (construct_fp mixfixes map_bs rel_bs pred_bs set_bss) fp_bs  | 
2404  | 
(map dest_TFree As) (map dest_TFree killed_As) (map dest_TFree Xs) ctrXs_repTs  | 
|
| 63818 | 2405  | 
empty_comp_cache lthy  | 
| 53222 | 2406  | 
handle BAD_DEAD (X, X_backdrop) =>  | 
2407  | 
(case X_backdrop of  | 
|
2408  | 
Type (bad_tc, _) =>  | 
|
2409  | 
let  | 
|
2410  | 
val fake_T = qsoty (unfreeze_fp X);  | 
|
2411  | 
val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop);  | 
|
2412  | 
fun register_hint () =  | 
|
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59877 
diff
changeset
 | 
2413  | 
              "\nUse the " ^ quote (#1 @{command_keyword bnf}) ^ " command to register " ^
 | 
| 53222 | 2414  | 
quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \  | 
2415  | 
\it";  | 
|
2416  | 
in  | 
|
| 63818 | 2417  | 
if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then  | 
| 53256 | 2418  | 
              error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
 | 
2419  | 
" in type expression " ^ fake_T_backdrop)  | 
|
| 63818 | 2420  | 
else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy)  | 
| 53222 | 2421  | 
bad_tc) then  | 
2422  | 
              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
 | 
|
2423  | 
" via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^  | 
|
2424  | 
fake_T_backdrop ^ register_hint ())  | 
|
2425  | 
else  | 
|
2426  | 
              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
 | 
|
2427  | 
" via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^  | 
|
2428  | 
register_hint ())  | 
|
2429  | 
end);  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2430  | 
|
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53126 
diff
changeset
 | 
2431  | 
val time = time lthy;  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2432  | 
val timer = time (Timer.startRealTimer ());  | 
| 49121 | 2433  | 
|
| 57397 | 2434  | 
val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs;  | 
2435  | 
val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As;  | 
|
| 49226 | 2436  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2437  | 
val pre_map_defs = map map_def_of_bnf pre_bnfs;  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2438  | 
val pre_set_defss = map set_defs_of_bnf pre_bnfs;  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2439  | 
val pre_rel_defs = map rel_def_of_bnf pre_bnfs;  | 
| 57397 | 2440  | 
val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs;  | 
| 62335 | 2441  | 
val fp_nesting_rel_eq_onps = map rel_eq_onp_of_bnf fp_nesting_bnfs;  | 
| 58347 | 2442  | 
val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs;  | 
| 58732 | 2443  | 
val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;  | 
| 57397 | 2444  | 
val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs;  | 
| 58328 | 2445  | 
val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs;  | 
| 62335 | 2446  | 
val live_nesting_rel_eq_onps = map rel_eq_onp_of_bnf live_nesting_bnfs;  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2447  | 
|
| 63824 | 2448  | 
val liveness = liveness_of_fp_bnf num_As any_fp_bnf;  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2449  | 
val live = live_of_bnf any_fp_bnf;  | 
| 52961 | 2450  | 
val _ =  | 
| 62324 | 2451  | 
if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs @ pred_bs) then  | 
2452  | 
warning "Map function, relator, and predicator names ignored"  | 
|
| 52961 | 2453  | 
else  | 
2454  | 
();  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2455  | 
|
| 63824 | 2456  | 
    val Bs = @{map 3} (fn alive => fn A as TFree (_, S) => fn B =>
 | 
2457  | 
if alive then resort_tfree_or_tvar S B else A)  | 
|
2458  | 
liveness As Bs0;  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2459  | 
|
| 58446 | 2460  | 
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);  | 
| 63839 | 2461  | 
val B_ify = Term.map_types B_ify_T;  | 
2462  | 
||
| 58732 | 2463  | 
val live_AsBs = filter (op <>) (As ~~ Bs);  | 
| 49167 | 2464  | 
|
| 63824 | 2465  | 
val abss = map #abs absT_infos;  | 
2466  | 
val reps = map #rep absT_infos;  | 
|
2467  | 
val absTs = map #absT absT_infos;  | 
|
2468  | 
val repTs = map #repT absT_infos;  | 
|
2469  | 
val abs_injects = map #abs_inject absT_infos;  | 
|
2470  | 
val abs_inverses = map #abs_inverse absT_infos;  | 
|
2471  | 
val type_definitions = map #type_definition absT_infos;  | 
|
2472  | 
||
| 49501 | 2473  | 
val ctors = map (mk_ctor As) ctors0;  | 
2474  | 
val dtors = map (mk_dtor As) dtors0;  | 
|
| 49124 | 2475  | 
|
| 49501 | 2476  | 
val fpTs = map (domain_type o fastype_of) dtors;  | 
| 58446 | 2477  | 
val fpBTs = map B_ify_T fpTs;  | 
| 
49362
 
1271aca16aed
make tactic more robust in the case where "asm_simp_tac" already finishes the job
 
blanchet 
parents: 
49361 
diff
changeset
 | 
2478  | 
|
| 63818 | 2479  | 
val real_unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fpTs);  | 
2480  | 
||
2481  | 
val ctr_Tsss = map (map (map real_unfreeze_fp)) ctrXs_Tsss;  | 
|
| 49203 | 2482  | 
val ns = map length ctr_Tsss;  | 
| 49212 | 2483  | 
val kss = map (fn n => 1 upto n) ns;  | 
| 49203 | 2484  | 
val mss = map (map length) ctr_Tsss;  | 
2485  | 
||
| 61551 | 2486  | 
val (xtor_co_recs, recs_args_types, corecs_args_types) =  | 
2487  | 
mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0;  | 
|
| 49210 | 2488  | 
|
| 63814 | 2489  | 
fun define_ctrs_dtrs_for_type_etc fp_bnf fp_b fpT C E ctor dtor xtor_co_rec ctor_dtor dtor_ctor  | 
2490  | 
ctor_inject pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm n ks ms  | 
|
2491  | 
abs abs_inject type_definition ctr_bindings ctr_mixfixes ctr_Tss disc_bindings sel_bindingss  | 
|
| 63818 | 2492  | 
raw_sel_default_eqs lthy =  | 
| 63814 | 2493  | 
let  | 
| 63815 | 2494  | 
val fp_b_name = Binding.name_of fp_b;  | 
2495  | 
||
| 63814 | 2496  | 
val ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) =  | 
| 63815 | 2497  | 
define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs  | 
| 63818 | 2498  | 
ctr_bindings ctr_mixfixes ctr_Tss lthy;  | 
| 63814 | 2499  | 
|
| 49203 | 2500  | 
val ctrs = map (mk_ctr As) ctrs0;  | 
| 49121 | 2501  | 
|
| 63818 | 2502  | 
val sel_default_eqs =  | 
2503  | 
let  | 
|
2504  | 
val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss;  | 
|
2505  | 
val sel_bTs =  | 
|
2506  | 
flat sel_bindingss ~~ flat sel_Tss  | 
|
2507  | 
|> filter_out (Binding.is_empty o fst)  | 
|
2508  | 
|> distinct (Binding.eq_name o apply2 fst);  | 
|
2509  | 
val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy  | 
|
2510  | 
in  | 
|
2511  | 
map (prepare_term sel_default_lthy) raw_sel_default_eqs  | 
|
2512  | 
end;  | 
|
| 63816 | 2513  | 
|
2514  | 
fun mk_binding pre =  | 
|
2515  | 
Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);  | 
|
2516  | 
||
2517  | 
fun massage_res (ctr_sugar, maps_sets_rels) =  | 
|
2518  | 
(maps_sets_rels, (ctrs, xss, ctor_iff_dtor_thm, ctr_defs, ctr_sugar));  | 
|
| 49119 | 2519  | 
in  | 
| 63839 | 2520  | 
(wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition  | 
| 63816 | 2521  | 
disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs  | 
| 58346 | 2522  | 
#> (fn (ctr_sugar, lthy) =>  | 
| 63814 | 2523  | 
derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs  | 
| 64416 | 2524  | 
fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps  | 
2525  | 
live_nesting_rel_eqs live_nesting_rel_eq_onps [] fp_b_name fp_bnf fp_bnfs fpT ctor  | 
|
2526  | 
ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms  | 
|
2527  | 
fp_rel_thm [] [] [] ctr_Tss ctr_sugar lthy  | 
|
| 58346 | 2528  | 
|>> pair ctr_sugar)  | 
| 52313 | 2529  | 
##>>  | 
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
2530  | 
(if fp = Least_FP then define_rec (the recs_args_types) mk_binding fpTs Cs reps  | 
| 
58352
 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 
blanchet 
parents: 
58347 
diff
changeset
 | 
2531  | 
else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec  | 
| 63816 | 2532  | 
#>> apfst massage_res, lthy)  | 
| 49119 | 2533  | 
end;  | 
| 49167 | 2534  | 
|
| 63817 | 2535  | 
fun wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types (wrap_one_etcs, lthy) =  | 
2536  | 
fold_map I wrap_one_etcs lthy  | 
|
| 62335 | 2537  | 
      |>> apsnd split_list o apfst (apsnd @{split_list 5} o apfst @{split_list 17} o split_list)
 | 
| 52342 | 2538  | 
o split_list;  | 
| 49226 | 2539  | 
|
| 58462 | 2540  | 
    fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects
 | 
| 58565 | 2541  | 
rel_distincts set_thmss =  | 
2542  | 
injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @  | 
|
2543  | 
set_thmss;  | 
|
| 
49479
 
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
 
blanchet 
parents: 
49478 
diff
changeset
 | 
2544  | 
|
| 58448 | 2545  | 
fun mk_co_rec_transfer_goals lthy co_recs =  | 
| 58446 | 2546  | 
let  | 
| 63839 | 2547  | 
val BE_ify = Term.subst_atomic_types (live_AsBs @ (Cs ~~ Es));  | 
| 58446 | 2548  | 
|
2549  | 
val ((Rs, Ss), names_lthy) = lthy  | 
|
| 58732 | 2550  | 
|> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs)  | 
| 58446 | 2551  | 
||>> mk_Frees "S" (map2 mk_pred2T Cs Es);  | 
2552  | 
||
| 63839 | 2553  | 
val co_recBs = map BE_ify co_recs;  | 
| 58446 | 2554  | 
in  | 
| 58448 | 2555  | 
(Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy)  | 
2556  | 
end;  | 
|
2557  | 
||
| 58966 | 2558  | 
fun derive_rec_transfer_thms lthy recs rec_defs (SOME (_, _, _, xsssss)) =  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2559  | 
let  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2560  | 
val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy recs;  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2561  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2562  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2563  | 
in  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2564  | 
Goal.prove_sorry lthy vars [] goal  | 
| 58446 | 2565  | 
          (fn {context = ctxt, prems = _} =>
 | 
| 61760 | 2566  | 
mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss)  | 
2567  | 
(map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs  | 
|
2568  | 
live_nesting_rel_eqs)  | 
|
| 
61116
 
6189d179c2b5
close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
 
wenzelm 
parents: 
61101 
diff
changeset
 | 
2569  | 
|> Thm.close_derivation  | 
| 58446 | 2570  | 
|> Conjunction.elim_balanced nn  | 
2571  | 
end;  | 
|
2572  | 
||
| 58732 | 2573  | 
fun derive_rec_o_map_thmss lthy recs rec_defs =  | 
| 62335 | 2574  | 
if live = 0 then  | 
2575  | 
replicate nn []  | 
|
| 58732 | 2576  | 
else  | 
2577  | 
let  | 
|
2578  | 
fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy);  | 
|
2579  | 
||
2580  | 
val maps0 = map map_of_bnf fp_bnfs;  | 
|
2581  | 
val f_names = variant_names num_As "f";  | 
|
| 63824 | 2582  | 
val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs));  | 
| 58732 | 2583  | 
val live_gs = AList.find (op =) (fs ~~ liveness) true;  | 
2584  | 
||
2585  | 
val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0;  | 
|
2586  | 
||
| 59819 | 2587  | 
val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs));  | 
| 58732 | 2588  | 
|
2589  | 
val num_rec_args = length rec_arg_Ts;  | 
|
2590  | 
val g_Ts = map B_ify_T rec_arg_Ts;  | 
|
2591  | 
val g_names = variant_names num_rec_args "g";  | 
|
2592  | 
val gs = map2 (curry Free) g_names g_Ts;  | 
|
2593  | 
val grecs = map (fn recx => Term.list_comb (Term.map_types B_ify_T recx, gs)) recs;  | 
|
2594  | 
||
2595  | 
val rec_o_map_lhss = map2 (curry HOLogic.mk_comp) grecs gmaps;  | 
|
2596  | 
||
| 63824 | 2597  | 
val ABfs = (As ~~ Bs) ~~ fs;  | 
| 58732 | 2598  | 
|
2599  | 
fun mk_rec_arg_arg (x as Free (_, T)) =  | 
|
2600  | 
let val U = B_ify_T T in  | 
|
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2601  | 
if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x  | 
| 58732 | 2602  | 
end;  | 
2603  | 
||
2604  | 
fun mk_rec_o_map_arg rec_arg_T h =  | 
|
2605  | 
let  | 
|
2606  | 
val x_Ts = binder_types rec_arg_T;  | 
|
2607  | 
val m = length x_Ts;  | 
|
2608  | 
val x_names = variant_names m "x";  | 
|
2609  | 
val xs = map2 (curry Free) x_names x_Ts;  | 
|
2610  | 
val xs' = map mk_rec_arg_arg xs;  | 
|
2611  | 
in  | 
|
2612  | 
fold_rev Term.lambda xs (Term.list_comb (h, xs'))  | 
|
2613  | 
end;  | 
|
2614  | 
||
2615  | 
fun mk_rec_o_map_rhs recx =  | 
|
2616  | 
let val args = map2 mk_rec_o_map_arg rec_arg_Ts gs in  | 
|
2617  | 
Term.list_comb (recx, args)  | 
|
2618  | 
end;  | 
|
2619  | 
||
2620  | 
val rec_o_map_rhss = map mk_rec_o_map_rhs recs;  | 
|
2621  | 
||
2622  | 
val rec_o_map_goals =  | 
|
2623  | 
map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo  | 
|
2624  | 
curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss;  | 
|
2625  | 
val rec_o_map_thms =  | 
|
2626  | 
            @{map 3} (fn goal => fn rec_def => fn ctor_rec_o_map =>
 | 
|
2627  | 
                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
 | 
|
| 58734 | 2628  | 
mk_co_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s  | 
2629  | 
abs_inverses ctor_rec_o_map)  | 
|
| 58732 | 2630  | 
|> Thm.close_derivation)  | 
2631  | 
rec_o_map_goals rec_defs xtor_co_rec_o_maps;  | 
|
2632  | 
in  | 
|
2633  | 
map single rec_o_map_thms  | 
|
2634  | 
end;  | 
|
2635  | 
||
| 55867 | 2636  | 
fun derive_note_induct_recs_thms_for_types  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
2637  | 
((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss,  | 
| 62335 | 2638  | 
rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss,  | 
| 58676 | 2639  | 
set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss),  | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2640  | 
(ctrss, _, ctor_iff_dtors, ctr_defss, ctr_sugars)),  | 
| 55864 | 2641  | 
(recs, rec_defs)), lthy) =  | 
| 
49202
 
f493cd25737f
some work towards iterator and recursor properties
 
blanchet 
parents: 
49201 
diff
changeset
 | 
2642  | 
let  | 
| 55869 | 2643  | 
val ((induct_thms, induct_thm, induct_attrs), (rec_thmss, rec_attrs)) =  | 
| 
58335
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
2644  | 
derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct  | 
| 
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
2645  | 
xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses  | 
| 
 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 
blanchet 
parents: 
58332 
diff
changeset
 | 
2646  | 
type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy;  | 
| 49201 | 2647  | 
|
| 58446 | 2648  | 
val rec_transfer_thmss =  | 
| 
58967
 
6b6032e99a4b
also generate '(co)rec_transfer' for (co)datatypes with 0 live type variables
 
desharna 
parents: 
58966 
diff
changeset
 | 
2649  | 
map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types);  | 
| 58446 | 2650  | 
|
| 51830 | 2651  | 
val induct_type_attr = Attrib.internal o K o Induct.induct_type;  | 
| 57471 | 2652  | 
val induct_pred_attr = Attrib.internal o K o Induct.induct_pred;  | 
2653  | 
||
2654  | 
val ((rel_induct_thmss, common_rel_induct_thms),  | 
|
2655  | 
(rel_induct_attrs, common_rel_induct_attrs)) =  | 
|
2656  | 
if live = 0 then  | 
|
2657  | 
((replicate nn [], []), ([], []))  | 
|
2658  | 
else  | 
|
2659  | 
let  | 
|
2660  | 
val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) =  | 
|
| 
62158
 
c25c62055180
generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
 
blanchet 
parents: 
62093 
diff
changeset
 | 
2661  | 
derive_rel_induct_thms_for_types lthy nn fpTs As Bs ctrss ctr_Tsss  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58577 
diff
changeset
 | 
2662  | 
(map #exhaust ctr_sugars) xtor_rel_co_induct ctr_defss ctor_injects  | 
| 58328 | 2663  | 
pre_rel_defs abs_inverses live_nesting_rel_eqs;  | 
| 57471 | 2664  | 
in  | 
2665  | 
((map single rel_induct_thms, single common_rel_induct_thm),  | 
|
2666  | 
(rel_induct_attrs, rel_induct_attrs))  | 
|
2667  | 
end;  | 
|
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2668  | 
|
| 58732 | 2669  | 
val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs;  | 
2670  | 
||
| 
51835
 
56523caf372f
Added maps, sets, rels to "simps" thm collection
 
blanchet 
parents: 
51833 
diff
changeset
 | 
2671  | 
val simp_thmss =  | 
| 61760 | 2672  | 
          @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss
 | 
2673  | 
set_thmss;  | 
|
| 49438 | 2674  | 
|
| 49337 | 2675  | 
val common_notes =  | 
| 57471 | 2676  | 
(if nn > 1 then  | 
| 57700 | 2677  | 
[(inductN, [induct_thm], K induct_attrs),  | 
2678  | 
(rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)]  | 
|
| 57633 | 2679  | 
else  | 
2680  | 
[])  | 
|
| 51780 | 2681  | 
|> massage_simple_notes fp_common_name;  | 
| 49337 | 2682  | 
|
| 49226 | 2683  | 
val notes =  | 
| 
55856
 
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
 
blanchet 
parents: 
55803 
diff
changeset
 | 
2684  | 
[(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]),  | 
| 55869 | 2685  | 
(recN, rec_thmss, K rec_attrs),  | 
| 58732 | 2686  | 
(rec_o_mapN, rec_o_map_thmss, K []),  | 
| 58446 | 2687  | 
(rec_transferN, rec_transfer_thmss, K []),  | 
| 57471 | 2688  | 
(rel_inductN, rel_induct_thmss, K (rel_induct_attrs @ [induct_pred_attr ""])),  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49583 
diff
changeset
 | 
2689  | 
(simpsN, simp_thmss, K [])]  | 
| 58346 | 2690  | 
|> massage_multi_notes fp_b_names fpTs;  | 
| 
49202
 
f493cd25737f
some work towards iterator and recursor properties
 
blanchet 
parents: 
49201 
diff
changeset
 | 
2691  | 
in  | 
| 51824 | 2692  | 
lthy  | 
| 
56641
 
029997d3b5d8
reverted rb458558cbcc2 -- better to keep 'case' and 'rec' distinct, otherwise we lose the connection between 'ctor_rec' (the low-level recursor) and 'rec'
 
blanchet 
parents: 
56640 
diff
changeset
 | 
2693  | 
|> Spec_Rules.add Spec_Rules.Equational (recs, flat rec_thmss)  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
2694  | 
|> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (flat rec_thmss))  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57568 
diff
changeset
 | 
2695  | 
|> Local_Theory.notes (common_notes @ notes)  | 
| 57633 | 2696  | 
(* for "datatype_realizer.ML": *)  | 
2697  | 
|>> name_noted_thms  | 
|
| 59794 | 2698  | 
(fst (dest_Type (hd fpTs)) ^ implode (map (prefix "_") (tl fp_b_names))) inductN  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
2699  | 
|-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos  | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2700  | 
fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars  | 
| 
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2701  | 
recs rec_defs map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn [])  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
2702  | 
(replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss  | 
| 62335 | 2703  | 
rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess  | 
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
2704  | 
ctr_transferss case_transferss disc_transferss sel_transferss (replicate nn [])  | 
| 
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
2705  | 
(replicate nn []) rec_transfer_thmss common_rel_induct_thms rel_induct_thmss []  | 
| 
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
2706  | 
(replicate nn []) rec_o_map_thmss  | 
| 
49202
 
f493cd25737f
some work towards iterator and recursor properties
 
blanchet 
parents: 
49201 
diff
changeset
 | 
2707  | 
end;  | 
| 
 
f493cd25737f
some work towards iterator and recursor properties
 
blanchet 
parents: 
49201 
diff
changeset
 | 
2708  | 
|
| 58448 | 2709  | 
fun derive_corec_transfer_thms lthy corecs corec_defs =  | 
2710  | 
let  | 
|
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2711  | 
val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy corecs;  | 
| 58448 | 2712  | 
val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types;  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2713  | 
val goal = Logic.mk_conjunction_balanced goals;  | 
| 
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2714  | 
val vars = Variable.add_free_names lthy goal [];  | 
| 58448 | 2715  | 
in  | 
| 
61334
 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 
traytel 
parents: 
61301 
diff
changeset
 | 
2716  | 
Goal.prove_sorry lthy vars [] goal  | 
| 58448 | 2717  | 
          (fn {context = ctxt, prems = _} =>
 | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
2718  | 
mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs)  | 
| 
58581
 
e2e2d775869c
rename 'xtor_co_rec_transfer_thms' to 'xtor_co_rec_transfers'
 
desharna 
parents: 
58580 
diff
changeset
 | 
2719  | 
type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs  | 
| 58448 | 2720  | 
live_nesting_rel_eqs (flat pgss) pss qssss gssss)  | 
| 
61116
 
6189d179c2b5
close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
 
wenzelm 
parents: 
61101 
diff
changeset
 | 
2721  | 
|> Thm.close_derivation  | 
| 58448 | 2722  | 
|> Conjunction.elim_balanced (length goals)  | 
2723  | 
end;  | 
|
2724  | 
||
| 58734 | 2725  | 
fun derive_map_o_corec_thmss lthy0 lthy2 corecs corec_defs =  | 
| 62335 | 2726  | 
if live = 0 then  | 
2727  | 
replicate nn []  | 
|
| 58734 | 2728  | 
else  | 
2729  | 
let  | 
|
2730  | 
fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0);  | 
|
| 63824 | 2731  | 
|
| 58734 | 2732  | 
val maps0 = map map_of_bnf fp_bnfs;  | 
2733  | 
val f_names = variant_names num_As "f";  | 
|
| 63824 | 2734  | 
val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs));  | 
| 58734 | 2735  | 
val live_fs = AList.find (op =) (fs ~~ liveness) true;  | 
2736  | 
||
2737  | 
val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0;  | 
|
2738  | 
||
| 59819 | 2739  | 
val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs));  | 
| 58734 | 2740  | 
|
2741  | 
val num_rec_args = length corec_arg_Ts;  | 
|
2742  | 
val g_names = variant_names num_rec_args "g";  | 
|
2743  | 
val gs = map2 (curry Free) g_names corec_arg_Ts;  | 
|
2744  | 
val gcorecs = map (fn corecx => Term.list_comb (corecx, gs)) corecs;  | 
|
2745  | 
||
2746  | 
val map_o_corec_lhss = map2 (curry HOLogic.mk_comp) fmaps gcorecs;  | 
|
2747  | 
||
| 63824 | 2748  | 
val ABfs = (As ~~ Bs) ~~ fs;  | 
| 58734 | 2749  | 
|
2750  | 
fun mk_map_o_corec_arg corec_argB_T g =  | 
|
2751  | 
let  | 
|
2752  | 
val T = range_type (fastype_of g);  | 
|
2753  | 
val U = range_type corec_argB_T;  | 
|
2754  | 
in  | 
|
| 
64627
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2755  | 
if T = U then  | 
| 
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2756  | 
g  | 
| 
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2757  | 
else  | 
| 
 
8d7cb22482e3
generalized ML function (towards nonuniform datatypes)
 
blanchet 
parents: 
64626 
diff
changeset
 | 
2758  | 
HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g)  | 
| 58734 | 2759  | 
end;  | 
2760  | 
||
2761  | 
fun mk_map_o_corec_rhs corecx =  | 
|
2762  | 
let val args = map2 (mk_map_o_corec_arg o B_ify_T) corec_arg_Ts gs in  | 
|
2763  | 
Term.list_comb (B_ify corecx, args)  | 
|
2764  | 
end;  | 
|
2765  | 
||
2766  | 
val map_o_corec_rhss = map mk_map_o_corec_rhs corecs;  | 
|
2767  | 
||
2768  | 
val map_o_corec_goals =  | 
|
2769  | 
map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo  | 
|
2770  | 
curry HOLogic.mk_eq) map_o_corec_lhss map_o_corec_rhss;  | 
|
2771  | 
||
2772  | 
val map_o_corec_thms =  | 
|
2773  | 
            @{map 3} (fn goal => fn corec_def => fn dtor_map_o_corec =>
 | 
|
2774  | 
                Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} =>
 | 
|
2775  | 
mk_co_rec_o_map_tac ctxt corec_def pre_map_defs live_nesting_map_ident0s  | 
|
2776  | 
abs_inverses dtor_map_o_corec)  | 
|
2777  | 
|> Thm.close_derivation)  | 
|
2778  | 
map_o_corec_goals corec_defs xtor_co_rec_o_maps;  | 
|
2779  | 
in  | 
|
2780  | 
map single map_o_corec_thms  | 
|
2781  | 
end;  | 
|
2782  | 
||
| 55867 | 2783  | 
fun derive_note_coinduct_corecs_thms_for_types  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
2784  | 
((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss,  | 
| 62335 | 2785  | 
rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss,  | 
2786  | 
set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss),  | 
|
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2787  | 
(_, _, ctor_iff_dtors, ctr_defss, ctr_sugars)),  | 
| 55867 | 2788  | 
(corecs, corec_defs)), lthy) =  | 
| 49212 | 2789  | 
let  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2790  | 
val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)],  | 
| 
57260
 
8747af0d1012
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
 
blanchet 
parents: 
57206 
diff
changeset
 | 
2791  | 
(coinduct_attrs, common_coinduct_attrs)),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2792  | 
corec_thmss, corec_disc_thmss,  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2793  | 
(corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) =  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2794  | 
derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct  | 
| 57397 | 2795  | 
dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss  | 
| 
64637
 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 
blanchet 
parents: 
64627 
diff
changeset
 | 
2796  | 
ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs;  | 
| 49212 | 2797  | 
|
| 58264 | 2798  | 
fun distinct_prems ctxt thm =  | 
| 
57890
 
1e13f63fb452
use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
 
traytel 
parents: 
57882 
diff
changeset
 | 
2799  | 
Goal.prove (*no sorry*) ctxt [] []  | 
| 58264 | 2800  | 
(thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies)  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58935 
diff
changeset
 | 
2801  | 
(fn _ => HEADGOAL (cut_tac thm THEN' assume_tac ctxt) THEN ALLGOALS eq_assume_tac);  | 
| 57489 | 2802  | 
|
2803  | 
fun eq_ifIN _ [thm] = thm  | 
|
2804  | 
| eq_ifIN ctxt (thm :: thms) =  | 
|
2805  | 
              distinct_prems ctxt (@{thm eq_ifI} OF
 | 
|
2806  | 
                (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]})
 | 
|
2807  | 
[thm, eq_ifIN ctxt thms]));  | 
|
2808  | 
||
| 
57890
 
1e13f63fb452
use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
 
traytel 
parents: 
57882 
diff
changeset
 | 
2809  | 
val corec_code_thms = map (eq_ifIN lthy) corec_thmss;  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2810  | 
val corec_sel_thmss = map flat corec_sel_thmsss;  | 
| 53475 | 2811  | 
|
| 51830 | 2812  | 
val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type;  | 
| 57302 | 2813  | 
val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred;  | 
| 
49479
 
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
 
blanchet 
parents: 
49478 
diff
changeset
 | 
2814  | 
|
| 55867 | 2815  | 
val flat_corec_thms = append oo append;  | 
| 
49479
 
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
 
blanchet 
parents: 
49478 
diff
changeset
 | 
2816  | 
|
| 
58967
 
6b6032e99a4b
also generate '(co)rec_transfer' for (co)datatypes with 0 live type variables
 
desharna 
parents: 
58966 
diff
changeset
 | 
2817  | 
val corec_transfer_thmss = map single (derive_corec_transfer_thms lthy corecs corec_defs);  | 
| 58448 | 2818  | 
|
| 57471 | 2819  | 
val ((rel_coinduct_thmss, common_rel_coinduct_thms),  | 
2820  | 
(rel_coinduct_attrs, common_rel_coinduct_attrs)) =  | 
|
2821  | 
if live = 0 then  | 
|
2822  | 
((replicate nn [], []), ([], []))  | 
|
2823  | 
else  | 
|
2824  | 
let  | 
|
2825  | 
val ((rel_coinduct_thms, common_rel_coinduct_thm),  | 
|
2826  | 
(rel_coinduct_attrs, common_rel_coinduct_attrs)) =  | 
|
| 
62158
 
c25c62055180
generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
 
blanchet 
parents: 
62093 
diff
changeset
 | 
2827  | 
derive_rel_coinduct_thms_for_types lthy nn fpTs ns As Bs mss ctr_sugars abs_inverses  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58577 
diff
changeset
 | 
2828  | 
abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss xtor_rel_co_induct  | 
| 58328 | 2829  | 
live_nesting_rel_eqs;  | 
| 57471 | 2830  | 
in  | 
2831  | 
((map single rel_coinduct_thms, single common_rel_coinduct_thm),  | 
|
2832  | 
(rel_coinduct_attrs, common_rel_coinduct_attrs))  | 
|
2833  | 
end;  | 
|
2834  | 
||
| 58734 | 2835  | 
val map_o_corec_thmss = derive_map_o_corec_thmss lthy lthy corecs corec_defs;  | 
2836  | 
||
| 57700 | 2837  | 
val (set_induct_thms, set_induct_attrss) =  | 
2838  | 
derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars)  | 
|
| 
58580
 
8ee2d984caa8
rename 'dtor_set_induct_thms' to 'dtor_set_inducts'
 
desharna 
parents: 
58579 
diff
changeset
 | 
2839  | 
(map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_inducts  | 
| 62320 | 2840  | 
(map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) dtor_ctors abs_inverses  | 
| 57700 | 2841  | 
|> split_list;  | 
2842  | 
||
| 
49479
 
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
 
blanchet 
parents: 
49478 
diff
changeset
 | 
2843  | 
val simp_thmss =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2844  | 
          @{map 6} mk_simp_thms ctr_sugars
 | 
| 
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58586 
diff
changeset
 | 
2845  | 
            (@{map 3} flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss)
 | 
| 58565 | 2846  | 
map_thmss rel_injectss rel_distinctss set_thmss;  | 
| 
49479
 
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
 
blanchet 
parents: 
49478 
diff
changeset
 | 
2847  | 
|
| 49342 | 2848  | 
val common_notes =  | 
| 57700 | 2849  | 
(set_inductN, set_induct_thms, nth set_induct_attrss) ::  | 
| 57302 | 2850  | 
(if nn > 1 then  | 
| 57700 | 2851  | 
[(coinductN, [coinduct_thm], K common_coinduct_attrs),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2852  | 
(coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2853  | 
(rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)]  | 
| 57700 | 2854  | 
else [])  | 
| 51780 | 2855  | 
|> massage_simple_notes fp_common_name;  | 
| 49342 | 2856  | 
|
| 49212 | 2857  | 
val notes =  | 
| 51780 | 2858  | 
[(coinductN, map single coinduct_thms,  | 
| 
51810
 
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
 
blanchet 
parents: 
51808 
diff
changeset
 | 
2859  | 
fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2860  | 
(coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs),  | 
| 57489 | 2861  | 
(corecN, corec_thmss, K []),  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
2862  | 
(corec_codeN, map single corec_code_thms, K (nitpicksimp_attrs)),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2863  | 
(corec_discN, corec_disc_thmss, K []),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2864  | 
(corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs),  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2865  | 
(corec_selN, corec_sel_thmss, K corec_sel_attrs),  | 
| 58448 | 2866  | 
(corec_transferN, corec_transfer_thmss, K []),  | 
| 58734 | 2867  | 
(map_o_corecN, map_o_corec_thmss, K []),  | 
| 57471 | 2868  | 
(rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])),  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2869  | 
(simpsN, simp_thmss, K [])]  | 
| 58346 | 2870  | 
|> massage_multi_notes fp_b_names fpTs;  | 
| 49212 | 2871  | 
in  | 
| 51824 | 2872  | 
lthy  | 
| 55867 | 2873  | 
|> fold (curry (Spec_Rules.add Spec_Rules.Equational) corecs)  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57932 
diff
changeset
 | 
2874  | 
[flat corec_sel_thmss, flat corec_thmss]  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64638 
diff
changeset
 | 
2875  | 
|> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) corec_code_thms)  | 
| 
57631
 
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
 
blanchet 
parents: 
57568 
diff
changeset
 | 
2876  | 
|> Local_Theory.notes (common_notes @ notes)  | 
| 
58189
 
9d714be4f028
added 'plugins' option to control which hooks are enabled
 
blanchet 
parents: 
58188 
diff
changeset
 | 
2877  | 
|-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos  | 
| 
62321
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2878  | 
fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars  | 
| 
 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 
blanchet 
parents: 
62320 
diff
changeset
 | 
2879  | 
corecs corec_defs map_thmss [coinduct_thm, coinduct_strong_thm]  | 
| 
58177
 
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
 
blanchet 
parents: 
58175 
diff
changeset
 | 
2880  | 
(transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss  | 
| 
58672
 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 
desharna 
parents: 
58671 
diff
changeset
 | 
2881  | 
corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss  | 
| 62335 | 2882  | 
rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess  | 
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
2883  | 
ctr_transferss case_transferss disc_transferss sel_transferss corec_disc_iff_thmss  | 
| 62335 | 2884  | 
(map single corec_code_thms) corec_transfer_thmss common_rel_coinduct_thms  | 
2885  | 
rel_coinduct_thmss set_induct_thms (replicate nn (if nn = 1 then set_induct_thms else []))  | 
|
| 
63859
 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 
blanchet 
parents: 
63853 
diff
changeset
 | 
2886  | 
map_o_corec_thmss  | 
| 49212 | 2887  | 
end;  | 
2888  | 
||
| 63817 | 2889  | 
val lthy = lthy  | 
| 
58177
 
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
 
blanchet 
parents: 
58175 
diff
changeset
 | 
2890  | 
|> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs  | 
| 63814 | 2891  | 
      |> @{fold_map 29} define_ctrs_dtrs_for_type_etc fp_bnfs fp_bs fpTs Cs Es ctors dtors
 | 
| 61364 | 2892  | 
xtor_co_recs ctor_dtors dtor_ctors ctor_injects pre_map_defs pre_set_defss pre_rel_defs  | 
2893  | 
xtor_maps xtor_setss xtor_rels ns kss mss abss abs_injects type_definitions ctr_bindingss  | 
|
2894  | 
ctr_mixfixess ctr_Tsss disc_bindingss sel_bindingsss raw_sel_default_eqss  | 
|
| 62335 | 2895  | 
|> wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types  | 
| 55966 | 2896  | 
|> case_fp fp derive_note_induct_recs_thms_for_types  | 
| 62383 | 2897  | 
derive_note_coinduct_corecs_thms_for_types;  | 
| 49167 | 2898  | 
|
2899  | 
    val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
 | 
|
| 52899 | 2900  | 
co_prefix fp ^ "datatype"));  | 
| 49112 | 2901  | 
in  | 
| 63817 | 2902  | 
lthy  | 
| 49112 | 2903  | 
end;  | 
2904  | 
||
| 
63826
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63825 
diff
changeset
 | 
2905  | 
fun co_datatypes fp = define_co_datatypes (K I) (K I) (K I) (K I) fp;  | 
| 
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63825 
diff
changeset
 | 
2906  | 
|
| 
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63825 
diff
changeset
 | 
2907  | 
fun co_datatype_cmd fp construct_fp options lthy =  | 
| 63786 | 2908  | 
define_co_datatypes Plugin_Name.make_filter Typedecl.read_constraint Syntax.parse_typ  | 
| 
63826
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63825 
diff
changeset
 | 
2909  | 
Syntax.parse_term fp construct_fp options lthy  | 
| 
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63825 
diff
changeset
 | 
2910  | 
  handle EMPTY_DATATYPE s => error ("Cannot define empty datatype " ^ quote s);
 | 
| 49119 | 2911  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2912  | 
val parse_ctr_arg =  | 
| 57205 | 2913  | 
  @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"}
 | 
2914  | 
|| Parse.typ >> pair Binding.empty;  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2915  | 
|
| 
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2916  | 
val parse_ctr_specs =  | 
| 57091 | 2917  | 
Parse.enum1 "|" (parse_ctr_spec Parse.binding parse_ctr_arg -- Parse.opt_mixfix);  | 
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
2918  | 
|
| 
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
2919  | 
val parse_spec =  | 
| 
57206
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57205 
diff
changeset
 | 
2920  | 
parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix --  | 
| 62324 | 2921  | 
  (@{keyword "="} |-- parse_ctr_specs) -- parse_map_rel_pred_bindings -- parse_sel_default_eqs;
 | 
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
2922  | 
|
| 
55469
 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 
blanchet 
parents: 
55468 
diff
changeset
 | 
2923  | 
val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec;  | 
| 49278 | 2924  | 
|
| 63786 | 2925  | 
fun parse_co_datatype_cmd fp construct_fp =  | 
2926  | 
parse_co_datatype >> co_datatype_cmd fp construct_fp;  | 
|
| 49112 | 2927  | 
|
2928  | 
end;  |