| author | Thomas Lindae <thomas.lindae@in.tum.de> | 
| Fri, 14 Jun 2024 10:21:03 +0200 | |
| changeset 81054 | 4bfcb14547c6 | 
| parent 80636 | 4041e7c8059d | 
| child 82643 | f1c14af17591 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_fp_util.ML  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 57668 | 4  | 
Author: Martin Desharnais, TU Muenchen  | 
| 
58256
 
08c0f0d4b9f4
generalized 'datatype' LaTeX antiquotation and added 'codatatype'
 
blanchet 
parents: 
58208 
diff
changeset
 | 
5  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 57668 | 6  | 
Copyright 2012, 2013, 2014  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
|
| 49389 | 8  | 
Shared library for the datatype and codatatype constructions.  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
|
| 
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: 
51839 
diff
changeset
 | 
11  | 
signature BNF_FP_UTIL =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
12  | 
sig  | 
| 
63826
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63813 
diff
changeset
 | 
13  | 
exception EMPTY_DATATYPE of string  | 
| 
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63813 
diff
changeset
 | 
14  | 
|
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49589 
diff
changeset
 | 
15  | 
type fp_result =  | 
| 51859 | 16  | 
    {Ts: typ list,
 | 
17  | 
bnfs: BNF_Def.bnf list,  | 
|
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
18  | 
pre_bnfs: BNF_Def.bnf list,  | 
| 
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
19  | 
absT_infos: BNF_Comp.absT_info list,  | 
| 51839 | 20  | 
ctors: term list,  | 
| 51819 | 21  | 
dtors: term list,  | 
| 62907 | 22  | 
xtor_un_folds: term list,  | 
| 55868 | 23  | 
xtor_co_recs: term list,  | 
| 53106 | 24  | 
xtor_co_induct: thm,  | 
| 51819 | 25  | 
dtor_ctors: thm list,  | 
26  | 
ctor_dtors: thm list,  | 
|
27  | 
ctor_injects: thm list,  | 
|
| 
53203
 
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
 
traytel 
parents: 
53202 
diff
changeset
 | 
28  | 
dtor_injects: thm list,  | 
| 58583 | 29  | 
xtor_maps: thm list,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
30  | 
xtor_map_unique: thm,  | 
| 58584 | 31  | 
xtor_setss: thm list list,  | 
| 58585 | 32  | 
xtor_rels: thm list,  | 
| 62907 | 33  | 
xtor_un_fold_thms: thm list,  | 
| 55868 | 34  | 
xtor_co_rec_thms: thm list,  | 
| 62907 | 35  | 
xtor_un_fold_unique: thm,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
36  | 
xtor_co_rec_unique: thm,  | 
| 62907 | 37  | 
xtor_un_fold_o_maps: thm list,  | 
| 
58578
 
9ff8ca957c02
rename 'xtor_co_rec_o_map_thms' to 'xtor_co_rec_o_maps'
 
desharna 
parents: 
58448 
diff
changeset
 | 
38  | 
xtor_co_rec_o_maps: thm list,  | 
| 62907 | 39  | 
xtor_un_fold_transfers: thm list,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
40  | 
xtor_co_rec_transfers: thm list,  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
41  | 
xtor_rel_co_induct: thm,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
42  | 
dtor_set_inducts: thm list}  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49589 
diff
changeset
 | 
43  | 
|
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
44  | 
val morph_fp_result: morphism -> fp_result -> fp_result  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
45  | 
|
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53138 
diff
changeset
 | 
46  | 
val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
47  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
48  | 
  val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
49  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
50  | 
val IITN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
51  | 
val LevN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
52  | 
val algN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
53  | 
val behN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
54  | 
val bisN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
55  | 
val carTN: string  | 
| 49338 | 56  | 
val caseN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
val coN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
val coinductN: string  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
59  | 
val coinduct_strongN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
60  | 
val corecN: string  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
61  | 
val corec_discN: string  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
62  | 
val corec_disc_iffN: string  | 
| 49501 | 63  | 
val ctorN: string  | 
64  | 
val ctor_dtorN: string  | 
|
65  | 
val ctor_exhaustN: string  | 
|
66  | 
val ctor_induct2N: string  | 
|
67  | 
val ctor_inductN: string  | 
|
68  | 
val ctor_injectN: string  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
69  | 
val ctor_foldN: string  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
70  | 
val ctor_fold_o_mapN: string  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
71  | 
val ctor_fold_transferN: string  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
72  | 
val ctor_fold_uniqueN: string  | 
| 49541 | 73  | 
val ctor_mapN: string  | 
| 
49543
 
53b3c532a082
renamed low-level "map_unique" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49542 
diff
changeset
 | 
74  | 
val ctor_map_uniqueN: string  | 
| 49501 | 75  | 
val ctor_recN: string  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
76  | 
val ctor_rec_o_mapN: string  | 
| 58444 | 77  | 
val ctor_rec_transferN: string  | 
| 
51739
 
3514b90d0a8b
(co)rec is (just as the (un)fold) the unique morphism;
 
traytel 
parents: 
49635 
diff
changeset
 | 
78  | 
val ctor_rec_uniqueN: string  | 
| 
49518
 
b377da40244b
renamed LFP low-level rel property to have ctor not dtor in its name
 
blanchet 
parents: 
49516 
diff
changeset
 | 
79  | 
val ctor_relN: string  | 
| 
55901
 
8c6d49dd8ae1
renamed a pair of low-level theorems to have c/dtor in their names (like the others)
 
blanchet 
parents: 
55899 
diff
changeset
 | 
80  | 
val ctor_rel_inductN: string  | 
| 
49544
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
81  | 
val ctor_set_inclN: string  | 
| 
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
82  | 
val ctor_set_set_inclN: string  | 
| 49501 | 83  | 
val dtorN: string  | 
| 
49582
 
557302525778
renamed "dtor_rel_coinduct" etc. to "dtor_coinduct"
 
blanchet 
parents: 
49581 
diff
changeset
 | 
84  | 
val dtor_coinductN: string  | 
| 49501 | 85  | 
val dtor_corecN: string  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
86  | 
val dtor_corec_o_mapN: string  | 
| 58445 | 87  | 
val dtor_corec_transferN: string  | 
| 
51739
 
3514b90d0a8b
(co)rec is (just as the (un)fold) the unique morphism;
 
traytel 
parents: 
49635 
diff
changeset
 | 
88  | 
val dtor_corec_uniqueN: string  | 
| 
49518
 
b377da40244b
renamed LFP low-level rel property to have ctor not dtor in its name
 
blanchet 
parents: 
49516 
diff
changeset
 | 
89  | 
val dtor_ctorN: string  | 
| 49501 | 90  | 
val dtor_exhaustN: string  | 
91  | 
val dtor_injectN: string  | 
|
| 
49545
 
8bb6e2d7346b
renamed coinduction principles to have "dtor" in the name
 
blanchet 
parents: 
49544 
diff
changeset
 | 
92  | 
val dtor_mapN: string  | 
| 
49581
 
4e5bd3883429
renamed "dtor_coinduct" etc. to "dtor_map_coinduct"
 
blanchet 
parents: 
49545 
diff
changeset
 | 
93  | 
val dtor_map_coinductN: string  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
94  | 
val dtor_map_coinduct_strongN: string  | 
| 
49543
 
53b3c532a082
renamed low-level "map_unique" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49542 
diff
changeset
 | 
95  | 
val dtor_map_uniqueN: string  | 
| 
49545
 
8bb6e2d7346b
renamed coinduction principles to have "dtor" in the name
 
blanchet 
parents: 
49544 
diff
changeset
 | 
96  | 
val dtor_relN: string  | 
| 
55901
 
8c6d49dd8ae1
renamed a pair of low-level theorems to have c/dtor in their names (like the others)
 
blanchet 
parents: 
55899 
diff
changeset
 | 
97  | 
val dtor_rel_coinductN: string  | 
| 
49544
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
98  | 
val dtor_set_inclN: string  | 
| 
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
99  | 
val dtor_set_set_inclN: string  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
100  | 
val dtor_coinduct_strongN: string  | 
| 
49516
 
d4859efc1096
renamed "rel_simp" to "dtor_rel" and similarly for "srel"
 
blanchet 
parents: 
49510 
diff
changeset
 | 
101  | 
val dtor_unfoldN: string  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
102  | 
val dtor_unfold_o_mapN: string  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
103  | 
val dtor_unfold_transferN: string  | 
| 
49516
 
d4859efc1096
renamed "rel_simp" to "dtor_rel" and similarly for "srel"
 
blanchet 
parents: 
49510 
diff
changeset
 | 
104  | 
val dtor_unfold_uniqueN: string  | 
| 
49020
 
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
 
blanchet 
parents: 
49019 
diff
changeset
 | 
105  | 
val exhaustN: string  | 
| 56113 | 106  | 
val colN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
107  | 
val inductN: string  | 
| 49019 | 108  | 
val injectN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
109  | 
val isNodeN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
110  | 
val lsbisN: string  | 
| 
49594
 
