author | desharna |
Mon, 06 Oct 2014 13:36:42 +0200 | |
changeset 58575 | 629891fd8c51 |
parent 58574 | e66ed9634a74 |
child 58576 | 1f4a2d8142fe |
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, |
|
13 |
ctr_defs: thm list, |
|
58569 | 14 |
ctr_sugar: Ctr_Sugar.ctr_sugar, |
58570 | 15 |
ctr_transfers: thm list, |
58571 | 16 |
case_transfers: thm list, |
17 |
disc_transfers: thm list} |
|
58458 | 18 |
|
19 |
type fp_bnf_sugar = |
|
58462 | 20 |
{map_thms: thm list, |
58560 | 21 |
map_disc_iffs: thm list, |
58561 | 22 |
map_sels: thm list, |
58462 | 23 |
rel_injects: thm list, |
58562 | 24 |
rel_distincts: thm list, |
58563 | 25 |
rel_sels: thm list, |
58564 | 26 |
rel_intros: thm list, |
58565 | 27 |
rel_cases: thm list, |
58566 | 28 |
set_thms: thm list, |
58567 | 29 |
set_sels: thm list, |
58568 | 30 |
set_intros: thm list, |
31 |
set_cases: thm list} |
|
58458 | 32 |
|
58459 | 33 |
type fp_co_induct_sugar = |
58461 | 34 |
{co_rec: term, |
35 |
common_co_inducts: thm list, |
|
36 |
co_inducts: thm list, |
|
37 |
co_rec_def: thm, |
|
58459 | 38 |
co_rec_thms: thm list, |
39 |
co_rec_discs: thm list, |
|
58572 | 40 |
co_rec_disc_iffs: thm list, |
58573 | 41 |
co_rec_selss: thm list list, |
42 |
co_rec_codes: thm list, |
|
58574 | 43 |
co_rec_transfers: thm list, |
58575 | 44 |
common_rel_co_inducts: thm list, |
45 |
rel_co_inducts: thm list} |
|
58457 | 46 |
|
58180 | 47 |
type fp_sugar = |
48 |
{T: typ, |
|
49 |
BT: typ, |
|
50 |
X: typ, |
|
51 |
fp: BNF_Util.fp_kind, |
|
52 |
fp_res_index: int, |
|
53 |
fp_res: BNF_FP_Util.fp_result, |
|
54 |
pre_bnf: BNF_Def.bnf, |
|
55 |
absT_info: BNF_Comp.absT_info, |
|
56 |
fp_nesting_bnfs: BNF_Def.bnf list, |
|
57 |
live_nesting_bnfs: BNF_Def.bnf list, |
|
58457 | 58 |
fp_ctr_sugar: fp_ctr_sugar, |
59 |
fp_bnf_sugar: fp_bnf_sugar, |
|
60 |
fp_co_induct_sugar: fp_co_induct_sugar}; |
|
58180 | 61 |
|
62 |
val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar |
|
63 |
val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar |
|
64 |
val fp_sugar_of: Proof.context -> string -> fp_sugar option |
|
65 |
val fp_sugar_of_global: theory -> string -> fp_sugar option |
|
66 |
val fp_sugars_of: Proof.context -> fp_sugar list |
|
67 |
val fp_sugars_of_global: theory -> fp_sugar list |
|
58186 | 68 |
val fp_sugars_interpretation: string -> (fp_sugar list -> theory -> theory) -> |
58180 | 69 |
(fp_sugar list -> local_theory -> local_theory)-> theory -> theory |
58188 | 70 |
val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory |
58182 | 71 |
val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory |
58188 | 72 |
val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory |
51852 | 73 |
|
53106 | 74 |
val co_induct_of: 'a list -> 'a |
75 |
val strong_co_induct_of: 'a list -> 'a |
|
76 |
||
52868 | 77 |
val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list -> |
78 |
'a list |
|
58230 | 79 |
val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list |
57397 | 80 |
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
|
81 |
|
58214 | 82 |
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
|
83 |
|
54256 | 84 |
val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms |
58115 | 85 |
val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms |
54256 | 86 |
|
53746
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents:
53741
diff
changeset
|
87 |
type gfp_sugar_thms = |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
88 |
((thm list * thm) list * (Token.src list * Token.src list)) |
57489 | 89 |
* thm list list |
90 |
* thm list list |
|
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
91 |
* (thm list list * Token.src list) |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
92 |
* (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
|
93 |
|
54256 | 94 |
val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms |
58115 | 95 |
val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms |
54256 | 96 |
|
55867 | 97 |
val mk_co_recs_prelims: BNF_Util.fp_kind -> typ list list list -> typ list -> typ list -> |
98 |
typ list -> typ list -> int list -> int list list -> term list -> Proof.context -> |
|
99 |
(term list |
|
100 |
* (typ list list * typ list list list list * term list list * term list list list list) option |
|
55862
b458558cbcc2
optimized simple non-recursive datatypes by reusing 'case' for 'rec' constant
blanchet
parents:
55861
diff
changeset
|
101 |
* (string * term list * term list list |
58448 | 102 |
* (((term list list * term list list * term list list list list * term list list list list) |
103 |
* term list list list) * typ list)) option) |
|
55862
b458558cbcc2
optimized simple non-recursive datatypes by reusing 'case' for 'rec' constant
blanchet
parents:
55861
diff
changeset
|
104 |
* Proof.context |
55772
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
blanchet
parents:
55701
diff
changeset
|
105 |
val repair_nullary_single_ctr: typ list list -> typ list list |
55867 | 106 |
val mk_corec_p_pred_types: typ list -> int list -> typ list list |
107 |
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
|
108 |
int list list -> term -> |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
109 |
typ list list |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
110 |
* (typ list list list list * typ list list list * typ list list list list * typ list) |
58210 | 111 |
val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory -> |
58209 | 112 |
(term * thm) * local_theory |
55867 | 113 |
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
|
114 |
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
|
115 |
(string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context -> |
55864 | 116 |
(term * thm) * Proof.context |
55867 | 117 |
val define_corec: 'a * term list * term list list |
58448 | 118 |
* (((term list list * term list list * term list list list list * term list list list list) |
119 |
* term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list -> |
|
120 |
term list -> term -> local_theory -> (term * thm) * local_theory |
|
58214 | 121 |
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
|
122 |
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
|
123 |
Token.src list * Token.src list |
58335
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58332
diff
changeset
|
124 |
val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list -> |
55867 | 125 |
('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
|
126 |
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
|
127 |
typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list -> |
55864 | 128 |
term list -> thm list -> Proof.context -> lfp_sugar_thms |
55867 | 129 |
val derive_coinduct_corecs_thms_for_types: BNF_Def.bnf list -> |
58448 | 130 |
string * term list * term list list |
131 |
* (((term list list * term list list * term list list list list * term list list list list) |
|
132 |
* term list list list) * typ list) -> |
|
55867 | 133 |
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
|
134 |
typ list -> typ list list list -> int list list -> int list list -> int list -> thm list -> |
55864 | 135 |
thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list -> |
56347
6edbd4d20717
added new-style (co)datatype interpretation hook
blanchet
parents:
56254
diff
changeset
|
136 |
thm list -> (thm list -> thm list) -> Proof.context -> gfp_sugar_thms |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
137 |
|
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55573
diff
changeset
|
138 |
val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> |
52207
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents:
52197
diff
changeset
|
139 |
binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
140 |
BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
141 |
BNF_FP_Util.fp_result * local_theory) -> |
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
142 |
Ctr_Sugar.ctr_options |
57206
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
143 |
* ((((((binding option * (typ * sort)) list * binding) * mixfix) |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
144 |
* ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding)) |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
145 |
* term list) list -> |
49297 | 146 |
local_theory -> local_theory |
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55573
diff
changeset
|
147 |
val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> |
51867 | 148 |
binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
149 |
BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
150 |
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
|
151 |
(local_theory -> local_theory) parser |
49112 | 152 |
end; |
153 |
||
49636 | 154 |
structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR = |
49112 | 155 |
struct |
156 |
||
54006 | 157 |
open Ctr_Sugar |
55571 | 158 |
open BNF_FP_Rec_Sugar_Util |
49119 | 159 |
open BNF_Util |
53222 | 160 |
open BNF_Comp |
49214
2a3cb4c71b87
construct the right iterator theorem in the recursive case
blanchet
parents:
49213
diff
changeset
|
161 |
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
|
162 |
open BNF_FP_Util |
49636 | 163 |
open BNF_FP_Def_Sugar_Tactics |
49119 | 164 |
|
51788 | 165 |
val EqN = "Eq_"; |
57489 | 166 |
|
58093 | 167 |
val case_transferN = "case_transfer"; |
57999 | 168 |
val ctr_transferN = "ctr_transfer"; |
58095 | 169 |
val disc_transferN = "disc_transfer"; |
57489 | 170 |
val corec_codeN = "corec_code"; |
58448 | 171 |
val corec_transferN = "corec_transfer"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
172 |
val map_disc_iffN = "map_disc_iff"; |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
173 |
val map_selN = "map_sel"; |
58446 | 174 |
val rec_transferN = "rec_transfer"; |
57893 | 175 |
val set_casesN = "set_cases"; |
57891 | 176 |
val set_introsN = "set_intros"; |
57700 | 177 |
val set_inductN = "set_induct"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
178 |
val set_selN = "set_sel"; |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51769
diff
changeset
|
179 |
|
58460 | 180 |
type fp_ctr_sugar = |
181 |
{ctrXs_Tss: typ list list, |
|
182 |
ctr_defs: thm list, |
|
58569 | 183 |
ctr_sugar: Ctr_Sugar.ctr_sugar, |
58570 | 184 |
ctr_transfers: thm list, |
58571 | 185 |
case_transfers: thm list, |
186 |
disc_transfers: thm list}; |
|
58458 | 187 |
|
188 |
type fp_bnf_sugar = |
|
58462 | 189 |
{map_thms: thm list, |
58560 | 190 |
map_disc_iffs: thm list, |
58561 | 191 |
map_sels: thm list, |
58462 | 192 |
rel_injects: thm list, |
58562 | 193 |
rel_distincts: thm list, |
58563 | 194 |
rel_sels: thm list, |
58564 | 195 |
rel_intros: thm list, |
58565 | 196 |
rel_cases: thm list, |
58566 | 197 |
set_thms: thm list, |
58567 | 198 |
set_sels: thm list, |
58568 | 199 |
set_intros: thm list, |
200 |
set_cases: thm list}; |
|
58458 | 201 |
|
58459 | 202 |
type fp_co_induct_sugar = |
58461 | 203 |
{co_rec: term, |
204 |
common_co_inducts: thm list, |
|
205 |
co_inducts: thm list, |
|
206 |
co_rec_def: thm, |
|
58459 | 207 |
co_rec_thms: thm list, |
208 |
co_rec_discs: thm list, |
|
58572 | 209 |
co_rec_disc_iffs: thm list, |
58573 | 210 |
co_rec_selss: thm list list, |
211 |
co_rec_codes: thm list, |
|
58574 | 212 |
co_rec_transfers: thm list, |
58575 | 213 |
common_rel_co_inducts: thm list, |
214 |
rel_co_inducts: thm list}; |
|
58457 | 215 |
|
58180 | 216 |
type fp_sugar = |
217 |
{T: typ, |
|
218 |
BT: typ, |
|
219 |
X: typ, |
|
220 |
fp: fp_kind, |
|
221 |
fp_res_index: int, |
|
222 |
fp_res: fp_result, |
|
223 |
pre_bnf: bnf, |
|
224 |
absT_info: absT_info, |
|
225 |
fp_nesting_bnfs: bnf list, |
|
226 |
live_nesting_bnfs: bnf list, |
|
58457 | 227 |
fp_ctr_sugar: fp_ctr_sugar, |
228 |
fp_bnf_sugar: fp_bnf_sugar, |
|
229 |
fp_co_induct_sugar: fp_co_induct_sugar}; |
|
58180 | 230 |
|
58562 | 231 |
fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_sels, rel_injects, rel_distincts, |
58568 | 232 |
rel_sels, rel_intros, rel_cases, set_thms, set_sels, set_intros, set_cases} : fp_bnf_sugar) = |
58462 | 233 |
{map_thms = map (Morphism.thm phi) map_thms, |
58560 | 234 |
map_disc_iffs = map (Morphism.thm phi) map_disc_iffs, |
58561 | 235 |
map_sels = map (Morphism.thm phi) map_sels, |
58462 | 236 |
rel_injects = map (Morphism.thm phi) rel_injects, |
58562 | 237 |
rel_distincts = map (Morphism.thm phi) rel_distincts, |
58563 | 238 |
rel_sels = map (Morphism.thm phi) rel_sels, |
58564 | 239 |
rel_intros = map (Morphism.thm phi) rel_intros, |
58565 | 240 |
rel_cases = map (Morphism.thm phi) rel_cases, |
58566 | 241 |
set_thms = map (Morphism.thm phi) set_thms, |
58567 | 242 |
set_sels = map (Morphism.thm phi) set_sels, |
58568 | 243 |
set_intros = map (Morphism.thm phi) set_intros, |
244 |
set_cases = map (Morphism.thm phi) set_cases}; |
|
58458 | 245 |
|
58461 | 246 |
fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms, |
58574 | 247 |
co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, |
58575 | 248 |
common_rel_co_inducts, rel_co_inducts} : fp_co_induct_sugar) = |
58461 | 249 |
{co_rec = Morphism.term phi co_rec, |
250 |
common_co_inducts = map (Morphism.thm phi) common_co_inducts, |
|
251 |
co_inducts = map (Morphism.thm phi) co_inducts, |
|
252 |
co_rec_def = Morphism.thm phi co_rec_def, |
|
58459 | 253 |
co_rec_thms = map (Morphism.thm phi) co_rec_thms, |
254 |
co_rec_discs = map (Morphism.thm phi) co_rec_discs, |
|
58572 | 255 |
co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs, |
58573 | 256 |
co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss, |
257 |
co_rec_codes = map (Morphism.thm phi) co_rec_codes, |
|
58574 | 258 |
co_rec_transfers = map (Morphism.thm phi) co_rec_transfers, |
58575 | 259 |
common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts, |
260 |
rel_co_inducts = map (Morphism.thm phi) rel_co_inducts}; |
|
58459 | 261 |
|
58571 | 262 |
fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctr_defs, ctr_sugar, ctr_transfers, case_transfers, |
263 |
disc_transfers} : fp_ctr_sugar) = |
|
58460 | 264 |
{ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss, |
265 |
ctr_defs = map (Morphism.thm phi) ctr_defs, |
|
58569 | 266 |
ctr_sugar = morph_ctr_sugar phi ctr_sugar, |
58570 | 267 |
ctr_transfers = map (Morphism.thm phi) ctr_transfers, |
58571 | 268 |
case_transfers = map (Morphism.thm phi) case_transfers, |
269 |
disc_transfers = map (Morphism.thm phi) disc_transfers}; |
|
58460 | 270 |
|
58180 | 271 |
fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, absT_info, fp_nesting_bnfs, |
58462 | 272 |
live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar, |
58460 | 273 |
fp_co_induct_sugar} : fp_sugar) = |
58180 | 274 |
{T = Morphism.typ phi T, |
275 |
BT = Morphism.typ phi BT, |
|
276 |
X = Morphism.typ phi X, |
|
277 |
fp = fp, |
|
278 |
fp_res = morph_fp_result phi fp_res, |
|
279 |
fp_res_index = fp_res_index, |
|
280 |
pre_bnf = morph_bnf phi pre_bnf, |
|
281 |
absT_info = morph_absT_info phi absT_info, |
|
282 |
fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs, |
|
283 |
live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs, |
|
58460 | 284 |
fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar, |
58458 | 285 |
fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar, |
58459 | 286 |
fp_co_induct_sugar = morph_fp_co_induct_sugar phi fp_co_induct_sugar}; |
58180 | 287 |
|
288 |
val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism; |
|
289 |
||
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
290 |
structure Data = Generic_Data |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
291 |
( |
51840 | 292 |
type T = fp_sugar Symtab.table; |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
293 |
val empty = Symtab.empty; |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
294 |
val extend = I; |
55394
d5ebe055b599
more liberal merging of BNFs and constructor sugar
blanchet
parents:
55356
diff
changeset
|
295 |
fun merge data : T = Symtab.merge (K true) data; |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
296 |
); |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
297 |
|
58116 | 298 |
fun fp_sugar_of_generic context = |
299 |
Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context) |
|
300 |
||
301 |
fun fp_sugars_of_generic context = |
|
302 |
Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; |
|
53907 | 303 |
|
58116 | 304 |
val fp_sugar_of = fp_sugar_of_generic o Context.Proof; |
305 |
val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory; |
|
306 |
||
307 |
val fp_sugars_of = fp_sugars_of_generic o Context.Proof; |
|
308 |
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
|
309 |
|
53106 | 310 |
fun co_induct_of (i :: _) = i; |
311 |
fun strong_co_induct_of [_, s] = s; |
|
312 |
||
58177
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents:
58175
diff
changeset
|
313 |
structure FP_Sugar_Interpretation = Local_Interpretation |
56347
6edbd4d20717
added new-style (co)datatype interpretation hook
blanchet
parents:
56254
diff
changeset
|
314 |
( |
56637
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
315 |
type T = fp_sugar list; |
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
316 |
val eq: T * T -> bool = op = o pairself (map #T); |
56347
6edbd4d20717
added new-style (co)datatype interpretation hook
blanchet
parents:
56254
diff
changeset
|
317 |
); |
6edbd4d20717
added new-style (co)datatype interpretation hook
blanchet
parents:
56254
diff
changeset
|
318 |
|
58186 | 319 |
fun with_transfer_fp_sugars g fp_sugars thy = g (map (transfer_fp_sugar thy) fp_sugars) thy; |
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
|
320 |
|
58186 | 321 |
fun fp_sugars_interpretation name = |
322 |
FP_Sugar_Interpretation.interpretation name o with_transfer_fp_sugars; |
|
323 |
||
58182 | 324 |
val interpret_fp_sugars = FP_Sugar_Interpretation.data; |
56347
6edbd4d20717
added new-style (co)datatype interpretation hook
blanchet
parents:
56254
diff
changeset
|
325 |
|
58182 | 326 |
fun register_fp_sugars_raw fp_sugars = |
56648
2ffa440b3074
manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents:
56641
diff
changeset
|
327 |
fold (fn fp_sugar as {T = Type (s, _), ...} => |
56637
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
328 |
Local_Theory.declaration {syntax = false, pervasive = true} |
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
329 |
(fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar)))) |
58182 | 330 |
fp_sugars; |
331 |
||
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
332 |
fun register_fp_sugars plugins fp_sugars = |
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
333 |
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
|
334 |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
335 |
fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs |
58462 | 336 |
live_nesting_bnfs fp_res ctrXs_Tsss ctr_defss ctr_sugars co_recs co_rec_defs map_thmss |
58177
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents:
58175
diff
changeset
|
337 |
common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss |
58566 | 338 |
rel_distinctss map_disc_iffss map_selss rel_selss rel_intross rel_casess set_thmss set_selss |
58573 | 339 |
set_intross set_casess ctr_transferss case_transferss disc_transferss co_rec_disc_iffss |
58575 | 340 |
co_rec_codess co_rec_transferss common_rel_co_inducts rel_co_inductss noted = |
56637
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
341 |
let |
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
342 |
val fp_sugars = |
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
343 |
map_index (fn (kk, T) => |
56648
2ffa440b3074
manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents:
56641
diff
changeset
|
344 |
{T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk, |
57397 | 345 |
pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, |
346 |
fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, |
|
58460 | 347 |
fp_ctr_sugar = |
348 |
{ctrXs_Tss = nth ctrXs_Tsss kk, |
|
349 |
ctr_defs = nth ctr_defss kk, |
|
58569 | 350 |
ctr_sugar = nth ctr_sugars kk, |
58570 | 351 |
ctr_transfers = nth ctr_transferss kk, |
58571 | 352 |
case_transfers = nth case_transferss kk, |
353 |
disc_transfers = nth disc_transferss kk}, |
|
58459 | 354 |
fp_bnf_sugar = |
58462 | 355 |
{map_thms = nth map_thmss kk, |
58560 | 356 |
map_disc_iffs = nth map_disc_iffss kk, |
58561 | 357 |
map_sels = nth map_selss kk, |
58462 | 358 |
rel_injects = nth rel_injectss kk, |
58562 | 359 |
rel_distincts = nth rel_distinctss kk, |
58563 | 360 |
rel_sels = nth rel_selss kk, |
58564 | 361 |
rel_intros = nth rel_intross kk, |
58565 | 362 |
rel_cases = nth rel_casess kk, |
58566 | 363 |
set_thms = nth set_thmss kk, |
58567 | 364 |
set_sels = nth set_selss kk, |
58568 | 365 |
set_intros = nth set_intross kk, |
366 |
set_cases = nth set_casess kk}, |
|
58459 | 367 |
fp_co_induct_sugar = |
58461 | 368 |
{co_rec = nth co_recs kk, |
369 |
common_co_inducts = common_co_inducts, |
|
370 |
co_inducts = nth co_inductss kk, |
|
371 |
co_rec_def = nth co_rec_defs kk, |
|
58459 | 372 |
co_rec_thms = nth co_rec_thmss kk, |
373 |
co_rec_discs = nth co_rec_discss kk, |
|
58572 | 374 |
co_rec_disc_iffs = nth co_rec_disc_iffss kk, |
58573 | 375 |
co_rec_selss = nth co_rec_selsss kk, |
376 |
co_rec_codes = nth co_rec_codess kk, |
|
58574 | 377 |
co_rec_transfers = nth co_rec_transferss kk, |
58575 | 378 |
common_rel_co_inducts = common_rel_co_inducts, |
379 |
rel_co_inducts = nth rel_co_inductss kk}} |
|
57633 | 380 |
|> morph_fp_sugar (substitute_noted_thm noted)) Ts; |
56637
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
381 |
in |
58182 | 382 |
register_fp_sugars_raw fp_sugars |
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
383 |
#> fold (interpret_bnf plugins) (#bnfs fp_res) |
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
384 |
#> interpret_fp_sugars plugins fp_sugars |
56637
d1d55f1bbc8a
invoke 'fp_sugar' interpretation hook on mutually recursive clique
blanchet
parents:
56523
diff
changeset
|
385 |
end; |
51824 | 386 |
|
49622 | 387 |
fun quasi_unambiguous_case_names names = |
388 |
let |
|
389 |
val ps = map (`Long_Name.base_name) names; |
|
390 |
val dups = Library.duplicates (op =) (map fst ps); |
|
391 |
fun underscore s = |
|
392 |
let val ss = space_explode Long_Name.separator s in |
|
393 |
space_implode "_" (drop (length ss - 2) ss) |
|
394 |
end; |
|
395 |
in |
|
396 |
map (fn (base, full) => if member (op =) dups base then underscore full else base) ps |
|
58285 | 397 |
|> Name.variant_list [] |
49622 | 398 |
end; |
399 |
||
58180 | 400 |
fun zipper_map f = |
401 |
let |
|
402 |
fun zed _ [] = [] |
|
403 |
| zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys |
|
404 |
in zed [] end; |
|
405 |
||
58265
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
blanchet
parents:
58264
diff
changeset
|
406 |
fun uncurry_thm 0 thm = thm |
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
blanchet
parents:
58264
diff
changeset
|
407 |
| uncurry_thm 1 thm = thm |
cae4f3d14b05
prevent infinite loop when type variables are of a non-'type' sort
blanchet
parents:
58264
diff
changeset
|
408 |
| 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
|
409 |
|
58180 | 410 |
fun choose_binary_fun fs AB = |
411 |
find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs; |
|
412 |
fun build_binary_fun_app fs t u = |
|
413 |
Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u)); |
|
414 |
||
415 |
fun build_the_rel rel_table ctxt Rs Ts A B = |
|
416 |
build_rel rel_table ctxt Ts (the o choose_binary_fun Rs) (A, B); |
|
417 |
fun build_rel_app ctxt Rs Ts t u = |
|
418 |
build_the_rel [] ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u; |
|
419 |
||
420 |
fun mk_parametricity_goal ctxt Rs t u = |
|
421 |
let val prem = build_the_rel [] ctxt Rs [] (fastype_of t) (fastype_of u) in |
|
422 |
HOLogic.mk_Trueprop (prem $ t $ u) |
|
423 |
end; |
|
424 |
||
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
|
425 |
val name_of_set = name_of_const "set function" domain_type; |
58180 | 426 |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
427 |
val mp_conj = @{thm mp_conj}; |
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
428 |
|
55397
c2506f61fd26
generate 'fundec_cong' attribute for new-style (co)datatypes
blanchet
parents:
55394
diff
changeset
|
429 |
val fundefcong_attrs = @{attributes [fundef_cong]}; |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54107
diff
changeset
|
430 |
val nitpicksimp_attrs = @{attributes [nitpick_simp]}; |
49300 | 431 |
val simp_attrs = @{attributes [simp]}; |
432 |
||
49297 | 433 |
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
|
434 |
|
55871 | 435 |
fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss); |
51829 | 436 |
|
55871 | 437 |
fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss |
438 |
| flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) = |
|
439 |
p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss; |
|
51829 | 440 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
441 |
fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
442 |
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
|
443 |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
444 |
fun flip_rels lthy n thm = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
445 |
let |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
446 |
val Rs = Term.add_vars (prop_of thm) []; |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
447 |
val Rs' = rev (drop (length Rs - n) Rs); |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
448 |
val cRs = map (fn f => (certify lthy (Var f), certify lthy (mk_flip f))) Rs'; |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
449 |
in |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
450 |
Drule.cterm_instantiate cRs thm |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
451 |
end; |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
452 |
|
49536 | 453 |
fun mk_ctor_or_dtor get_T Ts t = |
454 |
let val Type (_, Ts0) = get_T (fastype_of t) in |
|
455 |
Term.subst_atomic_types (Ts0 ~~ Ts) t |
|
456 |
end; |
|
457 |
||
458 |
val mk_ctor = mk_ctor_or_dtor range_type; |
|
459 |
val mk_dtor = mk_ctor_or_dtor domain_type; |
|
460 |
||
55867 | 461 |
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
|
462 |
let |
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
blanchet
parents:
51910
diff
changeset
|
463 |
val nn = length fpTs; |
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
blanchet
parents:
51910
diff
changeset
|
464 |
val (fpTs0, Cs0) = |
52207
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents:
52197
diff
changeset
|
465 |
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
|
466 |
|> split_list; |
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
blanchet
parents:
51910
diff
changeset
|
467 |
val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs); |
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
blanchet
parents:
51910
diff
changeset
|
468 |
in |
52170
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
blanchet
parents:
52169
diff
changeset
|
469 |
map (Term.subst_TVars rho) ts0 |
51911
6c425d450a8c
more robust iterator construction (needed for mutualized FPs)
blanchet
parents:
51910
diff
changeset
|
470 |
end; |
51827 | 471 |
|
57665 | 472 |
fun mk_set Ts t = |
473 |
subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; |
|
474 |
||
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
475 |
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
|
476 |
(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
|
477 |
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
|
478 |
| _ => replicate n false); |
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
479 |
|
55700 | 480 |
fun cannot_merge_types fp = |
481 |
error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters"); |
|
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
482 |
|
55700 | 483 |
fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp; |
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
484 |
|
55700 | 485 |
fun merge_type_args fp (As, As') = |
486 |
if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp; |
|
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
487 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
488 |
fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
489 |
fun type_binding_of_spec (((((_, b), _), _), _), _) = b; |
57206
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
490 |
fun mixfix_of_spec ((((_, mx), _), _), _) = mx; |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
491 |
fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs; |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
492 |
fun map_binding_of_spec ((_, (b, _)), _) = b; |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
493 |
fun rel_binding_of_spec ((_, (_, b)), _) = b; |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
494 |
fun sel_default_eqs_of_spec (_, ts) = ts; |
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
495 |
|
57397 | 496 |
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
|
497 |
let |
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
498 |
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
|
499 |
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
|
500 |
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
|
501 |
end |
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
502 |
| 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
|
503 |
in snd oo add end; |
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
504 |
|
57397 | 505 |
fun nesting_bnfs ctxt ctr_Tsss Us = |
506 |
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
|
507 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
508 |
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
|
509 |
|
58346 | 510 |
fun massage_simple_notes base = |
511 |
filter_out (null o #2) |
|
512 |
#> map (fn (thmN, thms, f_attrs) => |
|
513 |
((Binding.qualify true base (Binding.name thmN), []), |
|
514 |
map_index (fn (i, thm) => ([thm], f_attrs i)) thms)); |
|
515 |
||
516 |
fun massage_multi_notes b_names Ts = |
|
517 |
maps (fn (thmN, thmss, attrs) => |
|
518 |
map3 (fn b_name => fn Type (T_name, _) => fn thms => |
|
519 |
((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])])) |
|
520 |
b_names Ts thmss) |
|
521 |
#> filter_out (null o fst o hd o snd); |
|
522 |
||
58347 | 523 |
fun derive_map_set_rel_thms plugins fp live As Bs abs_inverses ctr_defs' fp_nesting_set_maps |
524 |
live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs fp_b_name fp_bnf fpT ctor |
|
525 |
ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm |
|
526 |
ctr_Tss abs |
|
58346 | 527 |
({casex, case_thms, discs, selss, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss, injects, |
528 |
distincts, distinct_discsss, ...} : ctr_sugar) |
|
529 |
lthy = |
|
530 |
if live = 0 then |
|
58571 | 531 |
(([], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy) |
58346 | 532 |
else |
533 |
let |
|
58347 | 534 |
val n = length ctr_Tss; |
535 |
val ks = 1 upto n; |
|
536 |
val ms = map length ctr_Tss; |
|
537 |
||
58446 | 538 |
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); |
58346 | 539 |
|
58446 | 540 |
val fpBT = B_ify_T fpT; |
58346 | 541 |
val live_AsBs = filter (op <>) (As ~~ Bs); |
542 |
val fTs = map (op -->) live_AsBs; |
|
543 |
||
58347 | 544 |
val (((((((([C, E], xss), yss), fs), Rs), ta), tb), thesis), names_lthy) = lthy |
58435 | 545 |
|> fold (fold Variable.declare_typ) [As, Bs] |
58346 | 546 |
|> mk_TFrees 2 |
58347 | 547 |
||>> mk_Freess "x" ctr_Tss |
58446 | 548 |
||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss) |
58346 | 549 |
||>> mk_Frees "f" fTs |
550 |
||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) |
|
551 |
||>> yield_singleton (mk_Frees "a") fpT |
|
552 |
||>> yield_singleton (mk_Frees "b") fpBT |
|
553 |
||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT; |
|
554 |
||
555 |
val mapx = mk_map live As Bs (map_of_bnf fp_bnf); |
|
556 |
val ctrAs = map (mk_ctr As) ctrs; |
|
557 |
val ctrBs = map (mk_ctr Bs) ctrs; |
|
558 |
val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf); |
|
559 |
val setAs = map (mk_set As) (sets_of_bnf fp_bnf); |
|
560 |
val discAs = map (mk_disc_or_sel As) discs; |
|
561 |
val discBs = map (mk_disc_or_sel Bs) discs; |
|
562 |
val selAss = map (map (mk_disc_or_sel As)) selss; |
|
563 |
val discAs_selAss = flat (map2 (map o pair) discAs selAss); |
|
564 |
||
565 |
val ctor_cong = |
|
566 |
if fp = Least_FP then |
|
567 |
Drule.dummy_thm |
|
568 |
else |
|
569 |
let val ctor' = mk_ctor Bs ctor in |
|
570 |
cterm_instantiate_pos [NONE, NONE, SOME (certify lthy ctor')] arg_cong |
|
571 |
end; |
|
572 |
||
573 |
fun mk_cIn ctor k xs = |
|
574 |
let val absT = domain_type (fastype_of ctor) in |
|
575 |
mk_absumprod absT abs n k xs |
|
576 |
|> fp = Greatest_FP ? curry (op $) ctor |
|
577 |
|> certify lthy |
|
578 |
end; |
|
579 |
||
580 |
val cxIns = map2 (mk_cIn ctor) ks xss; |
|
58446 | 581 |
val cyIns = map2 (mk_cIn (Term.map_types B_ify_T ctor)) ks yss; |
58346 | 582 |
|
583 |
fun mk_map_thm ctr_def' cxIn = |
|
584 |
fold_thms lthy [ctr_def'] |
|
585 |
(unfold_thms lthy (o_apply :: pre_map_def :: |
|
586 |
(if fp = Least_FP then [] else [dtor_ctor]) @ sumprod_thms_map @ abs_inverses) |
|
587 |
(cterm_instantiate_pos (map (SOME o certify lthy) fs @ [SOME cxIn]) |
|
588 |
(if fp = Least_FP then fp_map_thm |
|
589 |
else fp_map_thm RS ctor_cong RS (ctor_dtor RS sym RS trans)))) |
|
590 |
|> singleton (Proof_Context.export names_lthy lthy); |
|
591 |
||
592 |
fun mk_set0_thm fp_set_thm ctr_def' cxIn = |
|
593 |
fold_thms lthy [ctr_def'] |
|
594 |
(unfold_thms lthy (pre_set_defs @ fp_nesting_set_maps @ live_nesting_set_maps @ |
|
595 |
(if fp = Least_FP then [] else [dtor_ctor]) @ basic_sumprod_thms_set @ |
|
596 |
@{thms UN_Un sup_assoc[THEN sym]} @ abs_inverses) |
|
597 |
(cterm_instantiate_pos [SOME cxIn] fp_set_thm)) |
|
598 |
|> singleton (Proof_Context.export names_lthy lthy); |
|
599 |
||
600 |
fun mk_set0_thms fp_set_thm = map2 (mk_set0_thm fp_set_thm) ctr_defs' cxIns; |
|
601 |
||
602 |
val map_thms = map2 mk_map_thm ctr_defs' cxIns; |
|
603 |
val set0_thmss = map mk_set0_thms fp_set_thms; |
|
604 |
val set0_thms = flat set0_thmss; |
|
605 |
val set_thms = set0_thms |
|
606 |
|> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left}); |
|
607 |
||
608 |
val rel_infos = (ctr_defs' ~~ cxIns, ctr_defs' ~~ cyIns); |
|
609 |
||
610 |
fun mk_rel_thm postproc ctr_defs' cxIn cyIn = |
|
611 |
fold_thms lthy ctr_defs' |
|
58347 | 612 |
(unfold_thms lthy (pre_rel_def :: abs_inverses @ |
58346 | 613 |
(if fp = Least_FP then [] else [dtor_ctor]) @ sumprod_thms_rel @ |
614 |
@{thms vimage2p_def sum.inject sum.distinct(1)[THEN eq_False[THEN iffD2]]}) |
|
615 |
(cterm_instantiate_pos (map (SOME o certify lthy) Rs @ [SOME cxIn, SOME cyIn]) |
|
616 |
fp_rel_thm)) |
|
617 |
|> postproc |
|
618 |
|> singleton (Proof_Context.export names_lthy lthy); |
|
619 |
||
620 |
fun mk_rel_inject_thm ((ctr_def', cxIn), (_, cyIn)) = |
|
621 |
mk_rel_thm (unfold_thms lthy @{thms eq_sym_Unity_conv}) [ctr_def'] cxIn cyIn; |
|
622 |
||
623 |
fun mk_rel_intro_thm m thm = |
|
624 |
uncurry_thm m (thm RS iffD2) handle THM _ => thm; |
|
625 |
||
626 |
fun mk_half_rel_distinct_thm ((xctr_def', cxIn), (yctr_def', cyIn)) = |
|
627 |
mk_rel_thm (fn thm => thm RS @{thm eq_False[THEN iffD1]}) [xctr_def', yctr_def'] cxIn cyIn; |
|
628 |
||
629 |
val rel_flip = rel_flip_of_bnf fp_bnf; |
|
630 |
||
631 |
fun mk_other_half_rel_distinct_thm thm = |
|
632 |
flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2); |
|
633 |
||
634 |
val rel_inject_thms = map mk_rel_inject_thm (op ~~ rel_infos); |
|
635 |
val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms; |
|
636 |
||
637 |
val half_rel_distinct_thmss = map (map mk_half_rel_distinct_thm) (mk_half_pairss rel_infos); |
|
638 |
val other_half_rel_distinct_thmss = |
|
639 |
map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss; |
|
640 |
val (rel_distinct_thms, _) = |
|
641 |
join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss; |
|
642 |
||
643 |
val rel_eq_thms = |
|
644 |
map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @ |
|
645 |
map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms; |
|
646 |
||
647 |
val ctr_transfer_thms = |
|
648 |
let val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs in |
|
649 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
650 |
(fn {context = ctxt, prems = _} => |
|
651 |
mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs) |
|
652 |
|> Conjunction.elim_balanced (length goals) |
|
653 |
|> Proof_Context.export names_lthy lthy |
|
654 |
|> map Thm.close_derivation |
|
655 |
end; |
|
656 |
||
657 |
val (set_cases_thms, set_cases_attrss) = |
|
658 |
let |
|
659 |
fun mk_prems assms elem t ctxt = |
|
660 |
(case fastype_of t of |
|
661 |
Type (type_name, xs) => |
|
662 |
(case bnf_of ctxt type_name of |
|
663 |
NONE => ([], ctxt) |
|
664 |
| SOME bnf => |
|
665 |
apfst flat (fold_map (fn set => fn ctxt => |
|
666 |
let |
|
667 |
val T = HOLogic.dest_setT (range_type (fastype_of set)); |
|
668 |
val new_var = not (T = fastype_of elem); |
|
669 |
val (x, ctxt') = |
|
670 |
if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt); |
|
671 |
in |
|
672 |
mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt' |
|
673 |
|>> map (new_var ? Logic.all x) |
|
674 |
end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt)) |
|
675 |
| T => |
|
676 |
rpair ctxt |
|
677 |
(if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else [])); |
|
678 |
in |
|
679 |
split_list (map (fn set => |
|
680 |
let |
|
681 |
val A = HOLogic.dest_setT (range_type (fastype_of set)); |
|
682 |
val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy; |
|
683 |
val premss = |
|
684 |
map (fn ctr => |
|
685 |
let |
|
686 |
val (args, names_lthy) = |
|
687 |
mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy; |
|
688 |
in |
|
689 |
flat (zipper_map (fn (prev_args, arg, next_args) => |
|
690 |
let |
|
691 |
val (args_with_elem, args_without_elem) = |
|
692 |
if fastype_of arg = A then |
|
693 |
(prev_args @ [elem] @ next_args, prev_args @ next_args) |
|
694 |
else |
|
695 |
`I (prev_args @ [arg] @ next_args); |
|
696 |
in |
|
697 |
mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))] |
|
698 |
elem arg names_lthy |
|
699 |
|> fst |
|
700 |
|> map (fold_rev Logic.all args_without_elem) |
|
701 |
end) args) |
|
702 |
end) ctrAs; |
|
703 |
val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); |
|
704 |
val thm = |
|
705 |
Goal.prove_sorry lthy [] (flat premss) goal (fn {context = ctxt, prems} => |
|
706 |
mk_set_cases_tac ctxt (certify ctxt ta) prems exhaust set_thms) |
|
707 |
|> singleton (Proof_Context.export names_lthy lthy) |
|
708 |
|> Thm.close_derivation |
|
709 |
|> rotate_prems ~1; |
|
710 |
||
711 |
val consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1)); |
|
712 |
val cases_set_attr = |
|
713 |
Attrib.internal (K (Induct.cases_pred (name_of_set set))); |
|
58456 | 714 |
|
715 |
val ctr_names = quasi_unambiguous_case_names (flat |
|
716 |
(map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs))); |
|
717 |
val case_names_attr = Attrib.internal (K (Rule_Cases.case_names ctr_names)); |
|
58346 | 718 |
in |
719 |
(* TODO: @{attributes [elim?]} *) |
|
58456 | 720 |
(thm, [consumes_attr, cases_set_attr, case_names_attr]) |
58346 | 721 |
end) setAs) |
722 |
end; |
|
723 |
||
724 |
val set_intros_thms = |
|
725 |
let |
|
726 |
fun mk_goals A setA ctr_args t ctxt = |
|
727 |
(case fastype_of t of |
|
728 |
Type (type_name, innerTs) => |
|
729 |
(case bnf_of ctxt type_name of |
|
730 |
NONE => ([], ctxt) |
|
731 |
| SOME bnf => |
|
732 |
apfst flat (fold_map (fn set => fn ctxt => |
|
733 |
let |
|
734 |
val T = HOLogic.dest_setT (range_type (fastype_of set)); |
|
735 |
val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; |
|
736 |
val assm = mk_Trueprop_mem (x, set $ t); |
|
737 |
in |
|
738 |
apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args x ctxt') |
|
739 |
end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt)) |
|
740 |
| T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt)); |
|
741 |
||
742 |
val (goals, names_lthy) = |
|
743 |
apfst flat (fold_map (fn set => fn ctxt => |
|
744 |
let val A = HOLogic.dest_setT (range_type (fastype_of set)) in |
|
745 |
apfst flat (fold_map (fn ctr => fn ctxt => |
|
746 |
let |
|
747 |
val (args, ctxt') = mk_Frees "a" (binder_types (fastype_of ctr)) ctxt; |
|
748 |
val ctr_args = Term.list_comb (ctr, args); |
|
749 |
in |
|
750 |
apfst flat (fold_map (mk_goals A set ctr_args) args ctxt') |
|
751 |
end) ctrAs ctxt) |
|
752 |
end) setAs lthy); |
|
753 |
in |
|
754 |
if null goals then [] |
|
755 |
else |
|
756 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
757 |
(fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms) |
|
758 |
|> Conjunction.elim_balanced (length goals) |
|
759 |
|> Proof_Context.export names_lthy lthy |
|
760 |
|> map Thm.close_derivation |
|
761 |
end; |
|
762 |
||
763 |
val rel_sel_thms = |
|
764 |
let |
|
765 |
val selBss = map (map (mk_disc_or_sel Bs)) selss; |
|
766 |
val n = length discAs; |
|
767 |
||
768 |
fun mk_conjunct n k discA selAs discB selBs = |
|
769 |
(if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @ |
|
770 |
(if null selAs then |
|
771 |
[] |
|
772 |
else |
|
773 |
[Library.foldr HOLogic.mk_imp |
|
774 |
(if n = 1 then [] else [discA $ ta, discB $ tb], |
|
775 |
Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs []) |
|
776 |
(map (rapp ta) selAs) (map (rapp tb) selBs)))]); |
|
777 |
||
778 |
val goals = |
|
779 |
if n = 0 then |
|
780 |
[] |
|
781 |
else |
|
782 |
[mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb, |
|
783 |
(case flat (map5 (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of |
|
784 |
[] => @{term True} |
|
785 |
| conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))]; |
|
786 |
||
787 |
fun prove goal = |
|
788 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
|
789 |
mk_rel_sel_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust (flat disc_thmss) |
|
790 |
(flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |
|
791 |
|> singleton (Proof_Context.export names_lthy lthy) |
|
792 |
|> Thm.close_derivation; |
|
793 |
in |
|
794 |
map prove goals |
|
795 |
end; |
|
796 |
||
797 |
val (rel_cases_thm, rel_cases_attrs) = |
|
798 |
let |
|
799 |
val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb; |
|
800 |
val ctrBs = map (mk_ctr Bs) ctrs; |
|
801 |
||
802 |
fun mk_assms ctrA ctrB ctxt = |
|
803 |
let |
|
804 |
val argA_Ts = binder_types (fastype_of ctrA); |
|
805 |
val argB_Ts = binder_types (fastype_of ctrB); |
|
806 |
val ((argAs, argBs), names_ctxt) = ctxt |
|
807 |
|> mk_Frees "x" argA_Ts |
|
808 |
||>> mk_Frees "y" argB_Ts; |
|
809 |
val ctrA_args = list_comb (ctrA, argAs); |
|
810 |
val ctrB_args = list_comb (ctrB, argBs); |
|
811 |
in |
|
812 |
(fold_rev Logic.all (argAs @ argBs) (Logic.list_implies |
|
813 |
(mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) :: |
|
814 |
map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs, |
|
815 |
thesis)), |
|
816 |
names_ctxt) |
|
817 |
end; |
|
818 |
||
819 |
val (assms, names_lthy) = fold_map2 mk_assms ctrAs ctrBs names_lthy; |
|
820 |
val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); |
|
821 |
val thm = |
|
822 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
|
823 |
mk_rel_cases_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust injects |
|
824 |
rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |
|
825 |
|> singleton (Proof_Context.export names_lthy lthy) |
|
826 |
|> Thm.close_derivation; |
|
827 |
||
828 |
val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs); |
|
829 |
val case_names_attr = Attrib.internal (K (Rule_Cases.case_names ctr_names)); |
|
830 |
val consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1)); |
|
831 |
val cases_pred_attr = Attrib.internal o K o Induct.cases_pred; |
|
832 |
in |
|
833 |
(thm, [consumes_attr, case_names_attr, cases_pred_attr ""]) |
|
834 |
end; |
|
835 |
||
58570 | 836 |
val case_transfer_thm = |
58346 | 837 |
let |
58435 | 838 |
val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy; |
58346 | 839 |
val caseA = mk_case As C casex; |
840 |
val caseB = mk_case Bs E casex; |
|
58435 | 841 |
val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB; |
58346 | 842 |
in |
843 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
|
844 |
mk_case_transfer_tac ctxt rel_cases_thm case_thms) |
|
845 |
|> singleton (Proof_Context.export names_lthy lthy) |
|
846 |
|> Thm.close_derivation |
|
847 |
end; |
|
848 |
||
849 |
val disc_transfer_thms = |
|
850 |
let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in |
|
851 |
if null goals then |
|
852 |
[] |
|
853 |
else |
|
854 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
855 |
(K (mk_disc_transfer_tac (the_single rel_sel_thms) (the_single exhaust_discs) |
|
856 |
(flat (flat distinct_discsss)))) |
|
857 |
|> Conjunction.elim_balanced (length goals) |
|
858 |
|> Proof_Context.export names_lthy lthy |
|
859 |
|> map Thm.close_derivation |
|
860 |
end; |
|
861 |
||
862 |
val map_disc_iff_thms = |
|
863 |
let |
|
864 |
val discsB = map (mk_disc_or_sel Bs) discs; |
|
865 |
val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs; |
|
866 |
||
867 |
fun mk_goal (discA_t, discB) = |
|
868 |
if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then |
|
869 |
NONE |
|
870 |
else |
|
871 |
SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t)); |
|
872 |
||
873 |
val goals = map_filter mk_goal (discsA_t ~~ discsB); |
|
874 |
in |
|
875 |
if null goals then |
|
876 |
[] |
|
877 |
else |
|
878 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
879 |
(fn {context = ctxt, prems = _} => |
|
58560 | 880 |
mk_map_disc_iff_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss) map_thms) |
58346 | 881 |
|> Conjunction.elim_balanced (length goals) |
882 |
|> Proof_Context.export names_lthy lthy |
|
883 |
|> map Thm.close_derivation |
|
884 |
end; |
|
885 |
||
886 |
val map_sel_thms = |
|
887 |
let |
|
888 |
fun mk_goal (discA, selA) = |
|
889 |
let |
|
890 |
val prem = Term.betapply (discA, ta); |
|
891 |
val selB = mk_disc_or_sel Bs selA; |
|
892 |
val lhs = selB $ (Term.list_comb (mapx, fs) $ ta); |
|
893 |
val lhsT = fastype_of lhs; |
|
894 |
val map_rhsT = |
|
895 |
map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT; |
|
896 |
val map_rhs = build_map lthy [] |
|
897 |
(the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT); |
|
898 |
val rhs = (case map_rhs of |
|
899 |
Const (@{const_name id}, _) => selA $ ta |
|
900 |
| _ => map_rhs $ (selA $ ta)); |
|
901 |
val concl = mk_Trueprop_eq (lhs, rhs); |
|
902 |
in |
|
903 |
if is_refl_bool prem then concl |
|
904 |
else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl) |
|
905 |
end; |
|
906 |
||
907 |
val goals = map mk_goal discAs_selAss; |
|
908 |
in |
|
909 |
if null goals then |
|
910 |
[] |
|
911 |
else |
|
912 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
913 |
(fn {context = ctxt, prems = _} => |
|
914 |
mk_map_sel_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss) map_thms |
|
915 |
(flat sel_thmss) live_nesting_map_id0s) |
|
916 |
|> Conjunction.elim_balanced (length goals) |
|
917 |
|> Proof_Context.export names_lthy lthy |
|
918 |
|> map Thm.close_derivation |
|
919 |
end; |
|
920 |
||
921 |
val set_sel_thms = |
|
922 |
let |
|
923 |
fun mk_goal discA selA setA ctxt = |
|
924 |
let |
|
925 |
val prem = Term.betapply (discA, ta); |
|
926 |
val sel_rangeT = range_type (fastype_of selA); |
|
927 |
val A = HOLogic.dest_setT (range_type (fastype_of setA)); |
|
928 |
||
929 |
fun travese_nested_types t ctxt = |
|
930 |
(case fastype_of t of |
|
931 |
Type (type_name, innerTs) => |
|
932 |
(case bnf_of ctxt type_name of |
|
933 |
NONE => ([], ctxt) |
|
934 |
| SOME bnf => |
|
935 |
let |
|
936 |
fun seq_assm a set ctxt = |
|
937 |
let |
|
938 |
val T = HOLogic.dest_setT (range_type (fastype_of set)); |
|
939 |
val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; |
|
940 |
val assm = mk_Trueprop_mem (x, set $ a); |
|
941 |
in |
|
942 |
travese_nested_types x ctxt' |
|
943 |
|>> map (Logic.mk_implies o pair assm) |
|
944 |
end; |
|
945 |
in |
|
946 |
fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |
|
947 |
|>> flat |
|
948 |
end) |
|
949 |
| T => |
|
950 |
if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt)); |
|
951 |
||
952 |
val (concls, ctxt') = |
|
953 |
if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt) |
|
954 |
else travese_nested_types (selA $ ta) ctxt; |
|
955 |
in |
|
956 |
if exists_subtype_in [A] sel_rangeT then |
|
957 |
if is_refl_bool prem then (concls, ctxt') |
|
958 |
else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt') |
|
959 |
else |
|
960 |
([], ctxt) |
|
961 |
end; |
|
962 |
val (goals, names_lthy) = apfst (flat o flat) (fold_map (fn (disc, sel) => |
|
963 |
fold_map (mk_goal disc sel) setAs) discAs_selAss names_lthy); |
|
964 |
in |
|
965 |
if null goals then |
|
966 |
[] |
|
967 |
else |
|
968 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
969 |
(fn {context = ctxt, prems = _} => |
|
970 |
mk_set_sel_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) |
|
971 |
set0_thms) |
|
972 |
|> Conjunction.elim_balanced (length goals) |
|
973 |
|> Proof_Context.export names_lthy lthy |
|
974 |
|> map Thm.close_derivation |
|
975 |
end; |
|
976 |
||
977 |
val code_attrs = if plugins code_plugin then [Code.add_default_eqn_attrib] else []; |
|
978 |
||
979 |
val anonymous_notes = |
|
980 |
[(rel_eq_thms, code_attrs @ nitpicksimp_attrs)] |
|
981 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
|
982 |
||
983 |
val notes = |
|
58570 | 984 |
[(case_transferN, [case_transfer_thm], K []), |
58346 | 985 |
(ctr_transferN, ctr_transfer_thms, K []), |
986 |
(disc_transferN, disc_transfer_thms, K []), |
|
987 |
(mapN, map_thms, K (code_attrs @ nitpicksimp_attrs @ simp_attrs)), |
|
988 |
(map_disc_iffN, map_disc_iff_thms, K simp_attrs), |
|
989 |
(map_selN, map_sel_thms, K []), |
|
990 |
(rel_casesN, [rel_cases_thm], K rel_cases_attrs), |
|
991 |
(rel_distinctN, rel_distinct_thms, K simp_attrs), |
|
992 |
(rel_injectN, rel_inject_thms, K simp_attrs), |
|
993 |
(rel_introsN, rel_intro_thms, K []), |
|
994 |
(rel_selN, rel_sel_thms, K []), |
|
995 |
(setN, set_thms, K (code_attrs @ nitpicksimp_attrs @ simp_attrs)), |
|
996 |
(set_casesN, set_cases_thms, nth set_cases_attrss), |
|
997 |
(set_introsN, set_intros_thms, K []), |
|
998 |
(set_selN, set_sel_thms, K [])] |
|
999 |
|> massage_simple_notes fp_b_name; |
|
1000 |
||
1001 |
val (noted, lthy') = |
|
1002 |
lthy |
|
1003 |
|> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) map_thms) |
|
1004 |
|> fp = Least_FP |
|
1005 |
? Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) rel_eq_thms) |
|
1006 |
|> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) set0_thms) |
|
1007 |
|> Local_Theory.notes (anonymous_notes @ notes); |
|
1008 |
||
1009 |
val subst = Morphism.thm (substitute_noted_thm noted); |
|
1010 |
in |
|
58561 | 1011 |
((map subst map_thms, |
1012 |
map subst map_disc_iff_thms, |
|
1013 |
map subst map_sel_thms, |
|
1014 |
map subst rel_inject_thms, |
|
1015 |
map subst rel_distinct_thms, |
|
58562 | 1016 |
map subst rel_sel_thms, |
58563 | 1017 |
map subst rel_intro_thms, |
58564 | 1018 |
[subst rel_cases_thm], |
58566 | 1019 |
map subst set_thms, |
58567 | 1020 |
map subst set_sel_thms, |
58568 | 1021 |
map subst set_intros_thms, |
58569 | 1022 |
map subst set_cases_thms, |
58570 | 1023 |
map subst ctr_transfer_thms, |
58571 | 1024 |
[subst case_transfer_thm], |
1025 |
map subst disc_transfer_thms), lthy') |
|
58346 | 1026 |
end; |
1027 |
||
58214 | 1028 |
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
|
1029 |
|
55867 | 1030 |
fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) = |
54256 | 1031 |
((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs), |
58291 | 1032 |
(map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms; |
54256 | 1033 |
|
58115 | 1034 |
val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism; |
54256 | 1035 |
|
53746
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents:
53741
diff
changeset
|
1036 |
type gfp_sugar_thms = |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
1037 |
((thm list * thm) list * (Token.src list * Token.src list)) |
57489 | 1038 |
* thm list list |
1039 |
* thm list list |
|
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
1040 |
* (thm list list * Token.src list) |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57999
diff
changeset
|
1041 |
* (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
|
1042 |
|
57260
8747af0d1012
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
blanchet
parents:
57206
diff
changeset
|
1043 |
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
|
1044 |
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
|
1045 |
(corec_selsss, corec_sel_attrs)) = |
54256 | 1046 |
((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
|
1047 |
coinduct_attrs_pair), |
57489 | 1048 |
map (map (Morphism.thm phi)) corecss, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1049 |
map (map (Morphism.thm phi)) corec_discss, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1050 |
(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
|
1051 |
(map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms; |
54256 | 1052 |
|
58115 | 1053 |
val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism; |
54256 | 1054 |
|
58352
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1055 |
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
|
1056 |
(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
|
1057 |
if x = y then [T] else Ts |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1058 |
| unzip_recT _ (Type (@{type_name prod}, Ts as [_, TFree _])) = Ts |
58159 | 1059 |
| unzip_recT _ T = [T]; |
1060 |
||
55867 | 1061 |
fun mk_recs_args_types ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts lthy = |
51832 | 1062 |
let |
52214
4cc5a80bba80
generalized recursors, effectively reverting inductive half of c7a034d01936
blanchet
parents:
52213
diff
changeset
|
1063 |
val Css = map2 replicate ns Cs; |
55869 | 1064 |
val x_Tssss = |
55867 | 1065 |
map6 (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
|
1066 |
map2 (map2 unzip_recT) |
55867 | 1067 |
ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T))) |
1068 |
absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts; |
|
51832 | 1069 |
|
55869 | 1070 |
val x_Tsss' = map (map flat_rec_arg_args) x_Tssss; |
1071 |
val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css; |
|
51832 | 1072 |
|
55869 | 1073 |
val ((fss, xssss), lthy) = |
52214
4cc5a80bba80
generalized recursors, effectively reverting inductive half of c7a034d01936
blanchet
parents:
52213
diff
changeset
|
1074 |
lthy |
55869 | 1075 |
|> mk_Freess "f" f_Tss |
1076 |
||>> mk_Freessss "x" x_Tssss; |
|
51832 | 1077 |
in |
55869 | 1078 |
((f_Tss, x_Tssss, fss, xssss), lthy) |
51832 | 1079 |
end; |
1080 |
||
58159 | 1081 |
fun unzip_corecT (Type (@{type_name sum}, _)) T = [T] |
1082 |
| unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts |
|
1083 |
| unzip_corecT _ T = [T]; |
|
1084 |
||
55867 | 1085 |
(*avoid "'a itself" arguments in corecursors*) |
55966 | 1086 |
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
|
1087 |
| 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
|
1088 |
|
55867 | 1089 |
fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts = |
51829 | 1090 |
let |
55772
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
blanchet
parents:
55701
diff
changeset
|
1091 |
val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; |
55869 | 1092 |
val g_absTs = map range_type fun_Ts; |
1093 |
val g_Tsss = map repair_nullary_single_ctr (map5 dest_absumprodT absTs repTs ns mss g_absTs); |
|
1094 |
val g_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) |
|
1095 |
Cs ctr_Tsss' g_Tsss; |
|
1096 |
val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; |
|
52870 | 1097 |
in |
55869 | 1098 |
(q_Tssss, g_Tsss, g_Tssss, g_absTs) |
52870 | 1099 |
end; |
51829 | 1100 |
|
55867 | 1101 |
fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs; |
52898 | 1102 |
|
55867 | 1103 |
fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec = |
1104 |
(mk_corec_p_pred_types Cs ns, |
|
1105 |
mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss |
|
1106 |
(binder_fun_types (fastype_of dtor_corec))); |
|
51829 | 1107 |
|
55867 | 1108 |
fun mk_corecs_args_types ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts lthy = |
1109 |
let |
|
1110 |
val p_Tss = mk_corec_p_pred_types Cs ns; |
|
51829 | 1111 |
|
55869 | 1112 |
val (q_Tssss, g_Tsss, g_Tssss, corec_types) = |
55867 | 1113 |
mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts; |
51831 | 1114 |
|
55869 | 1115 |
val (((((Free (x, _), cs), pss), qssss), gssss), lthy) = |
51829 | 1116 |
lthy |
55869 | 1117 |
|> 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
|
1118 |
||>> mk_Frees "a" Cs |
51829 | 1119 |
||>> mk_Freess "p" p_Tss |
55869 | 1120 |
||>> mk_Freessss "q" q_Tssss |
1121 |
||>> mk_Freessss "g" g_Tssss; |
|
51831 | 1122 |
|
1123 |
val cpss = map2 (map o rapp) cs pss; |
|
51829 | 1124 |
|
57303 | 1125 |
fun build_sum_inj mk_inj = build_map lthy [] (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
|
1126 |
|
55871 | 1127 |
fun build_dtor_corec_arg _ [] [cg] = cg |
1128 |
| build_dtor_corec_arg T [cq] [cg, cg'] = |
|
1129 |
mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg) |
|
1130 |
(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
|
1131 |
|
55869 | 1132 |
val pgss = map3 flat_corec_preds_predsss_gettersss pss qssss gssss; |
1133 |
val cqssss = map2 (map o map o map o rapp) cs qssss; |
|
1134 |
val cgssss = map2 (map o map o map o rapp) cs gssss; |
|
1135 |
val cqgsss = map3 (map3 (map3 build_dtor_corec_arg)) g_Tsss cqssss cgssss; |
|
51831 | 1136 |
in |
58448 | 1137 |
((x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types)), lthy) |
51831 | 1138 |
end; |
51829 | 1139 |
|
55867 | 1140 |
fun mk_co_recs_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 lthy = |
51903 | 1141 |
let |
52169 | 1142 |
val thy = Proof_Context.theory_of lthy; |
1143 |
||
55867 | 1144 |
val (xtor_co_rec_fun_Ts, xtor_co_recs) = |
1145 |
mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd); |
|
51903 | 1146 |
|
55867 | 1147 |
val ((recs_args_types, corecs_args_types), lthy') = |
52207
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents:
52197
diff
changeset
|
1148 |
if fp = Least_FP then |
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
|
1149 |
mk_recs_args_types ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts lthy |
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
|
1150 |
|>> (rpair NONE o SOME) |
51903 | 1151 |
else |
55867 | 1152 |
mk_corecs_args_types ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts lthy |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1153 |
|>> (pair NONE o SOME); |
51903 | 1154 |
in |
55867 | 1155 |
((xtor_co_recs, recs_args_types, corecs_args_types), lthy') |
51903 | 1156 |
end; |
1157 |
||
55869 | 1158 |
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
|
1159 |
let |
55869 | 1160 |
val n = length cqgss; |
1161 |
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
|
1162 |
in |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1163 |
Term.lambda c (mk_IfN absT cps ts) |
51900 | 1164 |
end; |
51886 | 1165 |
|
58210 | 1166 |
fun define_co_rec_as fp Cs fpT b rhs lthy0 = |
51897 | 1167 |
let |
52170
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
blanchet
parents:
52169
diff
changeset
|
1168 |
val thy = Proof_Context.theory_of lthy0; |
564be617ae84
generalized "mk_co_iter" to handle mutualized (co)iterators
blanchet
parents:
52169
diff
changeset
|
1169 |
|
53569
b4db0ade27bd
conceal definitions of high-level constructors and (co)iterators
traytel
parents:
53553
diff
changeset
|
1170 |
val maybe_conceal_def_binding = Thm.def_binding |
58208 | 1171 |
#> not (Config.get lthy0 bnf_note_all) ? Binding.conceal; |
53569
b4db0ade27bd
conceal definitions of high-level constructors and (co)iterators
traytel
parents:
53553
diff
changeset
|
1172 |
|
55864 | 1173 |
val ((cst, (_, def)), (lthy', lthy)) = lthy0 |
1174 |
|> Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs)) |
|
52327 | 1175 |
||> `Local_Theory.restore; |
1176 |
||
1177 |
val phi = Proof_Context.export_morphism lthy lthy'; |
|
1178 |
||
58211
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
blanchet
parents:
58210
diff
changeset
|
1179 |
val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst); |
55864 | 1180 |
val def' = Morphism.thm phi def; |
52327 | 1181 |
in |
55864 | 1182 |
((cst', def'), lthy') |
52327 | 1183 |
end; |
1184 |
||
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
|
1185 |
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
|
1186 |
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
|
1187 |
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
|
1188 |
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
|
1189 |
|>> 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
|
1190 |
in |
58210 | 1191 |
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
|
1192 |
(fold_rev (fold_rev Term.lambda) fss (Term.list_comb (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
|
1193 |
map4 (fn ctor_rec_absT => fn rep => fn fs => fn 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
|
1194 |
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
|
1195 |
(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
|
1196 |
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
|
1197 |
end; |
51897 | 1198 |
|
58448 | 1199 |
fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = |
51897 | 1200 |
let |
51899 | 1201 |
val nn = length fpTs; |
55867 | 1202 |
val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); |
51897 | 1203 |
in |
58210 | 1204 |
define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN) |
55869 | 1205 |
(fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec, |
1206 |
map5 mk_preds_getterss_join cs cpss f_absTs abss cqgsss))) |
|
52320 | 1207 |
end; |
51897 | 1208 |
|
57471 | 1209 |
fun postproc_co_induct lthy nn prop prop_conj = |
1210 |
Drule.zero_var_indexes |
|
1211 |
#> `(conj_dests nn) |
|
58472 | 1212 |
#>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop)) |
1213 |
##> (fn thm => Thm.permute_prems 0 (~ nn) |
|
57471 | 1214 |
(if nn = 1 then thm RS prop |
57932 | 1215 |
else funpow nn (fn thm => unfold_thms lthy @{thms conj_assoc} (thm RS prop_conj)) thm)); |
57471 | 1216 |
|
1217 |
fun mk_induct_attrs ctrss = |
|
1218 |
let |
|
1219 |
val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss); |
|
1220 |
val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases)); |
|
1221 |
in |
|
1222 |
[induct_case_names_attr] |
|
1223 |
end; |
|
1224 |
||
57535 | 1225 |
fun derive_rel_induct_thms_for_types lthy fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct |
1226 |
ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs = |
|
57471 | 1227 |
let |
58446 | 1228 |
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); |
1229 |
val B_ify = Term.subst_atomic_types (As ~~ Bs); |
|
1230 |
||
1231 |
val fpB_Ts = map B_ify_T fpA_Ts; |
|
1232 |
val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss; |
|
1233 |
val ctrBss = map (map B_ify) ctrAss; |
|
57471 | 1234 |
|
1235 |
val ((((Rs, IRs), ctrAsss), ctrBsss), names_lthy) = lthy |
|
1236 |
|> mk_Frees "R" (map2 mk_pred2T As Bs) |
|
1237 |
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) |
|
1238 |
||>> mk_Freesss "a" ctrAs_Tsss |
|
1239 |
||>> mk_Freesss "b" ctrBs_Tsss; |
|
1240 |
||
58291 | 1241 |
val prems = |
57471 | 1242 |
let |
1243 |
fun mk_prem ctrA ctrB argAs argBs = |
|
1244 |
fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies) |
|
57565 | 1245 |
(map2 (HOLogic.mk_Trueprop oo build_rel_app names_lthy (Rs @ IRs) fpA_Ts) argAs argBs) |
1246 |
(HOLogic.mk_Trueprop (build_rel_app names_lthy (Rs @ IRs) fpA_Ts |
|
1247 |
(Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs))))); |
|
57471 | 1248 |
in |
1249 |
flat (map4 (map4 mk_prem) ctrAss ctrBss ctrAsss ctrBsss) |
|
1250 |
end; |
|
1251 |
||
57525 | 1252 |
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq |
58104 | 1253 |
(map2 (build_the_rel [] lthy (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs)); |
57471 | 1254 |
|
57898
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1255 |
val rel_induct0_thm = |
58291 | 1256 |
Goal.prove_sorry lthy [] prems goal (fn {context = ctxt, prems} => |
57898
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1257 |
mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (certify ctxt) IRs) exhausts ctor_defss |
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1258 |
ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |
57471 | 1259 |
|> singleton (Proof_Context.export names_lthy lthy) |
1260 |
|> Thm.close_derivation; |
|
1261 |
in |
|
1262 |
(postproc_co_induct lthy (length fpA_Ts) @{thm predicate2D} @{thm predicate2D_conj} |
|
1263 |
rel_induct0_thm, |
|
1264 |
mk_induct_attrs ctrAss) |
|
1265 |
end; |
|
1266 |
||
58335
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58332
diff
changeset
|
1267 |
fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms |
57397 | 1268 |
live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses fp_type_definitions |
1269 |
abs_inverses ctrss ctr_defss recs rec_defs lthy = |
|
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1270 |
let |
51827 | 1271 |
val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss; |
1272 |
||
51815 | 1273 |
val nn = length pre_bnfs; |
51827 | 1274 |
val ns = map length ctr_Tsss; |
1275 |
val mss = map (map length) ctr_Tsss; |
|
51815 | 1276 |
|
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1277 |
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
|
1278 |
val pre_set_defss = map set_defs_of_bnf pre_bnfs; |
57399 | 1279 |
val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; |
1280 |
val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs; |
|
57397 | 1281 |
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
|
1282 |
|
51816 | 1283 |
val fp_b_names = map base_name_of_typ fpTs; |
51811 | 1284 |
|
51832 | 1285 |
val ((((ps, ps'), xsss), us'), names_lthy) = |
52315
fafab8eac3ee
avoid duplicate call to "mk_fold_rec_args_types" function
blanchet
parents:
52314
diff
changeset
|
1286 |
lthy |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1287 |
|> mk_Frees' "P" (map mk_pred1T fpTs) |
51827 | 1288 |
||>> mk_Freesss "x" ctr_Tsss |
51816 | 1289 |
||>> 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
|
1290 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1291 |
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
|
1292 |
|
57665 | 1293 |
fun mk_sets bnf = |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1294 |
let |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1295 |
val Type (T_name, Us) = T_of_bnf bnf; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1296 |
val lives = lives_of_bnf bnf; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1297 |
val sets = sets_of_bnf bnf; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1298 |
fun mk_set U = |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1299 |
(case find_index (curry (op =) U) lives of |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1300 |
~1 => Term.dummy |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1301 |
| i => nth sets i); |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1302 |
in |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1303 |
(T_name, map mk_set Us) |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1304 |
end; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1305 |
|
57665 | 1306 |
val setss_fp_nesting = map mk_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
|
1307 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1308 |
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
|
1309 |
let |
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
|
1310 |
fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) = |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1311 |
[([], (find_index (curry (op =) X) Xs + 1, x))] |
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
|
1312 |
| mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) = |
57665 | 1313 |
(case AList.lookup (op =) setss_fp_nesting T_name of |
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
|
1314 |
NONE => [] |
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
|
1315 |
| SOME raw_sets0 => |
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
|
1316 |
let |
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
|
1317 |
val (Xs_Ts, (Ts, raw_sets)) = |
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
|
1318 |
filter (exists_subtype_in Xs o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0)) |
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
|
1319 |
|> split_list ||> split_list; |
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
|
1320 |
val sets = map (mk_set Ts0) raw_sets; |
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
|
1321 |
val (ys, names_lthy') = names_lthy |> mk_Frees s Ts; |
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
|
1322 |
val xysets = map (pair x) (ys ~~ sets); |
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
|
1323 |
val ppremss = map2 (mk_raw_prem_prems names_lthy') ys Xs_Ts; |
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
|
1324 |
in |
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
|
1325 |
flat (map2 (map o apfst o cons) xysets ppremss) |
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
|
1326 |
end) |
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
|
1327 |
| mk_raw_prem_prems _ _ _ = []; |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1328 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1329 |
fun close_prem_prem xs t = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1330 |
fold_rev Logic.all (map Free (drop (nn + length xs) |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1331 |
(rev (Term.add_frees t (map dest_Free xs @ ps'))))) t; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1332 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1333 |
fun mk_prem_prem xs (xysets, (j, x)) = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1334 |
close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) => |
57567 | 1335 |
mk_Trueprop_mem (y, set $ x')) xysets, |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1336 |
HOLogic.mk_Trueprop (nth ps (j - 1) $ x))); |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1337 |
|
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
|
1338 |
fun mk_raw_prem phi ctr ctr_Ts ctrXs_Ts = |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1339 |
let |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1340 |
val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts; |
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
|
1341 |
val pprems = flat (map2 (mk_raw_prem_prems names_lthy') xs ctrXs_Ts); |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1342 |
in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1343 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1344 |
fun mk_prem (xs, raw_pprems, concl) = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1345 |
fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl)); |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1346 |
|
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
|
1347 |
val raw_premss = map4 (map3 o mk_raw_prem) ps ctrss ctr_Tsss ctrXs_Tsss; |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1348 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1349 |
val goal = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1350 |
Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss, |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1351 |
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
|
1352 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1353 |
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
|
1354 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1355 |
val ctor_induct' = ctor_induct OF (map2 mk_absumprodE fp_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
|
1356 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1357 |
val thm = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1358 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1359 |
mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' fp_abs_inverses |
57397 | 1360 |
abs_inverses fp_nesting_set_maps pre_set_defss) |
57633 | 1361 |
|> singleton (Proof_Context.export names_lthy lthy); |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1362 |
in |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1363 |
`(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
|
1364 |
end; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1365 |
|
52305 | 1366 |
val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss; |
1367 |
||
55867 | 1368 |
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
|
1369 |
let |
55867 | 1370 |
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
|
1371 |
|
55871 | 1372 |
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
|
1373 |
fold_rev (fold_rev Logic.all) (xs :: fss) |
55867 | 1374 |
(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
|
1375 |
|
52302 | 1376 |
fun maybe_tick (T, U) u f = |
1377 |
if try (fst o HOLogic.dest_prodT) U = SOME T then |
|
1378 |
Term.lambda u (HOLogic.mk_prod (u, f $ u)) |
|
1379 |
else |
|
1380 |
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
|
1381 |
|
55867 | 1382 |
fun build_rec (x as Free (_, T)) U = |
52368 | 1383 |
if T = U then |
1384 |
x |
|
1385 |
else |
|
58352
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1386 |
let |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1387 |
val build_simple = |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1388 |
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
|
1389 |
(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
|
1390 |
in |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1391 |
build_map lthy [] build_simple (T, U) $ x |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1392 |
end; |
52301
7935e82a4ae4
simpler, more robust iterator goal construction code
blanchet
parents:
52300
diff
changeset
|
1393 |
|
55867 | 1394 |
val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss; |
55871 | 1395 |
val goalss = map5 (map4 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
|
1396 |
|
57397 | 1397 |
val tacss = map4 (map ooo |
57399 | 1398 |
mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs) |
55867 | 1399 |
ctor_rec_thms fp_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
|
1400 |
|
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1401 |
fun prove goal tac = |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1402 |
Goal.prove_sorry lthy [] [] goal (tac o #context) |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1403 |
|> 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
|
1404 |
in |
52305 | 1405 |
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
|
1406 |
end; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1407 |
|
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
|
1408 |
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
|
1409 |
|
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58332
diff
changeset
|
1410 |
val code_attrs = if plugins code_plugin then [Code.add_default_eqn_attrib] else []; |
51808
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1411 |
in |
57471 | 1412 |
((induct_thms, induct_thm, mk_induct_attrs ctrss), |
58335
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58332
diff
changeset
|
1413 |
(rec_thmss, code_attrs @ 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
|
1414 |
end; |
355dcd6a9b3c
factored out derivation of induction principles, folds and recs, in preparation for reduction of nested to mutual
blanchet
parents:
51805
diff
changeset
|
1415 |
|
58283
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
blanchet
parents:
58267
diff
changeset
|
1416 |
fun mk_coinduct_attrs fpTs ctrss discss mss = |
57302 | 1417 |
let |
1418 |
val nn = length fpTs; |
|
1419 |
val fp_b_names = map base_name_of_typ fpTs; |
|
57535 | 1420 |
|
57302 | 1421 |
fun mk_coinduct_concls ms discs ctrs = |
1422 |
let |
|
1423 |
fun mk_disc_concl disc = [name_of_disc disc]; |
|
1424 |
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
|
1425 |
| mk_ctr_concl _ ctr = [name_of_ctr ctr]; |
57302 | 1426 |
val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]]; |
1427 |
val ctr_concls = map2 mk_ctr_concl ms ctrs; |
|
1428 |
in |
|
1429 |
flat (map2 append disc_concls ctr_concls) |
|
1430 |
end; |
|
57535 | 1431 |
|
57302 | 1432 |
val coinduct_cases = quasi_unambiguous_case_names (map (prefix EqN) fp_b_names); |
1433 |
val coinduct_conclss = |
|
1434 |
map3 (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss; |
|
1435 |
||
1436 |
val common_coinduct_consumes_attr = Attrib.internal (K (Rule_Cases.consumes nn)); |
|
1437 |
val coinduct_consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1)); |
|
1438 |
||
1439 |
val coinduct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names coinduct_cases)); |
|
1440 |
val coinduct_case_concl_attrs = |
|
1441 |
map2 (fn casex => fn concls => |
|
1442 |
Attrib.internal (K (Rule_Cases.case_conclusion (casex, concls)))) |
|
1443 |
coinduct_cases coinduct_conclss; |
|
1444 |
||
1445 |
val common_coinduct_attrs = |
|
1446 |
common_coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs; |
|
1447 |
val coinduct_attrs = |
|
1448 |
coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs; |
|
1449 |
in |
|
1450 |
(coinduct_attrs, common_coinduct_attrs) |
|
1451 |
end; |
|
1452 |
||
57535 | 1453 |
fun derive_rel_coinduct_thm_for_types lthy fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list) |
1454 |
abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct |
|
1455 |
live_nesting_rel_eqs = |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1456 |
let |
58446 | 1457 |
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); |
1458 |
val fpB_Ts = map B_ify_T fpA_Ts; |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1459 |
|
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1460 |
val (Rs, IRs, fpAs, fpBs, names_lthy) = |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1461 |
let |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1462 |
val fp_names = map base_name_of_typ fpA_Ts; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1463 |
val ((((Rs, IRs), fpAs_names), fpBs_names), names_lthy) = lthy |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1464 |
|> mk_Frees "R" (map2 mk_pred2T As Bs) |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1465 |
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1466 |
||>> Variable.variant_fixes fp_names |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1467 |
||>> Variable.variant_fixes (map (suffix "'") fp_names); |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1468 |
in |
58291 | 1469 |
(Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts, |
1470 |
names_lthy) |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1471 |
end; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1472 |
|
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1473 |
val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) = |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1474 |
let |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1475 |
val discss = map #discs ctr_sugars; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1476 |
val selsss = map #selss ctr_sugars; |
58291 | 1477 |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1478 |
fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss); |
58291 | 1479 |
fun mk_selsss ts Ts = |
1480 |
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
|
1481 |
in |
57302 | 1482 |
((mk_discss fpAs As, mk_selsss fpAs As), |
1483 |
(mk_discss fpBs Bs, mk_selsss fpBs Bs)) |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1484 |
end; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1485 |
|
58291 | 1486 |
val prems = |
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1487 |
let |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1488 |
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
|
1489 |
(if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @ |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1490 |
(case (selA_ts, selB_ts) of |
57303 | 1491 |
([], []) => [] |
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1492 |
| (_ :: _, _ :: _) => |
57565 | 1493 |
[Library.foldr HOLogic.mk_imp |
1494 |
(if n = 1 then [] else [discA_t, discB_t], |
|
1495 |
Library.foldr1 HOLogic.mk_conj |
|
1496 |
(map2 (build_rel_app lthy (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]); |
|
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1497 |
|
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1498 |
fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss = |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1499 |
Library.foldr1 HOLogic.mk_conj (flat (map5 (mk_prem_ctr_concls n) |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1500 |
(1 upto n) discA_ts selA_tss discB_ts selB_tss)) |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1501 |
handle List.Empty => @{term True}; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1502 |
|
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1503 |
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
|
1504 |
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
|
1505 |
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
|
1506 |
in |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1507 |
map8 mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1508 |
end; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1509 |
|
57525 | 1510 |
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq |
58104 | 1511 |
IRs (map2 (build_the_rel [] lthy (Rs @ IRs) []) fpA_Ts fpB_Ts))); |
57302 | 1512 |
|
57898
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1513 |
val rel_coinduct0_thm = |
58291 | 1514 |
Goal.prove_sorry lthy [] prems goal (fn {context = ctxt, prems} => |
57898
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1515 |
mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (certify ctxt) IRs) prems |
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1516 |
(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
|
1517 |
(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
|
1518 |
rel_pre_defs abs_inverses live_nesting_rel_eqs) |
57301
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1519 |
|> singleton (Proof_Context.export names_lthy lthy) |
57302 | 1520 |
|> Thm.close_derivation; |
1521 |
in |
|
57471 | 1522 |
(postproc_co_induct lthy (length fpA_Ts) @{thm predicate2D} @{thm predicate2D_conj} |
1523 |
rel_coinduct0_thm, |
|
58283
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
blanchet
parents:
58267
diff
changeset
|
1524 |
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
|
1525 |
end; |
7b997028aaac
generate 'rel_coinduct0' theorem for codatatypes
desharna
parents:
57279
diff
changeset
|
1526 |
|
57700 | 1527 |
fun derive_set_induct_thms_for_types lthy nn fpTs ctrss setss dtor_set_inducts exhausts |
1528 |
set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = |
|
1529 |
let |
|
1530 |
fun mk_prems A Ps ctr_args t ctxt = |
|
1531 |
(case fastype_of t of |
|
57891 | 1532 |
Type (type_name, innerTs) => |
57700 | 1533 |
(case bnf_of ctxt type_name of |
1534 |
NONE => ([], ctxt) |
|
1535 |
| SOME bnf => |
|
1536 |
let |
|
1537 |
fun seq_assm a set ctxt = |
|
1538 |
let |
|
1539 |
val X = HOLogic.dest_setT (range_type (fastype_of set)); |
|
1540 |
val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt; |
|
1541 |
val assm = mk_Trueprop_mem (x, set $ a); |
|
1542 |
in |
|
1543 |
(case build_binary_fun_app Ps x a of |
|
1544 |
NONE => |
|
1545 |
mk_prems A Ps ctr_args x ctxt' |
|
1546 |
|>> map (Logic.all x o Logic.mk_implies o pair assm) |
|
1547 |
| SOME f => |
|
1548 |
([Logic.all x |
|
1549 |
(Logic.mk_implies (assm, |
|
1550 |
Logic.mk_implies (HOLogic.mk_Trueprop f, |
|
1551 |
HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))], |
|
1552 |
ctxt')) |
|
1553 |
end; |
|
1554 |
in |
|
57891 | 1555 |
fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |
57700 | 1556 |
|>> flat |
1557 |
end) |
|
1558 |
| T => |
|
1559 |
if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt) |
|
1560 |
else ([], ctxt)); |
|
1561 |
||
1562 |
fun mk_prems_for_ctr A Ps ctr ctxt = |
|
1563 |
let |
|
57893 | 1564 |
val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt; |
57700 | 1565 |
in |
1566 |
fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt' |
|
1567 |
|>> map (fold_rev Logic.all args) o flat |
|
1568 |
|>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr))) |
|
1569 |
end; |
|
1570 |
||
1571 |
fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt = |
|
1572 |
let |
|
1573 |
val ((x, fp), ctxt') = ctxt |
|
1574 |
|> yield_singleton (mk_Frees "x") A |
|
1575 |
||>> yield_singleton (mk_Frees "a") fpT; |
|
1576 |
val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x) |
|
1577 |
(the (build_binary_fun_app Ps x fp))); |
|
1578 |
in |
|
1579 |
fold_map (mk_prems_for_ctr A Ps) ctrs ctxt' |
|
1580 |
|>> split_list |
|
1581 |
|>> map_prod flat flat |
|
1582 |
|>> apfst (rpair concl) |
|
1583 |
end; |
|
1584 |
||
1585 |
fun mk_thm ctxt fpTs ctrss sets = |
|
1586 |
let |
|
1587 |
val A = HOLogic.dest_setT (range_type (fastype_of (hd sets))); |
|
1588 |
val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt; |
|
58291 | 1589 |
val (((prems, concl), case_names), ctxt'') = |
1590 |
fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt' |
|
57700 | 1591 |
|>> apfst split_list o split_list |
1592 |
|>> apfst (apfst flat) |
|
1593 |
|>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj)) |
|
1594 |
|>> apsnd flat; |
|
1595 |
||
57898
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1596 |
val thm = |
58291 | 1597 |
Goal.prove_sorry lthy [] 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
|
1598 |
(fn {context = ctxt, prems} => |
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1599 |
mk_set_induct0_tac ctxt (map (certify ctxt'') Ps) prems dtor_set_inducts exhausts |
f7f75b33d6c8
avoid needless (and wrong w.r.t. sorts) generation of type variables; tuned whitespaces;
blanchet
parents:
57893
diff
changeset
|
1600 |
set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |
57700 | 1601 |
|> singleton (Proof_Context.export ctxt'' ctxt) |
1602 |
|> Thm.close_derivation; |
|
1603 |
||
1604 |
val case_names_attr = |
|
1605 |
Attrib.internal (K (Rule_Cases.case_names (quasi_unambiguous_case_names case_names))); |
|
1606 |
val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets; |
|
1607 |
in |
|
1608 |
(thm, case_names_attr :: induct_set_attrs) |
|
1609 |
end |
|
1610 |
val consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1)); |
|
1611 |
in |
|
58472 | 1612 |
map (mk_thm lthy fpTs ctrss |
1613 |
#> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr)) |
|
1614 |
(transpose setss) |
|
57700 | 1615 |
end; |
1616 |
||
58448 | 1617 |
fun derive_coinduct_corecs_thms_for_types pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) |
57397 | 1618 |
dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss |
1619 |
kss mss ns fp_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) |
|
55864 | 1620 |
corecs corec_defs export_args lthy = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1621 |
let |
55867 | 1622 |
fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec = |
1623 |
iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]]; |
|
53203
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
traytel
parents:
53143
diff
changeset
|
1624 |
|
55867 | 1625 |
val ctor_dtor_corec_thms = map3 mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms; |
53203
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
traytel
parents:
53143
diff
changeset
|
1626 |
|
51815 | 1627 |
val nn = length pre_bnfs; |
1628 |
||
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1629 |
val pre_map_defs = map map_def_of_bnf pre_bnfs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1630 |
val pre_rel_defs = map rel_def_of_bnf pre_bnfs; |
57399 | 1631 |
val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; |
57397 | 1632 |
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
|
1633 |
|
51816 | 1634 |
val fp_b_names = map base_name_of_typ fpTs; |
51811 | 1635 |
|
53210
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
blanchet
parents:
53203
diff
changeset
|
1636 |
val ctrss = map #ctrs ctr_sugars; |
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
blanchet
parents:
53203
diff
changeset
|
1637 |
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
|
1638 |
val selsss = map #selss ctr_sugars; |
51840 | 1639 |
val exhausts = map #exhaust ctr_sugars; |
1640 |
val disc_thmsss = map #disc_thmss ctr_sugars; |
|
1641 |
val discIss = map #discIs ctr_sugars; |
|
1642 |
val sel_thmsss = map #sel_thmss ctr_sugars; |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1643 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1644 |
val (((rs, us'), vs'), names_lthy) = |
52341 | 1645 |
lthy |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1646 |
|> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs) |
51816 | 1647 |
||>> Variable.variant_fixes fp_b_names |
1648 |
||>> Variable.variant_fixes (map (suffix "'") fp_b_names); |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1649 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1650 |
val us = map2 (curry Free) us' fpTs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1651 |
val udiscss = map2 (map o rapp) us discss; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1652 |
val uselsss = map2 (map o map o rapp) us selsss; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1653 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1654 |
val vs = map2 (curry Free) vs' fpTs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1655 |
val vdiscss = map2 (map o rapp) vs discss; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1656 |
val vselsss = map2 (map o map o rapp) vs selsss; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1657 |
|
52345 | 1658 |
val coinduct_thms_pairs = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1659 |
let |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1660 |
val uvrs = map3 (fn r => fn u => fn v => r $ u $ v) rs us vs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1661 |
val uv_eqs = map2 (curry HOLogic.mk_eq) us vs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1662 |
val strong_rs = |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1663 |
map4 (fn u => fn v => fn uvr => fn uv_eq => |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1664 |
fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1665 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1666 |
fun build_the_rel rs' T Xs_T = |
58104 | 1667 |
build_rel [] lthy [] (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T) |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1668 |
|> Term.subst_atomic_types (Xs ~~ fpTs); |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1669 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1670 |
fun build_rel_app rs' usel vsel Xs_T = |
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1671 |
fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T); |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1672 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1673 |
fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1674 |
(if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @ |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1675 |
(if null usels then |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1676 |
[] |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1677 |
else |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1678 |
[Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1679 |
Library.foldr1 HOLogic.mk_conj (map3 (build_rel_app rs') usels vsels ctrXs_Ts))]); |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1680 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1681 |
fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss = |
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1682 |
Library.foldr1 HOLogic.mk_conj (flat (map6 (mk_prem_ctr_concls rs' n) |
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1683 |
(1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss)) |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1684 |
handle List.Empty => @{term True}; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1685 |
|
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1686 |
fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1687 |
fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr, |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1688 |
HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss))); |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1689 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1690 |
val concl = |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1691 |
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1692 |
(map3 (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v))) |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1693 |
uvrs us vs)); |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1694 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1695 |
fun mk_goal rs' = |
54237
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1696 |
Logic.list_implies (map9 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss |
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
blanchet
parents:
54236
diff
changeset
|
1697 |
ctrXs_Tsss, concl); |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1698 |
|
52345 | 1699 |
val goals = map mk_goal [rs, strong_rs]; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1700 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1701 |
fun prove dtor_coinduct' goal = |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1702 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
57397 | 1703 |
mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs |
1704 |
fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss) |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1705 |
|> singleton (Proof_Context.export names_lthy lthy) |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1706 |
|> Thm.close_derivation; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1707 |
|
53106 | 1708 |
val rel_eqs = map rel_eq_of_bnf pre_bnfs; |
1709 |
val rel_monos = map rel_mono_of_bnf pre_bnfs; |
|
1710 |
val dtor_coinducts = |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1711 |
[dtor_coinduct, mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p lthy] |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1712 |
in |
57471 | 1713 |
map2 (postproc_co_induct lthy nn mp mp_conj oo prove) dtor_coinducts goals |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1714 |
end; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1715 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1716 |
fun mk_maybe_not pos = not pos ? HOLogic.mk_not; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1717 |
|
55869 | 1718 |
val gcorecs = map (lists_bmoc pgss) corecs; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1719 |
|
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
1720 |
val corec_thmss = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1721 |
let |
55869 | 1722 |
fun mk_goal c cps gcorec n k ctr m cfs' = |
1723 |
fold_rev (fold_rev Logic.all) ([c] :: pgss) |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1724 |
(Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps, |
55869 | 1725 |
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
|
1726 |
|
55869 | 1727 |
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
|
1728 |
|
55869 | 1729 |
fun tack (c, u) f = |
1730 |
let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in |
|
1731 |
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
|
1732 |
end; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1733 |
|
55869 | 1734 |
fun build_corec cqg = |
1735 |
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
|
1736 |
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
|
1737 |
let |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1738 |
val U = mk_U T; |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1739 |
val build_simple = |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1740 |
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
|
1741 |
(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
|
1742 |
in |
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
58347
diff
changeset
|
1743 |
build_map lthy [] build_simple (T, U) $ cqg |
52368 | 1744 |
end |
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
1745 |
else |
55869 | 1746 |
cqg |
52347
ead18e3b2c1b
changed back type of corecursor for nested case, effectively reverting aa66ea552357 and 78a3d5006cf1
blanchet
parents:
52346
diff
changeset
|
1747 |
end; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1748 |
|
55869 | 1749 |
val cqgsss' = map (map (map build_corec)) cqgsss; |
1750 |
val goalss = map8 (map4 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
|
1751 |
|
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
1752 |
val tacss = |
57399 | 1753 |
map4 (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s) |
55867 | 1754 |
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
|
1755 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1756 |
fun prove goal tac = |
51815 | 1757 |
Goal.prove_sorry lthy [] [] goal (tac o #context) |
1758 |
|> Thm.close_derivation; |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1759 |
in |
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
1760 |
map2 (map2 prove) goalss tacss |
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
1761 |
|> map (map (unfold_thms lthy @{thms case_sum_if})) |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1762 |
end; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1763 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1764 |
val corec_disc_iff_thmss = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1765 |
let |
55869 | 1766 |
fun mk_goal c cps gcorec n k disc = |
1767 |
mk_Trueprop_eq (disc $ (gcorec $ c), |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1768 |
if n = 1 then @{const True} |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1769 |
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
|
1770 |
|
55869 | 1771 |
val goalss = map6 (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1772 |
|
51828 | 1773 |
fun mk_case_split' cp = Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split}; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1774 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1775 |
val case_splitss' = map (map mk_case_split') cpss; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1776 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1777 |
val tacss = map3 (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
|
1778 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1779 |
fun prove goal tac = |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1780 |
Goal.prove_sorry lthy [] [] goal (tac o #context) |
52357
a5d3730043c2
use right context when exporting variables (cf. AFP Coinductive_List failures)
blanchet
parents:
52356
diff
changeset
|
1781 |
|> singleton export_args |
51829 | 1782 |
|> singleton (Proof_Context.export names_lthy lthy) |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1783 |
|> Thm.close_derivation; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1784 |
|
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1785 |
fun proves [_] [_] = [] |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1786 |
| proves goals tacs = map2 prove goals tacs; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1787 |
in |
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
1788 |
map2 proves goalss tacss |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1789 |
end; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1790 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1791 |
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
|
1792 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1793 |
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
|
1794 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1795 |
fun mk_corec_sel_thm corec_thm sel sel_thm = |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1796 |
let |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1797 |
val (domT, ranT) = dest_funT (fastype_of sel); |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1798 |
val arg_cong' = |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1799 |
Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT]) |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1800 |
[NONE, NONE, SOME (certify lthy sel)] arg_cong |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1801 |
|> Thm.varifyT_global; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1802 |
val sel_thm' = sel_thm RSN (2, trans); |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1803 |
in |
55867 | 1804 |
corec_thm RS arg_cong' RS sel_thm' |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1805 |
end; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1806 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1807 |
fun mk_corec_sel_thms corec_thmss = |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1808 |
map3 (map3 (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
|
1809 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1810 |
val corec_sel_thmsss = mk_corec_sel_thms corec_thmss; |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1811 |
in |
57535 | 1812 |
((coinduct_thms_pairs, |
58283
71d74e641538
preserve case names in '(co)induct' theorems generated by prim(co)rec'
blanchet
parents:
58267
diff
changeset
|
1813 |
mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss), |
57489 | 1814 |
corec_thmss, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1815 |
corec_disc_thmss, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1816 |
(corec_disc_iff_thmss, simp_attrs), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
1817 |
(corec_sel_thmsss, simp_attrs)) |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1818 |
end; |
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
1819 |
|
52207
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents:
52197
diff
changeset
|
1820 |
fun define_co_datatypes prepare_constraint prepare_typ prepare_term fp construct_fp |
58434 | 1821 |
((plugins, discs_sels0), specs) no_defs_lthy = |
49112 | 1822 |
let |
49298
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1823 |
(* TODO: sanity checks on arguments *) |
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1824 |
|
57094
589ec121ce1a
don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents:
57091
diff
changeset
|
1825 |
val discs_sels = discs_sels0 orelse fp = Greatest_FP; |
49278 | 1826 |
|
49367 | 1827 |
val nn = length specs; |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1828 |
val fp_bs = map type_binding_of_spec specs; |
49498 | 1829 |
val fp_b_names = map Binding.name_of fp_bs; |
1830 |
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
|
1831 |
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
|
1832 |
val rel_bs = map rel_binding_of_spec specs; |
49298
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1833 |
|
51758 | 1834 |
fun prepare_type_arg (_, (ty, c)) = |
58434 | 1835 |
let val TFree (s, _) = prepare_typ no_defs_lthy ty in |
1836 |
TFree (s, prepare_constraint no_defs_lthy c) |
|
49298
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1837 |
end; |
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1838 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1839 |
val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs; |
58234 | 1840 |
val unsorted_Ass0 = map (map (resort_tfree_or_tvar @{sort type})) Ass0; |
55700 | 1841 |
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
|
1842 |
val num_As = length unsorted_As; |
55699
1f9cc432ecd4
reuse same parser for '(co)datatype(_new)' as for 'bnf_decl'
blanchet
parents:
55575
diff
changeset
|
1843 |
|
1f9cc432ecd4
reuse same parser for '(co)datatype(_new)' as for 'bnf_decl'
blanchet
parents:
55575
diff
changeset
|
1844 |
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
|
1845 |
val set_bss = map (map (the_default Binding.empty)) set_boss; |
49119 | 1846 |
|
58446 | 1847 |
val ((((Bs0, Cs), Es), Xs), names_no_defs_lthy) = |
58434 | 1848 |
no_defs_lthy |
58234 | 1849 |
|> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As |
53266
89f7f1cc9ae3
cleaner handling of bootstrapping "fake" context, with fewer (no?) obscure bugs
blanchet
parents:
53264
diff
changeset
|
1850 |
|> mk_TFrees num_As |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1851 |
||>> mk_TFrees nn |
58446 | 1852 |
||>> mk_TFrees nn |
51858
7a08fe1e19b1
added and moved library functions (used in primrec code)
blanchet
parents:
51857
diff
changeset
|
1853 |
||>> variant_tfrees fp_b_names; |
49119 | 1854 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1855 |
fun add_fake_type spec = |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1856 |
Typedecl.basic_typedecl (type_binding_of_spec spec, num_As, mixfix_of_spec spec); |
51768 | 1857 |
|
58434 | 1858 |
val (fake_T_names, fake_lthy) = fold_map add_fake_type specs no_defs_lthy; |
49119 | 1859 |
|
53262
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1860 |
val qsoty = quote o Syntax.string_of_typ fake_lthy; |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1861 |
|
54253 | 1862 |
val _ = (case Library.duplicates (op =) unsorted_As of [] => () |
53262
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1863 |
| A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^ |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1864 |
"datatype specification")); |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1865 |
|
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1866 |
val bad_args = |
58434 | 1867 |
map (Logic.type_map (singleton (Variable.polymorphic no_defs_lthy))) unsorted_As |
53262
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1868 |
|> filter_out Term.is_TVar; |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1869 |
val _ = null bad_args orelse |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1870 |
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
|
1871 |
"datatype specification"); |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1872 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1873 |
val mixfixes = map mixfix_of_spec specs; |
49119 | 1874 |
|
54253 | 1875 |
val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => () |
49119 | 1876 |
| b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b))); |
1877 |
||
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1878 |
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
|
1879 |
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
|
1880 |
val ctr_mixfixess = map (map snd) mx_ctr_specss; |
49119 | 1881 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1882 |
val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss; |
49336 | 1883 |
val ctr_bindingss = |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1884 |
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
|
1885 |
ctr_specss; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1886 |
val ctr_argsss = map (map args_of_ctr_spec) ctr_specss; |
49119 | 1887 |
|
49336 | 1888 |
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
|
1889 |
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
|
1890 |
val raw_sel_default_eqss = map sel_default_eqs_of_spec specs; |
49286 | 1891 |
|
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
49302
diff
changeset
|
1892 |
val (As :: _) :: fake_ctr_Tsss = |
49298
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1893 |
burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0); |
53262
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1894 |
val As' = map dest_TFree As; |
49298
36e551d3af3b
support for sort constraints in new (co)data commands
blanchet
parents:
49297
diff
changeset
|
1895 |
|
49183
0cc46e2dee7e
careful about constructor types w.r.t. fake context (third step of localization)
blanchet
parents:
49182
diff
changeset
|
1896 |
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
|
1897 |
val _ = (case subtract (op =) As' rhs_As' of [] => () |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1898 |
| extras => error ("Extra type variables on right-hand side: " ^ |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1899 |
commas (map (qsoty o TFree) extras))); |
49165 | 1900 |
|
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53266
diff
changeset
|
1901 |
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
|
1902 |
|
53262
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1903 |
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
|
1904 |
s = s' andalso (Ts = Ts' orelse |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1905 |
error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^ |
23927b18dce2
rationalize message generation + added a warning
blanchet
parents:
53260
diff
changeset
|
1906 |
" (expected " ^ qsoty T' ^ ")")) |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
1907 |
| eq_fpT_check _ _ = false; |
49146 | 1908 |
|
53222 | 1909 |
fun freeze_fp (T as Type (s, Ts)) = |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
1910 |
(case find_index (eq_fpT_check T) fake_Ts of |
53222 | 1911 |
~1 => Type (s, map freeze_fp Ts) |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
1912 |
| kk => nth Xs kk) |
49204 | 1913 |
| freeze_fp T = T; |
49121 | 1914 |
|
53222 | 1915 |
val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts); |
1916 |
||
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
|
1917 |
val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss; |
55966 | 1918 |
val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; |
49119 | 1919 |
|
58507 | 1920 |
val fp_eqs = map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctrXs_repTs; |
49121 | 1921 |
|
55701 | 1922 |
val killed_As = |
1923 |
map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE) |
|
1924 |
(unsorted_As ~~ transpose set_boss); |
|
1925 |
||
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1926 |
val ((pre_bnfs, absT_infos), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, |
55868 | 1927 |
dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1928 |
ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms, |
58446 | 1929 |
xtor_co_rec_thms, rel_xtor_co_induct_thm, dtor_set_induct_thms, |
58448 | 1930 |
xtor_co_rec_transfer_thms, ...}, |
53203
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
traytel
parents:
53143
diff
changeset
|
1931 |
lthy)) = |
55701 | 1932 |
fp_bnf (construct_fp mixfixes map_bs rel_bs set_bss) fp_bs (map dest_TFree unsorted_As) |
58434 | 1933 |
(map dest_TFree killed_As) fp_eqs no_defs_lthy |
53222 | 1934 |
handle BAD_DEAD (X, X_backdrop) => |
1935 |
(case X_backdrop of |
|
1936 |
Type (bad_tc, _) => |
|
1937 |
let |
|
1938 |
val fake_T = qsoty (unfreeze_fp X); |
|
1939 |
val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop); |
|
1940 |
fun register_hint () = |
|
1941 |
"\nUse the " ^ quote (fst (fst @{command_spec "bnf"})) ^ " command to register " ^ |
|
1942 |
quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ |
|
1943 |
\it"; |
|
1944 |
in |
|
1945 |
if is_some (bnf_of no_defs_lthy bad_tc) orelse |
|
1946 |
is_some (fp_sugar_of no_defs_lthy bad_tc) then |
|
53256 | 1947 |
error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ |
1948 |
" in type expression " ^ fake_T_backdrop) |
|
58112
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents:
58104
diff
changeset
|
1949 |
else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of no_defs_lthy) |
53222 | 1950 |
bad_tc) then |
1951 |
error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ |
|
1952 |
" via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^ |
|
1953 |
fake_T_backdrop ^ register_hint ()) |
|
1954 |
else |
|
1955 |
error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ |
|
1956 |
" via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ |
|
1957 |
register_hint ()) |
|
1958 |
end); |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1959 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1960 |
val abss = map #abs absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1961 |
val reps = map #rep absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1962 |
val absTs = map #absT absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1963 |
val repTs = map #repT absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1964 |
val abs_injects = map #abs_inject absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1965 |
val abs_inverses = map #abs_inverse absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1966 |
val type_definitions = map #type_definition absT_infos; |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
1967 |
|
53143
ba80154a1118
configuration option to control timing output for (co)datatypes
traytel
parents:
53126
diff
changeset
|
1968 |
val time = time lthy; |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1969 |
val timer = time (Timer.startRealTimer ()); |
49121 | 1970 |
|
57397 | 1971 |
val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; |
1972 |
val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; |
|
49226 | 1973 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1974 |
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
|
1975 |
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
|
1976 |
val pre_rel_defs = map rel_def_of_bnf pre_bnfs; |
57397 | 1977 |
val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; |
58347 | 1978 |
val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs; |
57397 | 1979 |
val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs; |
58328 | 1980 |
val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1981 |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1982 |
val live = live_of_bnf any_fp_bnf; |
52961 | 1983 |
val _ = |
1984 |
if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs) then |
|
53266
89f7f1cc9ae3
cleaner handling of bootstrapping "fake" context, with fewer (no?) obscure bugs
blanchet
parents:
53264
diff
changeset
|
1985 |
warning "Map function and relator names ignored" |
52961 | 1986 |
else |
1987 |
(); |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1988 |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1989 |
val Bs = |
58234 | 1990 |
map3 (fn alive => fn A as TFree (_, S) => fn B => |
1991 |
if alive then resort_tfree_or_tvar S B else A) |
|
53266
89f7f1cc9ae3
cleaner handling of bootstrapping "fake" context, with fewer (no?) obscure bugs
blanchet
parents:
53264
diff
changeset
|
1992 |
(liveness_of_fp_bnf num_As any_fp_bnf) As Bs0; |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
1993 |
|
58446 | 1994 |
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); |
49167 | 1995 |
|
49501 | 1996 |
val ctors = map (mk_ctor As) ctors0; |
1997 |
val dtors = map (mk_dtor As) dtors0; |
|
49124 | 1998 |
|
49501 | 1999 |
val fpTs = map (domain_type o fastype_of) dtors; |
58446 | 2000 |
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
|
2001 |
|
58337 | 2002 |
val code_attrs = if plugins code_plugin then [Code.add_default_eqn_attrib] else []; |
2003 |
||
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
|
2004 |
val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Xs ~~ fpTs)))) ctrXs_Tsss; |
49203 | 2005 |
val ns = map length ctr_Tsss; |
49212 | 2006 |
val kss = map (fn n => 1 upto n) ns; |
49203 | 2007 |
val mss = map (map length) ctr_Tsss; |
2008 |
||
55867 | 2009 |
val ((xtor_co_recs, recs_args_types, corecs_args_types), lthy') = |
55868 | 2010 |
mk_co_recs_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 lthy; |
49210 | 2011 |
|
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2012 |
fun define_ctrs_dtrs_for_type (((((((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor), |
55867 | 2013 |
xtor_co_rec), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs), |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2014 |
pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), abs), |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2015 |
abs_inject), abs_inverse), type_definition), ctr_bindings), ctr_mixfixes), ctr_Tss), |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
2016 |
disc_bindings), sel_bindingss), raw_sel_default_eqs) no_defs_lthy = |
49176 | 2017 |
let |
49498 | 2018 |
val fp_b_name = Binding.name_of fp_b; |
2019 |
||
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2020 |
val ctr_absT = domain_type (fastype_of ctor); |
49119 | 2021 |
|
58347 | 2022 |
val (((w, xss), u'), _) = no_defs_lthy |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2023 |
|> yield_singleton (mk_Frees "w") ctr_absT |
49370 | 2024 |
||>> mk_Freess "x" ctr_Tss |
49498 | 2025 |
||>> yield_singleton Variable.variant_fixes fp_b_name; |
49370 | 2026 |
|
49498 | 2027 |
val u = Free (u', fpT); |
49121 | 2028 |
|
49129 | 2029 |
val ctr_rhss = |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2030 |
map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctr_absT abs n k xs)) |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2031 |
ks xss; |
49121 | 2032 |
|
53569
b4db0ade27bd
conceal definitions of high-level constructors and (co)iterators
traytel
parents:
53553
diff
changeset
|
2033 |
val maybe_conceal_def_binding = Thm.def_binding |
58208 | 2034 |
#> not (Config.get no_defs_lthy bnf_note_all) ? Binding.conceal; |
53569
b4db0ade27bd
conceal definitions of high-level constructors and (co)iterators
traytel
parents:
53553
diff
changeset
|
2035 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52964
diff
changeset
|
2036 |
val ((raw_ctrs, raw_ctr_defs), (lthy', lthy)) = no_defs_lthy |
49169 | 2037 |
|> apfst split_list o fold_map3 (fn b => fn mx => fn rhs => |
53569
b4db0ade27bd
conceal definitions of high-level constructors and (co)iterators
traytel
parents:
53553
diff
changeset
|
2038 |
Local_Theory.define ((b, mx), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd) |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52964
diff
changeset
|
2039 |
ctr_bindings ctr_mixfixes ctr_rhss |
49121 | 2040 |
||> `Local_Theory.restore; |
2041 |
||
2042 |
val phi = Proof_Context.export_morphism lthy lthy'; |
|
2043 |
||
2044 |
val ctr_defs = map (Morphism.thm phi) raw_ctr_defs; |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2045 |
val ctr_defs' = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2046 |
map2 (fn m => fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) ms ctr_defs; |
49130
3c26e17b2849
implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents:
49129
diff
changeset
|
2047 |
|
49203 | 2048 |
val ctrs0 = map (Morphism.term phi) raw_ctrs; |
2049 |
val ctrs = map (mk_ctr As) ctrs0; |
|
49121 | 2050 |
|
51897 | 2051 |
fun wrap_ctrs lthy = |
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49121
diff
changeset
|
2052 |
let |
50170 | 2053 |
fun exhaust_tac {context = ctxt, prems = _} = |
49135 | 2054 |
let |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2055 |
val ctor_iff_dtor_thm = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2056 |
let |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2057 |
val goal = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2058 |
fold_rev Logic.all [w, u] |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2059 |
(mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2060 |
in |
51551 | 2061 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2062 |
mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [ctr_absT, fpT]) |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2063 |
(certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor) |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55397
diff
changeset
|
2064 |
|> Morphism.thm phi |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2065 |
|> Thm.close_derivation |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2066 |
end; |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2067 |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2068 |
val sumEN_thm' = |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2069 |
unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms) |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2070 |
|> Morphism.thm phi; |
49135 | 2071 |
in |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2072 |
mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm' |
49135 | 2073 |
end; |
2074 |
||
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2075 |
val inject_tacss = |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2076 |
map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} => |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2077 |
mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms; |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2078 |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2079 |
val half_distinct_tacss = |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2080 |
map (map (fn (def, def') => fn {context = ctxt, ...} => |
55803
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2081 |
mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def'])) |
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents:
55772
diff
changeset
|
2082 |
(mk_half_pairss (`I ctr_defs)); |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2083 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52964
diff
changeset
|
2084 |
val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss; |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49583
diff
changeset
|
2085 |
|
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2086 |
val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss; |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2087 |
val sel_bTs = |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2088 |
flat sel_bindingss ~~ flat sel_Tss |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2089 |
|> filter_out (Binding.is_empty o fst) |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2090 |
|> distinct (Binding.eq_name o pairself fst); |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2091 |
val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy; |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2092 |
|
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57824
diff
changeset
|
2093 |
val sel_default_eqs = map (prepare_term sel_default_lthy) raw_sel_default_eqs; |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
2094 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
2095 |
fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs); |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
2096 |
val ctr_specs = map3 ctr_spec_of disc_bindings ctrs0 sel_bindingss; |
58267
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2097 |
|
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2098 |
val (ctr_sugar as {case_cong, ...}, lthy') = |
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2099 |
free_constructors tacss ((((plugins, discs_sels), standard_binding), ctr_specs), |
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2100 |
sel_default_eqs) lthy |
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2101 |
|
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2102 |
val anonymous_notes = |
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2103 |
[([case_cong], fundefcong_attrs)] |
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2104 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49121
diff
changeset
|
2105 |
in |
58267
4a6c9bcb4189
set 'fundef_cong' attribute also for (co)datatypes with no live type variables
blanchet
parents:
58265
diff
changeset
|
2106 |
(ctr_sugar, lthy' |> Local_Theory.notes anonymous_notes |> snd) |
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49121
diff
changeset
|
2107 |
end; |
49121 | 2108 |
|
55410 | 2109 |
fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b); |
49287
ebe2a5cec4bf
allow defaults for one datatype to involve the constructor of another one in the mutually recursive case
blanchet
parents:
49286
diff
changeset
|
2110 |
|
58346 | 2111 |
fun massage_res (((ctr_sugar, maps_sets_rels), co_rec_res), lthy) = |
55867 | 2112 |
(((maps_sets_rels, (ctrs, xss, ctr_defs, ctr_sugar)), co_rec_res), lthy); |
49119 | 2113 |
in |
51897 | 2114 |
(wrap_ctrs |
58346 | 2115 |
#> (fn (ctr_sugar, lthy) => |
58347 | 2116 |
derive_map_set_rel_thms plugins fp live As Bs abs_inverses ctr_defs' fp_nesting_set_maps |
2117 |
live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs fp_b_name fp_bnf fpT |
|
2118 |
ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms |
|
2119 |
fp_rel_thm ctr_Tss abs ctr_sugar lthy |
|
58346 | 2120 |
|>> pair ctr_sugar) |
52313 | 2121 |
##>> |
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
|
2122 |
(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
|
2123 |
else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec |
51835
56523caf372f
Added maps, sets, rels to "simps" thm collection
blanchet
parents:
51833
diff
changeset
|
2124 |
#> massage_res, lthy') |
49119 | 2125 |
end; |
49167 | 2126 |
|
58345 | 2127 |
fun wrap_ctrs_derive_map_set_rel_thms_define_co_rec_for_types (wrap_one_etc, lthy) = |
2128 |
fold_map I wrap_one_etc lthy |
|
58571 | 2129 |
|>> apsnd split_list o apfst (apsnd split_list4 o apfst split_list15 o split_list) |
52342 | 2130 |
o split_list; |
49226 | 2131 |
|
58462 | 2132 |
fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects |
58565 | 2133 |
rel_distincts set_thmss = |
2134 |
injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @ |
|
2135 |
set_thmss; |
|
49479
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
blanchet
parents:
49478
diff
changeset
|
2136 |
|
58448 | 2137 |
fun mk_co_rec_transfer_goals lthy co_recs = |
58446 | 2138 |
let |
2139 |
val liveAsBs = filter (op <>) (As ~~ Bs); |
|
2140 |
val B_ify = Term.subst_atomic_types (liveAsBs @ (Cs ~~ Es)); |
|
2141 |
||
2142 |
val ((Rs, Ss), names_lthy) = lthy |
|
2143 |
|> mk_Frees "R" (map (uncurry mk_pred2T) liveAsBs) |
|
2144 |
||>> mk_Frees "S" (map2 mk_pred2T Cs Es); |
|
2145 |
||
58448 | 2146 |
val co_recBs = map B_ify co_recs; |
58446 | 2147 |
in |
58448 | 2148 |
(Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy) |
2149 |
end; |
|
2150 |
||
2151 |
fun derive_rec_transfer_thms lthy recs rec_defs = |
|
2152 |
let val (Rs, Ss, goals, names_lthy) = mk_co_rec_transfer_goals lthy recs in |
|
58446 | 2153 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
2154 |
(fn {context = ctxt, prems = _} => |
|
58448 | 2155 |
mk_rec_transfer_tac ctxt nn ns (map (certify ctxt) Ss) (map (certify ctxt) Rs) rec_defs |
2156 |
xtor_co_rec_transfer_thms pre_rel_defs live_nesting_rel_eqs) |
|
58446 | 2157 |
|> Conjunction.elim_balanced nn |
2158 |
|> Proof_Context.export names_lthy lthy |
|
2159 |
|> map Thm.close_derivation |
|
2160 |
end; |
|
2161 |
||
55867 | 2162 |
fun derive_note_induct_recs_thms_for_types |
58563 | 2163 |
((((map_thmss, map_disc_iffss, map_selss, rel_injectss, rel_distinctss, rel_selss, |
58570 | 2164 |
rel_intross, rel_casess, set_thmss, set_selss, set_intross, set_casess, ctr_transferss, |
58571 | 2165 |
case_transferss, disc_transferss), |
58560 | 2166 |
(ctrss, _, ctr_defss, ctr_sugars)), |
55864 | 2167 |
(recs, rec_defs)), lthy) = |
49202
f493cd25737f
some work towards iterator and recursor properties
blanchet
parents:
49201
diff
changeset
|
2168 |
let |
55869 | 2169 |
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
|
2170 |
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
|
2171 |
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
|
2172 |
type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy; |
49201 | 2173 |
|
58446 | 2174 |
val rec_transfer_thmss = |
2175 |
if live = 0 then replicate nn [] |
|
58448 | 2176 |
else map single (derive_rec_transfer_thms lthy recs rec_defs); |
58446 | 2177 |
|
51830 | 2178 |
val induct_type_attr = Attrib.internal o K o Induct.induct_type; |
57471 | 2179 |
val induct_pred_attr = Attrib.internal o K o Induct.induct_pred; |
2180 |
||
2181 |
val ((rel_induct_thmss, common_rel_induct_thms), |
|
2182 |
(rel_induct_attrs, common_rel_induct_attrs)) = |
|
2183 |
if live = 0 then |
|
2184 |
((replicate nn [], []), ([], [])) |
|
2185 |
else |
|
2186 |
let |
|
2187 |
val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) = |
|
57535 | 2188 |
derive_rel_induct_thms_for_types lthy fpTs As Bs ctrss ctr_Tsss |
2189 |
(map #exhaust ctr_sugars) rel_xtor_co_induct_thm ctr_defss ctor_injects |
|
58328 | 2190 |
pre_rel_defs abs_inverses live_nesting_rel_eqs; |
57471 | 2191 |
in |
2192 |
((map single rel_induct_thms, single common_rel_induct_thm), |
|
2193 |
(rel_induct_attrs, rel_induct_attrs)) |
|
2194 |
end; |
|
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
2195 |
|
51835
56523caf372f
Added maps, sets, rels to "simps" thm collection
blanchet
parents:
51833
diff
changeset
|
2196 |
val simp_thmss = |
58565 | 2197 |
map6 mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; |
49438 | 2198 |
|
49337 | 2199 |
val common_notes = |
57471 | 2200 |
(if nn > 1 then |
57700 | 2201 |
[(inductN, [induct_thm], K induct_attrs), |
2202 |
(rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)] |
|
57633 | 2203 |
else |
2204 |
[]) |
|
51780 | 2205 |
|> massage_simple_notes fp_common_name; |
49337 | 2206 |
|
49226 | 2207 |
val notes = |
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
2208 |
[(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]), |
55869 | 2209 |
(recN, rec_thmss, K rec_attrs), |
58446 | 2210 |
(rec_transferN, rec_transfer_thmss, K []), |
57471 | 2211 |
(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
|
2212 |
(simpsN, simp_thmss, K [])] |
58346 | 2213 |
|> massage_multi_notes fp_b_names fpTs; |
49202
f493cd25737f
some work towards iterator and recursor properties
blanchet
parents:
49201
diff
changeset
|
2214 |
in |
51824 | 2215 |
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
|
2216 |
|> Spec_Rules.add Spec_Rules.Equational (recs, flat rec_thmss) |
57631
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents:
57568
diff
changeset
|
2217 |
|> Local_Theory.notes (common_notes @ notes) |
57633 | 2218 |
(* for "datatype_realizer.ML": *) |
2219 |
|>> name_noted_thms |
|
2220 |
(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
|
2221 |
|-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos |
58177
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents:
58175
diff
changeset
|
2222 |
fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctr_defss ctr_sugars recs rec_defs |
58565 | 2223 |
map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn []) |
2224 |
(replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selss rel_selss |
|
58569 | 2225 |
rel_intross rel_casess set_thmss set_selss set_intross set_casess ctr_transferss |
58573 | 2226 |
case_transferss disc_transferss (replicate nn []) (replicate nn []) rec_transfer_thmss |
58575 | 2227 |
common_rel_induct_thms rel_induct_thmss |
49202
f493cd25737f
some work towards iterator and recursor properties
blanchet
parents:
49201
diff
changeset
|
2228 |
end; |
f493cd25737f
some work towards iterator and recursor properties
blanchet
parents:
49201
diff
changeset
|
2229 |
|
58448 | 2230 |
fun derive_corec_transfer_thms lthy corecs corec_defs = |
2231 |
let |
|
2232 |
val (Rs, Ss, goals, names_lthy) = mk_co_rec_transfer_goals lthy corecs; |
|
2233 |
val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types; |
|
2234 |
in |
|
2235 |
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) |
|
2236 |
(fn {context = ctxt, prems = _} => |
|
2237 |
mk_corec_transfer_tac ctxt (map (certify ctxt) Ss) (map (certify ctxt) Rs) |
|
2238 |
type_definitions corec_defs xtor_co_rec_transfer_thms pre_rel_defs |
|
2239 |
live_nesting_rel_eqs (flat pgss) pss qssss gssss) |
|
2240 |
|> Conjunction.elim_balanced (length goals) |
|
2241 |
|> Proof_Context.export names_lthy lthy |
|
2242 |
|> map Thm.close_derivation |
|
2243 |
end; |
|
2244 |
||
55867 | 2245 |
fun derive_note_coinduct_corecs_thms_for_types |
58563 | 2246 |
((((map_thmss, map_disc_iffss, map_selss, rel_injectss, rel_distinctss, rel_selss, |
58570 | 2247 |
rel_intross, rel_casess, set_thmss, set_selss, set_intross, set_casess, ctr_transferss, |
58571 | 2248 |
case_transferss, disc_transferss), |
58560 | 2249 |
(_, _, ctr_defss, ctr_sugars)), |
55867 | 2250 |
(corecs, corec_defs)), lthy) = |
49212 | 2251 |
let |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2252 |
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
|
2253 |
(coinduct_attrs, common_coinduct_attrs)), |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2254 |
corec_thmss, corec_disc_thmss, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2255 |
(corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) = |
55867 | 2256 |
derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types) xtor_co_induct |
57397 | 2257 |
dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss |
2258 |
ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs |
|
58434 | 2259 |
(Proof_Context.export lthy' names_no_defs_lthy) lthy; |
49212 | 2260 |
|
58264 | 2261 |
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
|
2262 |
Goal.prove (*no sorry*) ctxt [] [] |
58264 | 2263 |
(thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies) |
2264 |
(fn _ => HEADGOAL (cut_tac thm THEN' atac) THEN ALLGOALS eq_assume_tac); |
|
57489 | 2265 |
|
2266 |
fun eq_ifIN _ [thm] = thm |
|
2267 |
| eq_ifIN ctxt (thm :: thms) = |
|
2268 |
distinct_prems ctxt (@{thm eq_ifI} OF |
|
2269 |
(map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]}) |
|
2270 |
[thm, eq_ifIN ctxt thms])); |
|
2271 |
||
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
|
2272 |
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
|
2273 |
val corec_sel_thmss = map flat corec_sel_thmsss; |
53475 | 2274 |
|
51830 | 2275 |
val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type; |
57302 | 2276 |
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
|
2277 |
|
55867 | 2278 |
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
|
2279 |
|
58448 | 2280 |
val corec_transfer_thmss = |
2281 |
if live = 0 then replicate nn [] |
|
2282 |
else map single (derive_corec_transfer_thms lthy corecs corec_defs); |
|
2283 |
||
57471 | 2284 |
val ((rel_coinduct_thmss, common_rel_coinduct_thms), |
2285 |
(rel_coinduct_attrs, common_rel_coinduct_attrs)) = |
|
2286 |
if live = 0 then |
|
2287 |
((replicate nn [], []), ([], [])) |
|
2288 |
else |
|
2289 |
let |
|
2290 |
val ((rel_coinduct_thms, common_rel_coinduct_thm), |
|
2291 |
(rel_coinduct_attrs, common_rel_coinduct_attrs)) = |
|
2292 |
derive_rel_coinduct_thm_for_types lthy fpTs ns As Bs mss ctr_sugars abs_inverses |
|
2293 |
abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss rel_xtor_co_induct_thm |
|
58328 | 2294 |
live_nesting_rel_eqs; |
57471 | 2295 |
in |
2296 |
((map single rel_coinduct_thms, single common_rel_coinduct_thm), |
|
2297 |
(rel_coinduct_attrs, common_rel_coinduct_attrs)) |
|
2298 |
end; |
|
2299 |
||
57700 | 2300 |
val (set_induct_thms, set_induct_attrss) = |
2301 |
derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars) |
|
2302 |
(map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_induct_thms |
|
2303 |
(map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) |
|
2304 |
dtor_ctors abs_inverses |
|
2305 |
|> split_list; |
|
2306 |
||
49479
504f0a38f608
added "simp"s to coiter/corec theorems + export under "simps" name
blanchet
parents:
49478
diff
changeset
|
2307 |
val simp_thmss = |
55856
bddaada24074
got rid of automatically generated fold constant and theorems (to reduce overhead)
blanchet
parents:
55803
diff
changeset
|
2308 |
map6 mk_simp_thms ctr_sugars |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2309 |
(map3 flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss) |
58565 | 2310 |
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
|
2311 |
|
49342 | 2312 |
val common_notes = |
57700 | 2313 |
(set_inductN, set_induct_thms, nth set_induct_attrss) :: |
57302 | 2314 |
(if nn > 1 then |
57700 | 2315 |
[(coinductN, [coinduct_thm], K common_coinduct_attrs), |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2316 |
(coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2317 |
(rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)] |
57700 | 2318 |
else []) |
51780 | 2319 |
|> massage_simple_notes fp_common_name; |
49342 | 2320 |
|
49212 | 2321 |
val notes = |
51780 | 2322 |
[(coinductN, map single coinduct_thms, |
51810
7b75fab5ebf5
factored out derivation of coinduction, unfold, corec
blanchet
parents:
51808
diff
changeset
|
2323 |
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
|
2324 |
(coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs), |
57489 | 2325 |
(corecN, corec_thmss, K []), |
58335
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58332
diff
changeset
|
2326 |
(corec_codeN, map single corec_code_thms, K (code_attrs @ nitpicksimp_attrs)), |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2327 |
(corec_discN, corec_disc_thmss, K []), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57932
diff
changeset
|
2328 |
(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
|
2329 |
(corec_selN, corec_sel_thmss, K corec_sel_attrs), |
58448 | 2330 |
(corec_transferN, corec_transfer_thmss, K []), |
57471 | 2331 |
(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
|
2332 |
(simpsN, simp_thmss, K [])] |
58346 | 2333 |
|> massage_multi_notes fp_b_names fpTs; |
49212 | 2334 |
in |
51824 | 2335 |
lthy |
55867 | 2336 |
|> 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
|
2337 |
[flat corec_sel_thmss, flat corec_thmss] |
57631
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents:
57568
diff
changeset
|
2338 |
|> Local_Theory.notes (common_notes @ notes) |
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
2339 |
|-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos |
58177
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents:
58175
diff
changeset
|
2340 |
fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctr_defss ctr_sugars corecs corec_defs |
58462 | 2341 |
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
|
2342 |
(transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss |
58562 | 2343 |
corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selss rel_selss |
58569 | 2344 |
rel_intross rel_casess set_thmss set_selss set_intross set_casess ctr_transferss |
58573 | 2345 |
case_transferss disc_transferss corec_disc_iff_thmss (map single corec_code_thms) |
58575 | 2346 |
corec_transfer_thmss common_rel_coinduct_thms rel_coinduct_thmss |
49212 | 2347 |
end; |
2348 |
||
52367 | 2349 |
val lthy'' = lthy' |
58177
166131276380
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
blanchet
parents:
58175
diff
changeset
|
2350 |
|> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52964
diff
changeset
|
2351 |
|> fold_map define_ctrs_dtrs_for_type (fp_bnfs ~~ fp_bs ~~ fpTs ~~ ctors ~~ dtors ~~ |
55867 | 2352 |
xtor_co_recs ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~ pre_set_defss ~~ |
2353 |
pre_rel_defs ~~ xtor_map_thms ~~ xtor_set_thmss ~~ xtor_rel_thms ~~ ns ~~ kss ~~ mss ~~ |
|
2354 |
abss ~~ abs_injects ~~ abs_inverses ~~ type_definitions ~~ ctr_bindingss ~~ ctr_mixfixess ~~ |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57199
diff
changeset
|
2355 |
ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~ raw_sel_default_eqss) |
58345 | 2356 |
|> wrap_ctrs_derive_map_set_rel_thms_define_co_rec_for_types |
55966 | 2357 |
|> case_fp fp derive_note_induct_recs_thms_for_types |
55867 | 2358 |
derive_note_coinduct_corecs_thms_for_types; |
49167 | 2359 |
|
2360 |
val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^ |
|
52899 | 2361 |
co_prefix fp ^ "datatype")); |
49112 | 2362 |
in |
52367 | 2363 |
timer; lthy'' |
49112 | 2364 |
end; |
2365 |
||
53974 | 2366 |
fun co_datatypes x = define_co_datatypes (K I) (K I) (K I) x; |
49297 | 2367 |
|
53974 | 2368 |
fun co_datatype_cmd x = |
2369 |
define_co_datatypes Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term x; |
|
49119 | 2370 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
2371 |
val parse_ctr_arg = |
57205 | 2372 |
@{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} |
2373 |
|| 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
|
2374 |
|
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
2375 |
val parse_ctr_specs = |
57091 | 2376 |
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
|
2377 |
|
bbcdd8519253
honor user-specified name for relator + generalize syntax
blanchet
parents:
51766
diff
changeset
|
2378 |
val parse_spec = |
57206
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
2379 |
parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix -- |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57205
diff
changeset
|
2380 |
(@{keyword "="} |-- parse_ctr_specs) -- parse_map_rel_bindings -- parse_sel_default_eqs; |
51767
bbcdd8519253
honor user-specified name for relator + generalize syntax
blanchet
parents:
51766
diff
changeset
|
2381 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
2382 |
val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec; |
49278 | 2383 |
|
52207
21026c312cc3
tuning -- avoided unreadable true/false all over the place for LFP/GFP
blanchet
parents:
52197
diff
changeset
|
2384 |
fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp; |
49112 | 2385 |
|
58179 | 2386 |
val _ = Theory.setup FP_Sugar_Interpretation.init; |
56522
f54003750e7d
don't forget to init Interpretation and transfer theorems in the interpretation hook
kuncar
parents:
56521
diff
changeset
|
2387 |
|
49112 | 2388 |
end; |