55e798614c45
tweaked theorem names (in particular, dropped s's)
 
blanchet 
parents: 
49592 
diff
changeset
 | 
111  | 
val mapN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
112  | 
val map_uniqueN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
113  | 
val min_algN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
114  | 
val morN: string  | 
| 
49020
 
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
 
blanchet 
parents: 
49019 
diff
changeset
 | 
115  | 
val nchotomyN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
116  | 
val recN: string  | 
| 57525 | 117  | 
val rel_casesN: string  | 
| 51918 | 118  | 
val rel_coinductN: string  | 
119  | 
val rel_inductN: string  | 
|
| 
49592
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
120  | 
val rel_injectN: string  | 
| 57493 | 121  | 
val rel_introsN: string  | 
| 
49592
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
122  | 
val rel_distinctN: string  | 
| 57563 | 123  | 
val rel_selN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
124  | 
val rvN: string  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
125  | 
val corec_selN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
126  | 
val set_inclN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
val set_set_inclN: string  | 
| 53694 | 128  | 
val setN: string  | 
| 49438 | 129  | 
val simpsN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
val strTN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
val str_initN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
val sum_bdN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
val sum_bdTN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
val uniqueN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
135  | 
|
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
136  | 
(* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *)  | 
| 
49584
 
4339aa335355
use singular since there is always only one theorem
 
blanchet 
parents: 
49582 
diff
changeset
 | 
137  | 
val mk_ctor_setN: int -> string  | 
| 
 
4339aa335355
use singular since there is always only one theorem
 
blanchet 
parents: 
49582 
diff
changeset
 | 
138  | 
val mk_dtor_setN: int -> string  | 
| 
49542
 
b39354db8629
renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49541 
diff
changeset
 | 
139  | 
val mk_dtor_set_inductN: int -> string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
140  | 
val mk_set_inductN: int -> string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
141  | 
|
| 
55575
 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 
blanchet 
parents: 
55570 
diff
changeset
 | 
142  | 
val co_prefix: BNF_Util.fp_kind -> string  | 
| 51863 | 143  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
144  | 
val split_conj_thm: thm -> thm list  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
145  | 
val split_conj_prems: int -> thm -> thm  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
146  | 
|
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
147  | 
val mk_sumTN: typ list -> typ  | 
| 58159 | 148  | 
val mk_sumTN_balanced: typ list -> typ  | 
| 55969 | 149  | 
val mk_tupleT_balanced: typ list -> typ  | 
| 55966 | 150  | 
val mk_sumprodT_balanced: typ list list -> typ  | 
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
151  | 
|
| 
54923
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
152  | 
val mk_proj: typ -> int -> int -> term  | 
| 
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
153  | 
|
| 53032 | 154  | 
val mk_convol: term * term -> term  | 
| 58443 | 155  | 
val mk_rel_prod: term -> term -> term  | 
156  | 
val mk_rel_sum: term -> term -> term  | 
|
| 49368 | 157  | 
|
| 49121 | 158  | 
val Inl_const: typ -> typ -> term  | 
159  | 
val Inr_const: typ -> typ -> term  | 
|
| 55969 | 160  | 
val mk_tuple_balanced: term list -> term  | 
161  | 
val mk_tuple1_balanced: typ list -> term list -> term  | 
|
| 59873 | 162  | 
val abs_curried_balanced: typ list -> term -> term  | 
| 62124 | 163  | 
val mk_tupled_fun: term -> term -> term list -> term  | 
| 49121 | 164  | 
|
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
165  | 
val mk_case_sum: term * term -> term  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
166  | 
val mk_case_sumN: term list -> term  | 
| 55968 | 167  | 
val mk_case_absumprod: typ -> term -> term list -> term list list -> term list list -> term  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
168  | 
|
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
169  | 
val mk_Inl: typ -> term -> term  | 
| 
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
170  | 
val mk_Inr: typ -> term -> term  | 
| 59673 | 171  | 
val mk_sumprod_balanced: typ -> int -> int -> term list -> term  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
172  | 
val mk_absumprod: typ -> term -> int -> int -> term list -> term  | 
| 49121 | 173  | 
|
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
174  | 
val dest_sumT: typ -> typ * typ  | 
| 58159 | 175  | 
val dest_sumTN_balanced: int -> typ -> typ list  | 
176  | 
val dest_tupleT_balanced: int -> typ -> typ list  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
177  | 
val dest_absumprodT: typ -> typ -> int -> int list -> typ -> typ list list  | 
| 49176 | 178  | 
|
| 
53202
 
2333fae25719
export one more ML function, needed for primcorec
 
blanchet 
parents: 
53143 
diff
changeset
 | 
179  | 
val If_const: typ -> term  | 
| 
 
2333fae25719
export one more ML function, needed for primcorec
 
blanchet 
parents: 
53143 
diff
changeset
 | 
180  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
181  | 
val mk_Field: term -> term  | 
| 49275 | 182  | 
val mk_If: term -> term -> term -> term  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
184  | 
val mk_absumprodE: thm -> int list -> thm  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
185  | 
|
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
186  | 
val mk_sum_caseN: int -> int -> thm  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
187  | 
val mk_sum_caseN_balanced: int -> int -> thm  | 
| 49125 | 188  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
189  | 
val mk_sum_Cinfinite: thm list -> thm  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
190  | 
val mk_sum_card_order: thm list -> thm  | 
| 75624 | 191  | 
val mk_sum_Cinfinite_regularCard: (thm * thm) list -> thm * thm  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
192  | 
|
| 62905 | 193  | 
val force_typ: Proof.context -> typ -> term -> term  | 
194  | 
||
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
195  | 
val mk_xtor_rel_co_induct_thm: BNF_Util.fp_kind -> term list -> term list -> term list ->  | 
| 
55575
 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 
blanchet 
parents: 
55570 
diff
changeset
 | 
196  | 
term list -> term list -> term list -> term list -> term list ->  | 
| 
52505
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
197  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm
 | 
| 62827 | 198  | 
val mk_xtor_co_iter_transfer_thms: BNF_Util.fp_kind -> term list -> term list -> term list ->  | 
| 58443 | 199  | 
term list -> term list -> term list -> term list ->  | 
200  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm list
 | 
|
| 62827 | 201  | 
val mk_xtor_co_iter_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list ->  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
202  | 
thm list -> thm list -> thm list  | 
| 62905 | 203  | 
val derive_xtor_co_recs: BNF_Util.fp_kind -> binding list -> (typ list -> typ list) ->  | 
204  | 
(typ list list * typ list) -> BNF_Def.bnf list -> term list -> term list ->  | 
|
| 63045 | 205  | 
thm -> thm list -> thm list -> thm list -> thm list ->  | 
206  | 
(BNF_Comp.absT_info * BNF_Comp.absT_info) option list ->  | 
|
207  | 
local_theory ->  | 
|
| 62905 | 208  | 
(term list * (thm list * thm * thm list * thm list)) * local_theory  | 
| 63796 | 209  | 
val raw_qualify: (binding -> binding) -> binding -> binding -> binding  | 
| 
52505
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
210  | 
|
| 63802 | 211  | 
val fixpoint_bnf: bool -> (binding -> binding) ->  | 
| 62722 | 212  | 
(binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->  | 
213  | 
BNF_Comp.absT_info list -> local_theory -> 'a) ->  | 
|
| 63798 | 214  | 
binding list -> (string * sort) list -> (string * sort) list -> (string * sort) list ->  | 
215  | 
typ list -> BNF_Comp.comp_cache -> local_theory ->  | 
|
| 62720 | 216  | 
((BNF_Def.bnf list * BNF_Comp.absT_info list) * BNF_Comp.comp_cache) * 'a  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
217  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
218  | 
|
| 
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: 
51839 
diff
changeset
 | 
219  | 
structure BNF_FP_Util : BNF_FP_UTIL =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
220  | 
struct  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
221  | 
|
| 
56650
 
1f9ab71d43a5
no need to make 'size' generation an interpretation -- overkill
 
blanchet 
parents: 
56113 
diff
changeset
 | 
222  | 
open Ctr_Sugar  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
223  | 
open BNF_Comp  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
224  | 
open BNF_Def  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
225  | 
open BNF_Util  | 
| 62905 | 226  | 
open BNF_FP_Util_Tactics  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
227  | 
|
| 
63826
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63813 
diff
changeset
 | 
228  | 
exception EMPTY_DATATYPE of string;  | 
| 
 
9321b3d50abd
made it easier to catch 'empty datatype' exception
 
blanchet 
parents: 
63813 
diff
changeset
 | 
229  | 
|
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49589 
diff
changeset
 | 
230  | 
type fp_result =  | 
| 51859 | 231  | 
  {Ts: typ list,
 | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
232  | 
bnfs: bnf list,  | 
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
233  | 
pre_bnfs: BNF_Def.bnf list,  | 
| 
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
234  | 
absT_infos: BNF_Comp.absT_info list,  | 
| 51839 | 235  | 
ctors: term list,  | 
| 51819 | 236  | 
dtors: term list,  | 
| 62907 | 237  | 
xtor_un_folds: term list,  | 
| 55868 | 238  | 
xtor_co_recs: term list,  | 
| 53106 | 239  | 
xtor_co_induct: thm,  | 
| 51819 | 240  | 
dtor_ctors: thm list,  | 
241  | 
ctor_dtors: thm list,  | 
|
242  | 
ctor_injects: thm list,  | 
|
| 
53203
 
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
 
traytel 
parents: 
53202 
diff
changeset
 | 
243  | 
dtor_injects: thm list,  | 
| 58583 | 244  | 
xtor_maps: thm list,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
245  | 
xtor_map_unique: thm,  | 
| 58584 | 246  | 
xtor_setss: thm list list,  | 
| 58585 | 247  | 
xtor_rels: thm list,  | 
| 62907 | 248  | 
xtor_un_fold_thms: thm list,  | 
| 55868 | 249  | 
xtor_co_rec_thms: thm list,  | 
| 62907 | 250  | 
xtor_un_fold_unique: thm,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
251  | 
xtor_co_rec_unique: thm,  | 
| 62907 | 252  | 
xtor_un_fold_o_maps: thm list,  | 
| 
58578
 
9ff8ca957c02
rename 'xtor_co_rec_o_map_thms' to 'xtor_co_rec_o_maps'
 
desharna 
parents: 
58448 
diff
changeset
 | 
253  | 
xtor_co_rec_o_maps: thm list,  | 
| 62907 | 254  | 
xtor_un_fold_transfers: thm list,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
255  | 
xtor_co_rec_transfers: thm list,  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
256  | 
xtor_rel_co_induct: thm,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
257  | 
dtor_set_inducts: thm list};  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49589 
diff
changeset
 | 
258  | 
|
| 62907 | 259  | 
fun morph_fp_result phi {Ts, bnfs, pre_bnfs, absT_infos, ctors, dtors, xtor_un_folds,
 | 
| 62721 | 260  | 
xtor_co_recs, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps,  | 
| 62907 | 261  | 
xtor_map_unique, xtor_setss, xtor_rels, xtor_un_fold_thms, xtor_co_rec_thms,  | 
262  | 
xtor_un_fold_unique, xtor_co_rec_unique, xtor_un_fold_o_maps,  | 
|
263  | 
xtor_co_rec_o_maps, xtor_un_fold_transfers, xtor_co_rec_transfers, xtor_rel_co_induct,  | 
|
| 62721 | 264  | 
dtor_set_inducts} =  | 
| 51859 | 265  | 
  {Ts = map (Morphism.typ phi) Ts,
 | 
266  | 
bnfs = map (morph_bnf phi) bnfs,  | 
|
| 
62684
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
267  | 
pre_bnfs = map (morph_bnf phi) pre_bnfs,  | 
| 
 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 
traytel 
parents: 
62621 
diff
changeset
 | 
268  | 
absT_infos = map (morph_absT_info phi) absT_infos,  | 
| 51839 | 269  | 
ctors = map (Morphism.term phi) ctors,  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
270  | 
dtors = map (Morphism.term phi) dtors,  | 
| 62907 | 271  | 
xtor_un_folds = map (Morphism.term phi) xtor_un_folds,  | 
| 55868 | 272  | 
xtor_co_recs = map (Morphism.term phi) xtor_co_recs,  | 
| 53106 | 273  | 
xtor_co_induct = Morphism.thm phi xtor_co_induct,  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
274  | 
dtor_ctors = map (Morphism.thm phi) dtor_ctors,  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
275  | 
ctor_dtors = map (Morphism.thm phi) ctor_dtors,  | 
| 
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
276  | 
ctor_injects = map (Morphism.thm phi) ctor_injects,  | 
| 
53203
 
222ea6acbdd6
moved derivation of ctor_dtor_unfold to sugar (streamlines fp_result interface)
 
traytel 
parents: 
53202 
diff
changeset
 | 
277  | 
dtor_injects = map (Morphism.thm phi) dtor_injects,  | 
| 58583 | 278  | 
xtor_maps = map (Morphism.thm phi) xtor_maps,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
279  | 
xtor_map_unique = Morphism.thm phi xtor_map_unique,  | 
| 58584 | 280  | 
xtor_setss = map (map (Morphism.thm phi)) xtor_setss,  | 
| 58585 | 281  | 
xtor_rels = map (Morphism.thm phi) xtor_rels,  | 
| 62907 | 282  | 
xtor_un_fold_thms = map (Morphism.thm phi) xtor_un_fold_thms,  | 
| 55868 | 283  | 
xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms,  | 
| 62907 | 284  | 
xtor_un_fold_unique = Morphism.thm phi xtor_un_fold_unique,  | 
| 
62863
 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 
traytel 
parents: 
62827 
diff
changeset
 | 
285  | 
xtor_co_rec_unique = Morphism.thm phi xtor_co_rec_unique,  | 
| 62907 | 286  | 
xtor_un_fold_o_maps = map (Morphism.thm phi) xtor_un_fold_o_maps,  | 
| 
58578
 
9ff8ca957c02
rename 'xtor_co_rec_o_map_thms' to 'xtor_co_rec_o_maps'
 
desharna 
parents: 
58448 
diff
changeset
 | 
287  | 
xtor_co_rec_o_maps = map (Morphism.thm phi) xtor_co_rec_o_maps,  | 
| 62907 | 288  | 
xtor_un_fold_transfers = map (Morphism.thm phi) xtor_un_fold_transfers,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
289  | 
xtor_co_rec_transfers = map (Morphism.thm phi) xtor_co_rec_transfers,  | 
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
290  | 
xtor_rel_co_induct = Morphism.thm phi xtor_rel_co_induct,  | 
| 
59856
 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 
blanchet 
parents: 
59819 
diff
changeset
 | 
291  | 
dtor_set_inducts = map (Morphism.thm phi) dtor_set_inducts};  | 
| 
51823
 
38996458bc5c
create data structure for storing (co)datatype information
 
blanchet 
parents: 
51819 
diff
changeset
 | 
292  | 
|
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53138 
diff
changeset
 | 
293  | 
fun time ctxt timer msg = (if Config.get ctxt bnf_timing  | 
| 55811 | 294  | 
then warning (msg ^ ": " ^ string_of_int (Time.toMilliseconds (Timer.checkRealTimer timer)) ^  | 
| 62685 | 295  | 
" ms")  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
296  | 
else (); Timer.startRealTimer ());  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
297  | 
|
| 49223 | 298  | 
val preN = "pre_"  | 
299  | 
val rawN = "raw_"  | 
|
| 49218 | 300  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
301  | 
val coN = "co"  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
302  | 
val unN = "un"  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
303  | 
val algN = "alg"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
304  | 
val IITN = "IITN"  | 
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
305  | 
val foldN = "fold"  | 
| 
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
306  | 
val unfoldN = unN ^ foldN  | 
| 62125 | 307  | 
val uniqueN = "unique"  | 
308  | 
val transferN = "transfer"  | 
|
| 49438 | 309  | 
val simpsN = "simps"  | 
| 49501 | 310  | 
val ctorN = "ctor"  | 
311  | 
val dtorN = "dtor"  | 
|
| 
49504
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
312  | 
val ctor_foldN = ctorN ^ "_" ^ foldN  | 
| 
 
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
 
blanchet 
parents: 
49502 
diff
changeset
 | 
313  | 
val dtor_unfoldN = dtorN ^ "_" ^ unfoldN  | 
| 62125 | 314  | 
val ctor_fold_uniqueN = ctor_foldN ^ "_" ^ uniqueN  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
315  | 
val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN  | 
| 62125 | 316  | 
val dtor_unfold_uniqueN = dtor_unfoldN ^ "_" ^ uniqueN  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
317  | 
val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN  | 
| 62125 | 318  | 
val ctor_fold_transferN = ctor_foldN ^ "_" ^ transferN  | 
319  | 
val dtor_unfold_transferN = dtor_unfoldN ^ "_" ^ transferN  | 
|
| 49541 | 320  | 
val ctor_mapN = ctorN ^ "_" ^ mapN  | 
321  | 
val dtor_mapN = dtorN ^ "_" ^ mapN  | 
|
| 62125 | 322  | 
val map_uniqueN = mapN ^ "_" ^ uniqueN  | 
| 
49543
 
53b3c532a082
renamed low-level "map_unique" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49542 
diff
changeset
 | 
323  | 
val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN  | 
| 
 
53b3c532a082
renamed low-level "map_unique" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49542 
diff
changeset
 | 
324  | 
val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
325  | 
val min_algN = "min_alg"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
326  | 
val morN = "mor"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
327  | 
val bisN = "bis"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
328  | 
val lsbisN = "lsbis"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
329  | 
val sum_bdTN = "sbdT"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
330  | 
val sum_bdN = "sbd"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
331  | 
val carTN = "carT"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
332  | 
val strTN = "strT"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
333  | 
val isNodeN = "isNode"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
334  | 
val LevN = "Lev"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
335  | 
val rvN = "recover"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
336  | 
val behN = "beh"  | 
| 53694 | 337  | 
val setN = "set"  | 
| 
49584
 
4339aa335355
use singular since there is always only one theorem
 
blanchet 
parents: 
49582 
diff
changeset
 | 
338  | 
val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN  | 
| 
 
4339aa335355
use singular since there is always only one theorem
 
blanchet 
parents: 
49582 
diff
changeset
 | 
339  | 
val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
340  | 
fun mk_set_inductN i = mk_setN i ^ "_induct"  | 
| 
49542
 
b39354db8629
renamed low-level "set_simps" and "set_induct" to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49541 
diff
changeset
 | 
341  | 
val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
342  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
343  | 
val str_initN = "str_init"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
344  | 
val recN = "rec"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
345  | 
val corecN = coN ^ recN  | 
| 49501 | 346  | 
val ctor_recN = ctorN ^ "_" ^ recN  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
347  | 
val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN  | 
| 62125 | 348  | 
val ctor_rec_transferN = ctor_recN ^ "_" ^ transferN  | 
349  | 
val ctor_rec_uniqueN = ctor_recN ^ "_" ^ uniqueN  | 
|
| 49501 | 350  | 
val dtor_corecN = dtorN ^ "_" ^ corecN  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
351  | 
val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN  | 
| 62125 | 352  | 
val dtor_corec_transferN = dtor_corecN ^ "_" ^ transferN  | 
353  | 
val dtor_corec_uniqueN = dtor_corecN ^ "_" ^ uniqueN  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
354  | 
|
| 49501 | 355  | 
val ctor_dtorN = ctorN ^ "_" ^ dtorN  | 
356  | 
val dtor_ctorN = dtorN ^ "_" ^ ctorN  | 
|
| 
49020
 
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
 
blanchet 
parents: 
49019 
diff
changeset
 | 
357  | 
val nchotomyN = "nchotomy"  | 
| 49019 | 358  | 
val injectN = "inject"  | 
| 
49020
 
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
 
blanchet 
parents: 
49019 
diff
changeset
 | 
359  | 
val exhaustN = "exhaust"  | 
| 
49585
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
360  | 
val ctor_injectN = ctorN ^ "_" ^ injectN  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
361  | 
val ctor_exhaustN = ctorN ^ "_" ^ exhaustN  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
362  | 
val dtor_injectN = dtorN ^ "_" ^ injectN  | 
| 
 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 
blanchet 
parents: 
49584 
diff
changeset
 | 
363  | 
val dtor_exhaustN = dtorN ^ "_" ^ exhaustN  | 
| 
49545
 
8bb6e2d7346b
renamed coinduction principles to have "dtor" in the name
 
blanchet 
parents: 
49544 
diff
changeset
 | 
364  | 
val ctor_relN = ctorN ^ "_" ^ relN  | 
| 
 
8bb6e2d7346b
renamed coinduction principles to have "dtor" in the name
 
blanchet 
parents: 
49544 
diff
changeset
 | 
365  | 
val dtor_relN = dtorN ^ "_" ^ relN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
366  | 
val inductN = "induct"  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
367  | 
val coinductN = coN ^ inductN  | 
| 49501 | 368  | 
val ctor_inductN = ctorN ^ "_" ^ inductN  | 
369  | 
val ctor_induct2N = ctor_inductN ^ "2"  | 
|
| 
49581
 
4e5bd3883429
renamed "dtor_coinduct" etc. to "dtor_map_coinduct"
 
blanchet 
parents: 
49545 
diff
changeset
 | 
370  | 
val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN  | 
| 
49582
 
557302525778
renamed "dtor_rel_coinduct" etc. to "dtor_coinduct"
 
blanchet 
parents: 
49581 
diff
changeset
 | 
371  | 
val dtor_coinductN = dtorN ^ "_" ^ coinductN  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
372  | 
val coinduct_strongN = coinductN ^ "_strong"  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
373  | 
val dtor_map_coinduct_strongN = dtor_mapN ^ "_" ^ coinduct_strongN  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
374  | 
val dtor_coinduct_strongN = dtorN ^ "_" ^ coinduct_strongN  | 
| 56113 | 375  | 
val colN = "col"  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
376  | 
val set_inclN = "set_incl"  | 
| 
49544
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
377  | 
val ctor_set_inclN = ctorN ^ "_" ^ set_inclN  | 
| 
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
378  | 
val dtor_set_inclN = dtorN ^ "_" ^ set_inclN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
379  | 
val set_set_inclN = "set_set_incl"  | 
| 
49544
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
380  | 
val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN  | 
| 
 
24094fa47e0d
renamed "set_incl" etc. to have "ctor" or "dtor" in the name
 
blanchet 
parents: 
49543 
diff
changeset
 | 
381  | 
val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
382  | 
|
| 49338 | 383  | 
val caseN = "case"  | 
| 49342 | 384  | 
val discN = "disc"  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
385  | 
val corec_discN = corecN ^ "_" ^ discN  | 
| 
49594
 
55e798614c45
tweaked theorem names (in particular, dropped s's)
 
blanchet 
parents: 
49592 
diff
changeset
 | 
386  | 
val iffN = "_iff"  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
387  | 
val corec_disc_iffN = corec_discN ^ iffN  | 
| 
49592
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
388  | 
val distinctN = "distinct"  | 
| 
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
389  | 
val rel_distinctN = relN ^ "_" ^ distinctN  | 
| 
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
390  | 
val injectN = "inject"  | 
| 57525 | 391  | 
val rel_casesN = relN ^ "_cases"  | 
| 
49592
 
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
 
blanchet 
parents: 
49591 
diff
changeset
 | 
392  | 
val rel_injectN = relN ^ "_" ^ injectN  | 
| 57493 | 393  | 
val rel_introsN = relN ^ "_intros"  | 
| 51918 | 394  | 
val rel_coinductN = relN ^ "_" ^ coinductN  | 
| 57563 | 395  | 
val rel_selN = relN ^ "_sel"  | 
| 
55901
 
8c6d49dd8ae1
renamed a pair of low-level theorems to have c/dtor in their names (like the others)
 
blanchet 
parents: 
55899 
diff
changeset
 | 
396  | 
val dtor_rel_coinductN = dtorN ^ "_" ^ rel_coinductN  | 
| 51918 | 397  | 
val rel_inductN = relN ^ "_" ^ inductN  | 
| 
55901
 
8c6d49dd8ae1
renamed a pair of low-level theorems to have c/dtor in their names (like the others)
 
blanchet 
parents: 
55899 
diff
changeset
 | 
398  | 
val ctor_rel_inductN = ctorN ^ "_" ^ rel_inductN  | 
| 49342 | 399  | 
val selN = "sel"  | 
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57700 
diff
changeset
 | 
400  | 
val corec_selN = corecN ^ "_" ^ selN  | 
| 49338 | 401  | 
|
| 55966 | 402  | 
fun co_prefix fp = case_fp fp "" "co";  | 
| 51863 | 403  | 
|
| 69593 | 404  | 
fun dest_sumT (Type (\<^type_name>\<open>sum\<close>, [T, T'])) = (T, T');  | 
| 49264 | 405  | 
|
406  | 
val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;  | 
|
407  | 
||
| 69593 | 408  | 
fun dest_tupleT_balanced 0 \<^typ>\<open>unit\<close> = []  | 
| 55966 | 409  | 
| dest_tupleT_balanced n T = Balanced_Tree.dest HOLogic.dest_prodT n T;  | 
| 49264 | 410  | 
|
| 55966 | 411  | 
fun dest_absumprodT absT repT n ms =  | 
412  | 
map2 dest_tupleT_balanced ms o dest_sumTN_balanced n o mk_repT absT repT;  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
413  | 
|
| 49264 | 414  | 
val mk_sumTN = Library.foldr1 mk_sumT;  | 
415  | 
val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;  | 
|
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
416  | 
|
| 55966 | 417  | 
fun mk_tupleT_balanced [] = HOLogic.unitT  | 
418  | 
| mk_tupleT_balanced Ts = Balanced_Tree.make HOLogic.mk_prodT Ts;  | 
|
419  | 
||
420  | 
val mk_sumprodT_balanced = mk_sumTN_balanced o map mk_tupleT_balanced;  | 
|
421  | 
||
| 
54923
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
422  | 
fun mk_proj T n k =  | 
| 
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
423  | 
let val (binders, _) = strip_typeN n T in  | 
| 
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
424  | 
fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1))  | 
| 
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
425  | 
end;  | 
| 
 
ffed2452f5f6
instantiate schematics as projections to avoid HOU trouble
 
blanchet 
parents: 
54171 
diff
changeset
 | 
426  | 
|
| 53032 | 427  | 
fun mk_convol (f, g) =  | 
428  | 
let  | 
|
429  | 
val (fU, fTU) = `range_type (fastype_of f);  | 
|
430  | 
val ((gT, gU), gTU) = `dest_funT (fastype_of g);  | 
|
431  | 
val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU);  | 
|
| 69593 | 432  | 
in Const (\<^const_name>\<open>convol\<close>, convolT) $ f $ g end;  | 
| 49368 | 433  | 
|
| 58443 | 434  | 
fun mk_rel_prod R S =  | 
435  | 
let  | 
|
436  | 
val ((A1, A2), RT) = `dest_pred2T (fastype_of R);  | 
|
437  | 
val ((B1, B2), ST) = `dest_pred2T (fastype_of S);  | 
|
438  | 
val rel_prodT = RT --> ST --> mk_pred2T (HOLogic.mk_prodT (A1, B1)) (HOLogic.mk_prodT (A2, B2));  | 
|
| 69593 | 439  | 
in Const (\<^const_name>\<open>rel_prod\<close>, rel_prodT) $ R $ S end;  | 
| 58443 | 440  | 
|
441  | 
fun mk_rel_sum R S =  | 
|
442  | 
let  | 
|
443  | 
val ((A1, A2), RT) = `dest_pred2T (fastype_of R);  | 
|
444  | 
val ((B1, B2), ST) = `dest_pred2T (fastype_of S);  | 
|
445  | 
val rel_sumT = RT --> ST --> mk_pred2T (mk_sumT (A1, B1)) (mk_sumT (A2, B2));  | 
|
| 69593 | 446  | 
in Const (\<^const_name>\<open>rel_sum\<close>, rel_sumT) $ R $ S end;  | 
| 58443 | 447  | 
|
| 69593 | 448  | 
fun Inl_const LT RT = Const (\<^const_name>\<open>Inl\<close>, LT --> mk_sumT (LT, RT));  | 
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
449  | 
fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t;  | 
| 49121 | 450  | 
|
| 69593 | 451  | 
fun Inr_const LT RT = Const (\<^const_name>\<open>Inr\<close>, RT --> mk_sumT (LT, RT));  | 
| 
49255
 
2ecc533d6697
use balanced sums for constructors (to gracefully handle 100 constructors or more)
 
blanchet 
parents: 
49240 
diff
changeset
 | 
452  | 
fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t;  | 
| 49121 | 453  | 
|
| 55969 | 454  | 
fun mk_prod1 bound_Ts (t, u) =  | 
455  | 
HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;  | 
|
456  | 
||
457  | 
fun mk_tuple1_balanced _ [] = HOLogic.unit  | 
|
458  | 
| mk_tuple1_balanced bound_Ts ts = Balanced_Tree.make (mk_prod1 bound_Ts) ts;  | 
|
459  | 
||
460  | 
val mk_tuple_balanced = mk_tuple1_balanced [];  | 
|
| 55966 | 461  | 
|
| 59873 | 462  | 
fun abs_curried_balanced Ts t =  | 
463  | 
t $ mk_tuple1_balanced (List.rev Ts) (map Bound (length Ts - 1 downto 0))  | 
|
464  | 
|> fold_rev (Term.abs o pair Name.uu) Ts;  | 
|
465  | 
||
| 59673 | 466  | 
fun mk_sumprod_balanced T n k ts = Sum_Tree.mk_inj T n k (mk_tuple_balanced ts);  | 
467  | 
||
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
468  | 
fun mk_absumprod absT abs0 n k ts =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
469  | 
let val abs = mk_abs absT abs0;  | 
| 59673 | 470  | 
in abs $ mk_sumprod_balanced (domain_type (fastype_of abs)) n k ts end;  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
471  | 
|
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
472  | 
fun mk_case_sum (f, g) =  | 
| 49129 | 473  | 
let  | 
| 55968 | 474  | 
val (fT, T') = dest_funT (fastype_of f);  | 
475  | 
val (gT, _) = dest_funT (fastype_of g);  | 
|
| 49129 | 476  | 
in  | 
| 55968 | 477  | 
Sum_Tree.mk_sumcase fT gT T' f g  | 
| 49129 | 478  | 
end;  | 
479  | 
||
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
480  | 
val mk_case_sumN = Library.foldr1 mk_case_sum;  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
481  | 
val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum;  | 
| 49176 | 482  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
483  | 
fun mk_tupled_fun f x xs =  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
484  | 
if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
485  | 
|
| 55968 | 486  | 
fun mk_case_absumprod absT rep fs xss xss' =  | 
| 62124 | 487  | 
HOLogic.mk_comp (mk_case_sumN_balanced  | 
488  | 
    (@{map 3} mk_tupled_fun fs (map mk_tuple_balanced xss) xss'), mk_rep absT rep);
 | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
489  | 
|
| 69593 | 490  | 
fun If_const T = Const (\<^const_name>\<open>If\<close>, HOLogic.boolT --> T --> T --> T);  | 
| 
53202
 
2333fae25719
export one more ML function, needed for primcorec
 
blanchet 
parents: 
53143 
diff
changeset
 | 
491  | 
fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;  | 
| 49275 | 492  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
493  | 
fun mk_Field r =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
494  | 
let val T = fst (dest_relT (fastype_of r));  | 
| 69593 | 495  | 
in Const (\<^const_name>\<open>Field\<close>, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
496  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
497  | 
(*dangerous; use with monotonic, converging functions only!*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
498  | 
fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
499  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
500  | 
(* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
501  | 
fun split_conj_thm th =  | 
| 49119 | 502  | 
((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
503  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
504  | 
fun split_conj_prems limit th =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
505  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
506  | 
fun split n i th =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
507  | 
if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
508  | 
in split limit 1 th end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
509  | 
|
| 49335 | 510  | 
fun mk_obj_sumEN_balanced n =  | 
511  | 
  Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
 | 
|
512  | 
(replicate n asm_rl);  | 
|
513  | 
||
| 55966 | 514  | 
fun mk_tupled_allIN_balanced 0 = @{thm unit_all_impI}
 | 
515  | 
| mk_tupled_allIN_balanced n =  | 
|
516  | 
let  | 
|
| 69593 | 517  | 
val (tfrees, _) = BNF_Util.mk_TFrees n \<^context>;  | 
| 55966 | 518  | 
val T = mk_tupleT_balanced tfrees;  | 
519  | 
in  | 
|
| 67091 | 520  | 
      @{thm asm_rl[of "\<forall>x. P x \<longrightarrow> Q x" for P Q]}
 | 
| 69593 | 521  | 
|> Thm.instantiate' [SOME (Thm.ctyp_of \<^context> T)] []  | 
522  | 
      |> Raw_Simplifier.rewrite_goals_rule \<^context> @{thms split_paired_All[THEN eq_reflection]}
 | 
|
| 55966 | 523  | 
|> (fn thm => impI RS funpow n (fn th => allI RS th) thm)  | 
524  | 
|> Thm.varifyT_global  | 
|
525  | 
end;  | 
|
| 49335 | 526  | 
|
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
527  | 
fun mk_absumprodE type_definition ms =  | 
| 49335 | 528  | 
let val n = length ms in  | 
| 55966 | 529  | 
mk_obj_sumEN_balanced n OF map mk_tupled_allIN_balanced ms RS  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
530  | 
      (type_definition RS @{thm type_copy_obj_one_point_absE})
 | 
| 49335 | 531  | 
end;  | 
| 49264 | 532  | 
|
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
533  | 
fun mk_sum_caseN 1 1 = refl  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
534  | 
  | mk_sum_caseN _ 1 = @{thm sum.case(1)}
 | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
535  | 
  | mk_sum_caseN 2 2 = @{thm sum.case(2)}
 | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
536  | 
  | mk_sum_caseN n k = trans OF [@{thm case_sum_step(2)}, mk_sum_caseN (n - 1) (k - 1)];
 | 
| 49264 | 537  | 
|
538  | 
fun mk_sum_step base step thm =  | 
|
539  | 
if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];  | 
|
540  | 
||
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
541  | 
fun mk_sum_caseN_balanced 1 1 = refl  | 
| 
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
55394 
diff
changeset
 | 
542  | 
| mk_sum_caseN_balanced n k =  | 
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55575 
diff
changeset
 | 
543  | 
    Balanced_Tree.access {left = mk_sum_step @{thm sum.case(1)} @{thm case_sum_step(1)},
 | 
| 
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55575 
diff
changeset
 | 
544  | 
      right = mk_sum_step @{thm sum.case(2)} @{thm case_sum_step(2)}, init = refl} n k;
 | 
| 
49130
 
3c26e17b2849
implemented "mk_case_tac" -- and got rid of "cheat_tac"
 
blanchet 
parents: 
49129 
diff
changeset
 | 
545  | 
|
| 
55851
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
546  | 
fun mk_sum_Cinfinite [thm] = thm  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
547  | 
  | mk_sum_Cinfinite (thm :: thms) = @{thm Cinfinite_csum_weak} OF [thm, mk_sum_Cinfinite thms];
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
548  | 
|
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
549  | 
fun mk_sum_card_order [thm] = thm  | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
550  | 
  | mk_sum_card_order (thm :: thms) = @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
 | 
| 
 
3d40cf74726c
optimize cardinal bounds involving natLeq (omega)
 
blanchet 
parents: 
55811 
diff
changeset
 | 
551  | 
|
| 75624 | 552  | 
fun mk_sum_Cinfinite_regularCard [x] = x  | 
553  | 
| mk_sum_Cinfinite_regularCard ((cinf, thm) :: thms) =  | 
|
554  | 
let val (cinf_sum, thm_sum) = mk_sum_Cinfinite_regularCard thms  | 
|
555  | 
in (  | 
|
556  | 
      @{thm Cinfinite_csum_weak} OF [cinf, cinf_sum],
 | 
|
557  | 
      @{thm regularCard_csum} OF [cinf, cinf_sum, thm, thm_sum]
 | 
|
558  | 
) end;  | 
|
559  | 
||
| 
58579
 
b7bc5ba2f3fb
rename 'rel_xtor_co_induct_thm' to 'xtor_rel_co_induct'
 
desharna 
parents: 
58578 
diff
changeset
 | 
560  | 
fun mk_xtor_rel_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy =  | 
| 
52505
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
561  | 
let  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
562  | 
val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
563  | 
val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
564  | 
fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
565  | 
val dtor = mk_xtor Greatest_FP;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
566  | 
val ctor = mk_xtor Least_FP;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
567  | 
fun flip f x y = if fp = Greatest_FP then f y x else f x y;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
568  | 
|
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
569  | 
fun mk_prem pre_relphi phi x y xtor xtor' =  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
570  | 
HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp)  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
571  | 
(pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y))));  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
572  | 
    val prems = @{map 6} mk_prem pre_relphis pre_phis xs ys xtors xtor's;
 | 
| 
52505
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
573  | 
|
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
574  | 
val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
575  | 
(map2 (flip mk_leq) relphis pre_phis));  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
576  | 
in  | 
| 52506 | 577  | 
Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac  | 
| 70494 | 578  | 
|> Thm.close_derivation \<^here>  | 
| 
52505
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
579  | 
    |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]}))
 | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
580  | 
end;  | 
| 
 
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
 
traytel 
parents: 
52344 
diff
changeset
 | 
581  | 
|
| 62827 | 582  | 
fun mk_xtor_co_iter_transfer_thms fp pre_rels pre_iphis pre_ophis rels phis un_folds un_folds' tac lthy =  | 
| 52731 | 583  | 
let  | 
| 58443 | 584  | 
val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_iphis)) pre_rels;  | 
| 52731 | 585  | 
val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;  | 
586  | 
fun flip f x y = if fp = Greatest_FP then f y x else f x y;  | 
|
587  | 
||
| 58443 | 588  | 
val arg_rels = map2 (flip mk_rel_fun) pre_relphis pre_ophis;  | 
| 52731 | 589  | 
fun mk_transfer relphi pre_phi un_fold un_fold' =  | 
| 55945 | 590  | 
fold_rev mk_rel_fun arg_rels (flip mk_rel_fun relphi pre_phi) $ un_fold $ un_fold';  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
591  | 
    val transfers = @{map 4} mk_transfer relphis pre_ophis un_folds un_folds';
 | 
| 52731 | 592  | 
|
| 58443 | 593  | 
val goal = fold_rev Logic.all (phis @ pre_ophis)  | 
| 52731 | 594  | 
(HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers));  | 
595  | 
in  | 
|
596  | 
Goal.prove_sorry lthy [] [] goal tac  | 
|
| 70494 | 597  | 
|> Thm.close_derivation \<^here>  | 
| 52731 | 598  | 
|> split_conj_thm  | 
599  | 
end;  | 
|
600  | 
||
| 62827 | 601  | 
fun mk_xtor_co_iter_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
602  | 
map_cong0s =  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
603  | 
let  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
604  | 
val n = length sym_map_comps;  | 
| 55966 | 605  | 
    val rewrite_comp_comp2 = case_fp fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2};
 | 
606  | 
    val rewrite_comp_comp = case_fp fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp};
 | 
|
607  | 
    val map_cong_passive_args1 = replicate m (case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong);
 | 
|
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
608  | 
val map_cong_active_args1 = replicate n (if is_rec  | 
| 55966 | 609  | 
      then case_fp fp @{thm convol_o} @{thm o_case_sum} RS fun_cong
 | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
610  | 
else refl);  | 
| 55966 | 611  | 
    val map_cong_passive_args2 = replicate m (case_fp fp @{thm comp_id} @{thm id_comp} RS fun_cong);
 | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
612  | 
val map_cong_active_args2 = replicate n (if is_rec  | 
| 55966 | 613  | 
      then case_fp fp @{thm map_prod_o_convol_id} @{thm case_sum_o_map_sum_id}
 | 
614  | 
      else case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong);
 | 
|
| 55990 | 615  | 
fun mk_map_congs passive active =  | 
616  | 
      map (fn thm => thm OF (passive @ active) RS @{thm ext}) map_cong0s;
 | 
|
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
617  | 
val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
618  | 
val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2;  | 
| 57489 | 619  | 
|
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
620  | 
fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong =>  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
621  | 
mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
622  | 
val rewrite1s = mk_rewrites map_cong1s;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
623  | 
val rewrite2s = mk_rewrites map_cong2s;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
624  | 
val unique_prems =  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
625  | 
      @{map 4} (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 =>
 | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
626  | 
mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold])  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
627  | 
(mk_trans rewrite1 (mk_sym rewrite2)))  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
628  | 
xtor_maps xtor_un_folds rewrite1s rewrite2s;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
629  | 
in  | 
| 55966 | 630  | 
split_conj_thm (un_fold_unique OF map (case_fp fp I mk_sym) unique_prems)  | 
| 
52913
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
631  | 
end;  | 
| 
 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 
traytel 
parents: 
52899 
diff
changeset
 | 
632  | 
|
| 62905 | 633  | 
fun force_typ ctxt T =  | 
634  | 
Term.map_types Type_Infer.paramify_vars  | 
|
635  | 
#> Type.constraint T  | 
|
636  | 
#> Syntax.check_term ctxt  | 
|
637  | 
#> singleton (Variable.polymorphic ctxt);  | 
|
638  | 
||
| 63045 | 639  | 
fun absT_info_encodeT thy (SOME (src : absT_info, dst : absT_info)) src_absT =  | 
640  | 
let  | 
|
641  | 
val src_repT = mk_repT (#absT src) (#repT src) src_absT;  | 
|
642  | 
val dst_absT = mk_absT thy (#repT dst) (#absT dst) src_repT;  | 
|
643  | 
in  | 
|
644  | 
dst_absT  | 
|
645  | 
end  | 
|
646  | 
| absT_info_encodeT _ NONE T = T;  | 
|
647  | 
||
648  | 
fun absT_info_decodeT thy = absT_info_encodeT thy o Option.map swap;  | 
|
649  | 
||
650  | 
fun absT_info_encode thy fp (opt as SOME (src : absT_info, dst : absT_info)) t =  | 
|
651  | 
let  | 
|
652  | 
val co_alg_funT = case_fp fp domain_type range_type;  | 
|
653  | 
fun co_swap pair = case_fp fp I swap pair;  | 
|
654  | 
val mk_co_comp = curry (HOLogic.mk_comp o co_swap);  | 
|
655  | 
val mk_co_abs = case_fp fp mk_abs mk_rep;  | 
|
656  | 
val mk_co_rep = case_fp fp mk_rep mk_abs;  | 
|
657  | 
val co_abs = case_fp fp #abs #rep;  | 
|
658  | 
val co_rep = case_fp fp #rep #abs;  | 
|
659  | 
val src_absT = co_alg_funT (fastype_of t);  | 
|
660  | 
val dst_absT = absT_info_encodeT thy opt src_absT;  | 
|
661  | 
val co_src_abs = mk_co_abs src_absT (co_abs src);  | 
|
662  | 
val co_dst_rep = mk_co_rep dst_absT (co_rep dst);  | 
|
663  | 
in  | 
|
664  | 
mk_co_comp (mk_co_comp t co_src_abs) co_dst_rep  | 
|
665  | 
end  | 
|
666  | 
| absT_info_encode _ _ NONE t = t;  | 
|
667  | 
||
668  | 
fun absT_info_decode thy fp = absT_info_encode thy fp o Option.map swap;  | 
|
669  | 
||
670  | 
fun mk_xtor_un_fold_xtor_thms ctxt fp un_folds xtors xtor_un_fold_unique map_id0s  | 
|
671  | 
absT_info_opts =  | 
|
672  | 
let  | 
|
673  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
674  | 
fun mk_goal un_fold =  | 
|
675  | 
let  | 
|
676  | 
        val rhs = list_comb (un_fold, @{map 2} (absT_info_encode thy fp) absT_info_opts xtors);
 | 
|
677  | 
val T = range_type (fastype_of rhs);  | 
|
678  | 
in  | 
|
679  | 
HOLogic.mk_eq (HOLogic.id_const T, rhs)  | 
|
| 63813 | 680  | 
end;  | 
| 63045 | 681  | 
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map mk_goal un_folds));  | 
682  | 
fun mk_inverses NONE = []  | 
|
683  | 
| mk_inverses (SOME (src, dst)) =  | 
|
684  | 
        [#type_definition dst RS @{thm type_definition.Abs_inverse[OF _ UNIV_I]},
 | 
|
685  | 
         #type_definition src RS @{thm type_definition.Rep_inverse}];
 | 
|
686  | 
val inverses = maps mk_inverses absT_info_opts;  | 
|
687  | 
in  | 
|
688  | 
    Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} =>
 | 
|
689  | 
mk_xtor_un_fold_xtor_tac ctxt xtor_un_fold_unique map_id0s inverses)  | 
|
690  | 
|> split_conj_thm |> map mk_sym  | 
|
691  | 
end;  | 
|
| 62905 | 692  | 
|
693  | 
fun derive_xtor_co_recs fp bs mk_Ts (Dss, resDs) pre_bnfs xtors0 un_folds0  | 
|
| 63045 | 694  | 
xtor_un_fold_unique xtor_un_folds xtor_un_fold_transfers xtor_maps xtor_rels  | 
695  | 
absT_info_opts lthy =  | 
|
| 62905 | 696  | 
let  | 
| 63045 | 697  | 
val thy = Proof_Context.theory_of lthy;  | 
| 62905 | 698  | 
fun co_swap pair = case_fp fp I swap pair;  | 
699  | 
val mk_co_comp = curry (HOLogic.mk_comp o co_swap);  | 
|
700  | 
fun mk_co_algT T U = case_fp fp (T --> U) (U --> T);  | 
|
701  | 
val co_alg_funT = case_fp fp domain_type range_type;  | 
|
702  | 
val mk_co_product = curry (case_fp fp mk_convol mk_case_sum);  | 
|
703  | 
val co_proj1_const = case_fp fp fst_const (uncurry Inl_const o dest_sumT) o co_alg_funT;  | 
|
704  | 
val co_proj2_const = case_fp fp snd_const (uncurry Inr_const o dest_sumT) o co_alg_funT;  | 
|
705  | 
val mk_co_productT = curry (case_fp fp HOLogic.mk_prodT mk_sumT);  | 
|
| 63045 | 706  | 
    val rewrite_comp_comp = case_fp fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp};
 | 
| 62905 | 707  | 
|
708  | 
val n = length pre_bnfs;  | 
|
709  | 
val live = live_of_bnf (hd pre_bnfs);  | 
|
710  | 
val m = live - n;  | 
|
711  | 
val ks = 1 upto n;  | 
|
712  | 
||
713  | 
val map_id0s = map map_id0_of_bnf pre_bnfs;  | 
|
714  | 
val map_comps = map map_comp_of_bnf pre_bnfs;  | 
|
715  | 
val map_cong0s = map map_cong0_of_bnf pre_bnfs;  | 
|
716  | 
val map_transfers = map map_transfer_of_bnf pre_bnfs;  | 
|
717  | 
val sym_map_comp0s = map (mk_sym o map_comp0_of_bnf) pre_bnfs;  | 
|
718  | 
||
719  | 
val deads = fold (union (op =)) Dss resDs;  | 
|
720  | 
val ((((As, Bs), Xs), Ys), names_lthy) = lthy  | 
|
721  | 
|> fold Variable.declare_typ deads  | 
|
722  | 
|> mk_TFrees m  | 
|
723  | 
||>> mk_TFrees m  | 
|
724  | 
||>> mk_TFrees n  | 
|
725  | 
||>> mk_TFrees n;  | 
|
726  | 
||
727  | 
    val XFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Xs)) Dss pre_bnfs;
 | 
|
728  | 
    val co_algXFTs = @{map 2} mk_co_algT XFTs Xs;
 | 
|
729  | 
val Ts = mk_Ts As;  | 
|
730  | 
    val un_foldTs = @{map 2} (fn T => fn X => co_algXFTs ---> mk_co_algT T X) Ts Xs;
 | 
|
731  | 
    val un_folds = @{map 2} (force_typ names_lthy) un_foldTs un_folds0;
 | 
|
732  | 
val ABs = As ~~ Bs;  | 
|
733  | 
val XYs = Xs ~~ Ys;  | 
|
734  | 
||
735  | 
val Us = map (typ_subst_atomic ABs) Ts;  | 
|
736  | 
||
737  | 
    val TFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Ts)) Dss pre_bnfs;
 | 
|
738  | 
||
| 63045 | 739  | 
    val TFTs' = @{map 2} (absT_info_decodeT thy) absT_info_opts TFTs;
 | 
740  | 
    val xtors = @{map 3} (force_typ names_lthy oo mk_co_algT) TFTs' Ts xtors0;
 | 
|
| 62905 | 741  | 
|
742  | 
val ids = map HOLogic.id_const As;  | 
|
743  | 
    val co_rec_Xs = @{map 2} mk_co_productT Ts Xs;
 | 
|
744  | 
    val co_rec_Ys = @{map 2} mk_co_productT Us Ys;
 | 
|
745  | 
    val co_rec_algXs = @{map 2} mk_co_algT co_rec_Xs Xs;
 | 
|
746  | 
val co_proj1s = map co_proj1_const co_rec_algXs;  | 
|
747  | 
    val co_rec_maps = @{map 2} (fn Ds =>
 | 
|
748  | 
mk_map_of_bnf Ds (As @ case_fp fp co_rec_Xs Ts) (As @ case_fp fp Ts co_rec_Xs)) Dss pre_bnfs;  | 
|
749  | 
    val co_rec_Ts = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ co_rec_Xs)) Dss pre_bnfs
 | 
|
750  | 
    val co_rec_argTs = @{map 2} mk_co_algT co_rec_Ts Xs;
 | 
|
751  | 
    val co_rec_resTs = @{map 2} mk_co_algT Ts Xs;
 | 
|
752  | 
||
753  | 
val (((co_rec_ss, fs), xs), names_lthy) = names_lthy  | 
|
754  | 
|> mk_Frees "s" co_rec_argTs  | 
|
755  | 
||>> mk_Frees "f" co_rec_resTs  | 
|
| 63045 | 756  | 
||>> mk_Frees "x" (case_fp fp TFTs' Xs);  | 
| 62905 | 757  | 
|
758  | 
val co_rec_strs =  | 
|
| 63045 | 759  | 
      @{map 4} (fn xtor => fn s => fn mapx => fn absT_info_opt =>
 | 
760  | 
mk_co_product (mk_co_comp (absT_info_encode thy fp absT_info_opt xtor)  | 
|
761  | 
(list_comb (mapx, ids @ co_proj1s))) s)  | 
|
762  | 
xtors co_rec_ss co_rec_maps absT_info_opts;  | 
|
| 62905 | 763  | 
|
764  | 
val theta = Xs ~~ co_rec_Xs;  | 
|
765  | 
val co_rec_un_folds = map (subst_atomic_types theta) un_folds;  | 
|
766  | 
||
767  | 
val co_rec_spec0s = map (fn un_fold => list_comb (un_fold, co_rec_strs)) co_rec_un_folds;  | 
|
768  | 
||
769  | 
    val co_rec_ids = @{map 2} (mk_co_comp o co_proj1_const) co_rec_algXs co_rec_spec0s;
 | 
|
770  | 
    val co_rec_specs = @{map 2} (mk_co_comp o co_proj2_const) co_rec_algXs co_rec_spec0s;
 | 
|
771  | 
||
772  | 
val co_recN = case_fp fp ctor_recN dtor_corecN;  | 
|
773  | 
fun co_rec_bind i = nth bs (i - 1) |> Binding.prefix_name (co_recN ^ "_");  | 
|
774  | 
val co_rec_def_bind = rpair [] o Binding.concealed o Thm.def_binding o co_rec_bind;  | 
|
775  | 
||
776  | 
fun co_rec_spec i =  | 
|
777  | 
fold_rev (Term.absfree o Term.dest_Free) co_rec_ss (nth co_rec_specs (i - 1));  | 
|
778  | 
||
779  | 
val ((co_rec_frees, (_, co_rec_def_frees)), (lthy, lthy_old)) =  | 
|
780  | 
lthy  | 
|
| 72450 | 781  | 
|> (snd o Local_Theory.begin_nested)  | 
| 62905 | 782  | 
|> fold_map (fn i =>  | 
783  | 
Local_Theory.define ((co_rec_bind i, NoSyn), (co_rec_def_bind i, co_rec_spec i))) ks  | 
|
784  | 
|>> apsnd split_list o split_list  | 
|
| 72450 | 785  | 
||> `Local_Theory.end_nested;  | 
| 62905 | 786  | 
|
787  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
|
| 
80636
 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 
wenzelm 
parents: 
75624 
diff
changeset
 | 
788  | 
val co_rec_names = map (dest_Const_name o Morphism.term phi) co_rec_frees;  | 
| 62905 | 789  | 
    val co_recs = @{map 2} (fn name => fn resT =>
 | 
790  | 
Const (name, co_rec_argTs ---> resT)) co_rec_names co_rec_resTs;  | 
|
791  | 
val co_rec_defs = map (fn def =>  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67505 
diff
changeset
 | 
792  | 
mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) co_rec_def_frees;  | 
| 62905 | 793  | 
|
| 63045 | 794  | 
val xtor_un_fold_xtor_thms =  | 
795  | 
mk_xtor_un_fold_xtor_thms lthy fp (map (Term.subst_atomic_types (Xs ~~ Ts)) un_folds)  | 
|
796  | 
xtors xtor_un_fold_unique map_id0s absT_info_opts;  | 
|
| 62905 | 797  | 
|
798  | 
val co_rec_id_thms =  | 
|
799  | 
let  | 
|
800  | 
        val goal = @{map 2} (fn T => fn t => HOLogic.mk_eq (t, HOLogic.id_const T)) Ts co_rec_ids
 | 
|
801  | 
|> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop;  | 
|
802  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
803  | 
in  | 
|
804  | 
Goal.prove_sorry lthy vars [] goal  | 
|
805  | 
          (fn {context = ctxt, prems = _} => mk_xtor_co_rec_id_tac ctxt xtor_un_fold_xtor_thms
 | 
|
806  | 
xtor_un_fold_unique xtor_un_folds map_comps)  | 
|
| 70494 | 807  | 
|> Thm.close_derivation \<^here>  | 
| 63045 | 808  | 
|> split_conj_thm  | 
| 62905 | 809  | 
end;  | 
810  | 
||
811  | 
val co_rec_app_ss = map (fn co_rec => list_comb (co_rec, co_rec_ss)) co_recs;  | 
|
812  | 
    val co_products = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts co_rec_app_ss;
 | 
|
813  | 
    val co_rec_maps_rev = @{map 2} (fn Ds =>
 | 
|
814  | 
mk_map_of_bnf Ds (As @ case_fp fp Ts co_rec_Xs) (As @ case_fp fp co_rec_Xs Ts)) Dss pre_bnfs;  | 
|
815  | 
fun mk_co_app f g x = case_fp fp (f $ (g $ x)) (g $ (f $ x));  | 
|
816  | 
val co_rec_expand_thms = map (fn thm => thm RS  | 
|
817  | 
      case_fp fp @{thm convol_expand_snd} @{thm case_sum_expand_Inr_pointfree}) co_rec_id_thms;
 | 
|
818  | 
val xtor_co_rec_thms =  | 
|
819  | 
let  | 
|
| 63045 | 820  | 
fun mk_goal co_rec s mapx xtor x absT_info_opt =  | 
| 62905 | 821  | 
let  | 
822  | 
val lhs = mk_co_app co_rec xtor x;  | 
|
| 63045 | 823  | 
val rhs = mk_co_app s  | 
824  | 
(list_comb (mapx, ids @ co_products) |> absT_info_decode thy fp absT_info_opt) x;  | 
|
| 62905 | 825  | 
in  | 
826  | 
mk_Trueprop_eq (lhs, rhs)  | 
|
827  | 
end;  | 
|
| 63045 | 828  | 
val goals =  | 
829  | 
          @{map 6} mk_goal co_rec_app_ss co_rec_ss co_rec_maps_rev xtors xs absT_info_opts;
 | 
|
| 62905 | 830  | 
in  | 
831  | 
map2 (fn goal => fn un_fold =>  | 
|
832  | 
Variable.add_free_names lthy goal []  | 
|
833  | 
|> (fn vars => Goal.prove_sorry lthy vars [] goal  | 
|
834  | 
            (fn {context = ctxt, prems = _} =>
 | 
|
835  | 
mk_xtor_co_rec_tac ctxt un_fold co_rec_defs co_rec_expand_thms))  | 
|
| 70494 | 836  | 
|> Thm.close_derivation \<^here>)  | 
| 62905 | 837  | 
goals xtor_un_folds  | 
838  | 
end;  | 
|
839  | 
||
840  | 
    val co_product_fs = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts fs;
 | 
|
841  | 
val co_rec_expand'_thms = map (fn thm =>  | 
|
842  | 
      thm RS case_fp fp @{thm convol_expand_snd'} @{thm case_sum_expand_Inr'}) co_rec_id_thms;
 | 
|
843  | 
val xtor_co_rec_unique_thm =  | 
|
844  | 
let  | 
|
| 63045 | 845  | 
fun mk_prem f s mapx xtor absT_info_opt =  | 
| 62905 | 846  | 
let  | 
847  | 
val lhs = mk_co_comp f xtor;  | 
|
| 63045 | 848  | 
val rhs = mk_co_comp s (list_comb (mapx, ids @ co_product_fs))  | 
849  | 
|> absT_info_decode thy fp absT_info_opt;  | 
|
| 62905 | 850  | 
in  | 
851  | 
mk_Trueprop_eq (co_swap (lhs, rhs))  | 
|
852  | 
end;  | 
|
| 63045 | 853  | 
        val prems = @{map 5} mk_prem fs co_rec_ss co_rec_maps_rev xtors absT_info_opts;
 | 
| 62905 | 854  | 
        val concl = @{map 2} (curry HOLogic.mk_eq) fs co_rec_app_ss
 | 
855  | 
|> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop;  | 
|
856  | 
val goal = Logic.list_implies (prems, concl);  | 
|
857  | 
val vars = Variable.add_free_names lthy goal [];  | 
|
| 63045 | 858  | 
fun mk_inverses NONE = []  | 
859  | 
| mk_inverses (SOME (src, dst)) =  | 
|
860  | 
            [#type_definition dst RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp,
 | 
|
861  | 
             #type_definition src RS @{thm type_copy_Abs_o_Rep}];
 | 
|
862  | 
val inverses = maps mk_inverses absT_info_opts;  | 
|
| 62905 | 863  | 
in  | 
864  | 
Goal.prove_sorry lthy vars [] goal  | 
|
865  | 
          (fn {context = ctxt, prems = _} => mk_xtor_co_rec_unique_tac ctxt fp co_rec_defs
 | 
|
| 63045 | 866  | 
co_rec_expand'_thms xtor_un_fold_unique map_id0s sym_map_comp0s inverses)  | 
| 70494 | 867  | 
|> Thm.close_derivation \<^here>  | 
| 62905 | 868  | 
end;  | 
869  | 
||
| 63045 | 870  | 
val xtor_co_rec_o_map_thms = if forall is_none absT_info_opts  | 
871  | 
then  | 
|
872  | 
mk_xtor_co_iter_o_map_thms fp true m xtor_co_rec_unique_thm  | 
|
| 64413 | 873  | 
(map (mk_pointfree2 lthy) xtor_maps) (map (mk_pointfree2 lthy) xtor_co_rec_thms)  | 
| 63045 | 874  | 
sym_map_comp0s map_cong0s  | 
875  | 
else  | 
|
876  | 
replicate n refl (* FIXME *);  | 
|
| 62905 | 877  | 
|
878  | 
    val ABphiTs = @{map 2} mk_pred2T As Bs;
 | 
|
879  | 
    val XYphiTs = @{map 2} mk_pred2T Xs Ys;
 | 
|
880  | 
||
881  | 
val ((ABphis, XYphis), names_lthy) = names_lthy  | 
|
882  | 
|> mk_Frees "R" ABphiTs  | 
|
883  | 
||>> mk_Frees "S" XYphiTs;  | 
|
884  | 
||
| 63045 | 885  | 
val xtor_co_rec_transfer_thms = if forall is_none absT_info_opts  | 
886  | 
then  | 
|
887  | 
let  | 
|
888  | 
val pre_rels =  | 
|
889  | 
            @{map 2} (fn Ds => mk_rel_of_bnf Ds (As @ co_rec_Xs) (Bs @ co_rec_Ys)) Dss pre_bnfs;
 | 
|
890  | 
          val rels = @{map 3} (fn T => fn T' => Thm.prop_of #> HOLogic.dest_Trueprop
 | 
|
891  | 
#> fst o dest_comb #> fst o dest_comb #> funpow n (snd o dest_comb)  | 
|
892  | 
#> case_fp fp (fst o dest_comb #> snd o dest_comb) (snd o dest_comb) #> head_of  | 
|
893  | 
#> force_typ names_lthy (ABphiTs ---> mk_pred2T T T'))  | 
|
894  | 
Ts Us xtor_un_fold_transfers;  | 
|
| 63813 | 895  | 
|
| 63045 | 896  | 
          fun tac {context = ctxt, prems = _} = mk_xtor_co_rec_transfer_tac ctxt fp n m co_rec_defs
 | 
897  | 
xtor_un_fold_transfers map_transfers xtor_rels;  | 
|
| 63813 | 898  | 
|
| 63045 | 899  | 
val mk_rel_co_product = case_fp fp mk_rel_prod mk_rel_sum;  | 
900  | 
val rec_phis =  | 
|
901  | 
map2 (fn rel => mk_rel_co_product (Term.list_comb (rel, ABphis))) rels XYphis;  | 
|
902  | 
in  | 
|
903  | 
mk_xtor_co_iter_transfer_thms fp pre_rels rec_phis XYphis rels ABphis  | 
|
904  | 
co_recs (map (subst_atomic_types (ABs @ XYs)) co_recs) tac lthy  | 
|
905  | 
end  | 
|
906  | 
else  | 
|
907  | 
replicate n TrueI (* FIXME *);  | 
|
| 62905 | 908  | 
|
909  | 
val notes =  | 
|
910  | 
[(case_fp fp ctor_recN dtor_corecN, xtor_co_rec_thms),  | 
|
911  | 
(case_fp fp ctor_rec_uniqueN dtor_corec_uniqueN, split_conj_thm xtor_co_rec_unique_thm),  | 
|
912  | 
(case_fp fp ctor_rec_o_mapN dtor_corec_o_mapN, xtor_co_rec_o_map_thms),  | 
|
913  | 
(case_fp fp ctor_rec_transferN dtor_corec_transferN, xtor_co_rec_transfer_thms)]  | 
|
914  | 
|> map (apsnd (map single))  | 
|
915  | 
|> maps (fn (thmN, thmss) =>  | 
|
916  | 
map2 (fn b => fn thms =>  | 
|
917  | 
((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))  | 
|
918  | 
bs thmss);  | 
|
919  | 
||
920  | 
val lthy = lthy |> Config.get lthy bnf_internals ? snd o Local_Theory.notes notes;  | 
|
921  | 
in  | 
|
922  | 
((co_recs,  | 
|
923  | 
(xtor_co_rec_thms, xtor_co_rec_unique_thm, xtor_co_rec_o_map_thms, xtor_co_rec_transfer_thms)),  | 
|
924  | 
lthy)  | 
|
925  | 
end;  | 
|
926  | 
||
| 63796 | 927  | 
fun raw_qualify extra_qualify base_b =  | 
928  | 
let  | 
|
929  | 
val qs = Binding.path_of base_b;  | 
|
930  | 
val n = Binding.name_of base_b;  | 
|
931  | 
in  | 
|
932  | 
Binding.prefix_name rawN  | 
|
933  | 
#> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)])  | 
|
934  | 
#> extra_qualify #> Binding.concealed  | 
|
935  | 
end;  | 
|
936  | 
||
| 63802 | 937  | 
fun fixpoint_bnf force_out_of_line extra_qualify construct_fp bs resBs Ds0 Xs rhsXs comp_cache0  | 
938  | 
lthy =  | 
|
| 51868 | 939  | 
let  | 
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53138 
diff
changeset
 | 
940  | 
val time = time lthy;  | 
| 51868 | 941  | 
val timer = time (Timer.startRealTimer ());  | 
| 58332 | 942  | 
|
943  | 
fun flatten_tyargs Ass =  | 
|
| 53222 | 944  | 
subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs;  | 
| 51868 | 945  | 
|
| 63796 | 946  | 
val ((bnfs, (deadss, livess)), (comp_cache_unfold_set, lthy')) =  | 
| 55706 | 947  | 
apfst (apsnd split_list o split_list)  | 
| 
58634
 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 
wenzelm 
parents: 
58585 
diff
changeset
 | 
948  | 
        (@{fold_map 2}
 | 
| 63796 | 949  | 
(fn b => bnf_of_typ true Smart_Inline (raw_qualify extra_qualify b) flatten_tyargs Xs Ds0)  | 
950  | 
bs rhsXs ((comp_cache0, empty_unfolds), lthy));  | 
|
| 51868 | 951  | 
|
| 62722 | 952  | 
fun norm_qualify i =  | 
953  | 
Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1))))  | 
|
954  | 
#> extra_qualify #> Binding.concealed;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
955  | 
|
| 49132 | 956  | 
val Ass = map (map dest_TFree) livess;  | 
| 58332 | 957  | 
val Ds' = fold (fold Term.add_tfreesT) deadss [];  | 
958  | 
val Ds = union (op =) Ds' Ds0;  | 
|
959  | 
val missing = resBs |> fold (subtract (op =)) (Ds' :: Ass);  | 
|
960  | 
val (dead_phantoms, live_phantoms) = List.partition (member (op =) Ds0) missing;  | 
|
961  | 
val resBs' = resBs |> fold (subtract (op =)) [dead_phantoms, Ds];  | 
|
| 49132 | 962  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
963  | 
val timer = time (timer "Construction of BNFs");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
964  | 
|
| 
63837
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
965  | 
val ((kill_posss, _), (bnfs', ((comp_cache', unfold_set'), lthy''))) =  | 
| 63796 | 966  | 
normalize_bnfs norm_qualify Ass Ds (K (resBs' @ Xs)) bnfs (comp_cache_unfold_set, lthy');  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
967  | 
|
| 
63837
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
968  | 
    val Dss = @{map 3} (fn lives => fn kill_posss => fn deads => deads @ map (nth lives) kill_posss)
 | 
| 
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
969  | 
livess kill_posss deadss;  | 
| 
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
970  | 
val all_Dss = Dss |> force_out_of_line ? map (fn Ds' => union (op =) Ds' (map TFree Ds0));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
971  | 
|
| 62722 | 972  | 
fun pre_qualify b =  | 
973  | 
Binding.qualify false (Binding.name_of b)  | 
|
974  | 
#> extra_qualify  | 
|
| 63796 | 975  | 
#> not (Config.get lthy'' bnf_internals) ? Binding.concealed;  | 
| 53264 | 976  | 
|
| 63802 | 977  | 
val ((pre_bnfs, (deadss, absT_infos)), lthy''') = lthy''  | 
| 
63837
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
978  | 
      |> @{fold_map 5} (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b))
 | 
| 
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
979  | 
bs (replicate (length rhsXs) (force_out_of_line orelse not (null live_phantoms))) Dss  | 
| 
 
fdf90aa59868
provide a mechanism for ensuring dead type variables occur in typedef if desired
 
blanchet 
parents: 
63826 
diff
changeset
 | 
980  | 
all_Dss bnfs'  | 
| 
55803
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
981  | 
|>> split_list  | 
| 
 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 
traytel 
parents: 
55706 
diff
changeset
 | 
982  | 
|>> apsnd split_list;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
983  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
984  | 
val timer = time (timer "Normalization & sealing of BNFs");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
985  | 
|
| 63796 | 986  | 
val res = construct_fp bs resBs (map TFree dead_phantoms, deadss) pre_bnfs absT_infos lthy''';  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
987  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
988  | 
val timer = time (timer "FP construction in total");  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
989  | 
in  | 
| 62720 | 990  | 
(((pre_bnfs, absT_infos), comp_cache'), res)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
991  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
992  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
993  | 
end;  |