author | blanchet |
Mon, 26 May 2014 16:32:55 +0200 | |
changeset 57091 | 1fa9c19ba2c9 |
parent 57090 | 0224caba67ca |
child 57094 | 589ec121ce1a |
permissions | -rw-r--r-- |
54701 | 1 |
(* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML |
49017 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
54397 | 3 |
Copyright 2012, 2013 |
49017 | 4 |
|
51797 | 5 |
Wrapping existing freely generated type's constructors. |
49017 | 6 |
*) |
7 |
||
54006 | 8 |
signature CTR_SUGAR = |
49017 | 9 |
sig |
51840 | 10 |
type ctr_sugar = |
51839 | 11 |
{ctrs: term list, |
52375 | 12 |
casex: term, |
51839 | 13 |
discs: term list, |
51819 | 14 |
selss: term list list, |
15 |
exhaust: thm, |
|
52375 | 16 |
nchotomy: thm, |
51819 | 17 |
injects: thm list, |
18 |
distincts: thm list, |
|
19 |
case_thms: thm list, |
|
52375 | 20 |
case_cong: thm, |
21 |
weak_case_cong: thm, |
|
22 |
split: thm, |
|
23 |
split_asm: thm, |
|
56858 | 24 |
disc_defs: thm list, |
51819 | 25 |
disc_thmss: thm list list, |
26 |
discIs: thm list, |
|
56858 | 27 |
sel_defs: thm list, |
53475 | 28 |
sel_thmss: thm list list, |
54900 | 29 |
disc_excludesss: thm list list list, |
53475 | 30 |
disc_exhausts: thm list, |
53916 | 31 |
sel_exhausts: thm list, |
53475 | 32 |
collapses: thm list, |
33 |
expands: thm list, |
|
53917 | 34 |
sel_splits: thm list, |
35 |
sel_split_asms: thm list, |
|
54491 | 36 |
case_eq_ifs: thm list}; |
51809 | 37 |
|
51840 | 38 |
val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar |
54256 | 39 |
val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
40 |
val ctr_sugar_of: Proof.context -> string -> ctr_sugar option |
53906 | 41 |
val ctr_sugars_of: Proof.context -> ctr_sugar list |
54403 | 42 |
val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option |
56345 | 43 |
val ctr_sugar_interpretation: (ctr_sugar -> theory -> theory) -> theory -> theory |
54399 | 44 |
val register_ctr_sugar: string -> ctr_sugar -> local_theory -> local_theory |
55444
ec73f81e49e7
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
blanchet
parents:
55410
diff
changeset
|
45 |
val default_register_ctr_sugar_global: string -> ctr_sugar -> theory -> theory |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
46 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
47 |
val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
48 |
val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
49 |
|
49203 | 50 |
val mk_ctr: typ list -> term -> term |
53864
a48d4bd3faaa
use case rather than sequence of ifs in expansion
blanchet
parents:
53857
diff
changeset
|
51 |
val mk_case: typ list -> typ -> term -> term |
49484 | 52 |
val mk_disc_or_sel: typ list -> term -> term |
49622 | 53 |
val name_of_ctr: term -> string |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
54 |
val name_of_disc: term -> string |
53888 | 55 |
val dest_ctr: Proof.context -> string -> term -> term * term list |
54970
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
blanchet
parents:
54900
diff
changeset
|
56 |
val dest_case: Proof.context -> string -> typ list -> term -> |
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
blanchet
parents:
54900
diff
changeset
|
57 |
(ctr_sugar * term list * term list) option |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
58 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
59 |
type ('c, 'a, 'v) ctr_spec = ((binding * 'c) * 'a list) * (binding * 'v) list |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
60 |
|
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
61 |
val disc_of_ctr_spec: ('c, 'a, 'v) ctr_spec -> binding |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
62 |
val ctr_of_ctr_spec: ('c, 'a, 'v) ctr_spec -> 'c |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
63 |
val args_of_ctr_spec: ('c, 'a, 'v) ctr_spec -> 'a list |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
64 |
val sel_defaults_of_ctr_spec: ('c, 'a, 'v) ctr_spec -> (binding * 'v) list |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
65 |
|
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
66 |
val free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list -> |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
67 |
((bool * bool) * binding) * (term, binding, term) ctr_spec list -> local_theory -> |
51840 | 68 |
ctr_sugar * local_theory |
49286 | 69 |
val parse_bound_term: (binding * string) parser |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
70 |
val parse_ctr_options: (bool * bool) parser |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
71 |
val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a, string) ctr_spec parser |
49017 | 72 |
end; |
73 |
||
54006 | 74 |
structure Ctr_Sugar : CTR_SUGAR = |
49017 | 75 |
struct |
76 |
||
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
77 |
open Ctr_Sugar_Util |
54006 | 78 |
open Ctr_Sugar_Tactics |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
54493
diff
changeset
|
79 |
open Ctr_Sugar_Code |
49017 | 80 |
|
51840 | 81 |
type ctr_sugar = |
51839 | 82 |
{ctrs: term list, |
52375 | 83 |
casex: term, |
51839 | 84 |
discs: term list, |
51819 | 85 |
selss: term list list, |
86 |
exhaust: thm, |
|
52375 | 87 |
nchotomy: thm, |
51819 | 88 |
injects: thm list, |
89 |
distincts: thm list, |
|
90 |
case_thms: thm list, |
|
52375 | 91 |
case_cong: thm, |
92 |
weak_case_cong: thm, |
|
93 |
split: thm, |
|
94 |
split_asm: thm, |
|
56858 | 95 |
disc_defs: thm list, |
51819 | 96 |
disc_thmss: thm list list, |
97 |
discIs: thm list, |
|
56858 | 98 |
sel_defs: thm list, |
53475 | 99 |
sel_thmss: thm list list, |
54900 | 100 |
disc_excludesss: thm list list list, |
53475 | 101 |
disc_exhausts: thm list, |
53916 | 102 |
sel_exhausts: thm list, |
53475 | 103 |
collapses: thm list, |
104 |
expands: thm list, |
|
53917 | 105 |
sel_splits: thm list, |
106 |
sel_split_asms: thm list, |
|
54491 | 107 |
case_eq_ifs: thm list}; |
51809 | 108 |
|
52375 | 109 |
fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, |
56858 | 110 |
case_thms, case_cong, weak_case_cong, split, split_asm, disc_defs, disc_thmss, discIs, sel_defs, |
111 |
sel_thmss, disc_excludesss, disc_exhausts, sel_exhausts, collapses, expands, sel_splits, |
|
112 |
sel_split_asms, case_eq_ifs} = |
|
51839 | 113 |
{ctrs = map (Morphism.term phi) ctrs, |
52375 | 114 |
casex = Morphism.term phi casex, |
51839 | 115 |
discs = map (Morphism.term phi) discs, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
116 |
selss = map (map (Morphism.term phi)) selss, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
117 |
exhaust = Morphism.thm phi exhaust, |
52375 | 118 |
nchotomy = Morphism.thm phi nchotomy, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
119 |
injects = map (Morphism.thm phi) injects, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
120 |
distincts = map (Morphism.thm phi) distincts, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
121 |
case_thms = map (Morphism.thm phi) case_thms, |
52375 | 122 |
case_cong = Morphism.thm phi case_cong, |
123 |
weak_case_cong = Morphism.thm phi weak_case_cong, |
|
124 |
split = Morphism.thm phi split, |
|
125 |
split_asm = Morphism.thm phi split_asm, |
|
56858 | 126 |
disc_defs = map (Morphism.thm phi) disc_defs, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
127 |
disc_thmss = map (map (Morphism.thm phi)) disc_thmss, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
128 |
discIs = map (Morphism.thm phi) discIs, |
56858 | 129 |
sel_defs = map (Morphism.thm phi) sel_defs, |
53475 | 130 |
sel_thmss = map (map (Morphism.thm phi)) sel_thmss, |
54900 | 131 |
disc_excludesss = map (map (map (Morphism.thm phi))) disc_excludesss, |
53475 | 132 |
disc_exhausts = map (Morphism.thm phi) disc_exhausts, |
53916 | 133 |
sel_exhausts = map (Morphism.thm phi) sel_exhausts, |
53475 | 134 |
collapses = map (Morphism.thm phi) collapses, |
135 |
expands = map (Morphism.thm phi) expands, |
|
53917 | 136 |
sel_splits = map (Morphism.thm phi) sel_splits, |
137 |
sel_split_asms = map (Morphism.thm phi) sel_split_asms, |
|
54491 | 138 |
case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
139 |
|
53906 | 140 |
val transfer_ctr_sugar = |
54740 | 141 |
morph_ctr_sugar o Morphism.transfer_morphism o Proof_Context.theory_of; |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
142 |
|
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
143 |
structure Data = Generic_Data |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
144 |
( |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
145 |
type T = ctr_sugar Symtab.table; |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
146 |
val empty = Symtab.empty; |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
147 |
val extend = I; |
55394
d5ebe055b599
more liberal merging of BNFs and constructor sugar
blanchet
parents:
55356
diff
changeset
|
148 |
fun merge data : T = Symtab.merge (K true) data; |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
149 |
); |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
150 |
|
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
151 |
fun ctr_sugar_of ctxt = |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
152 |
Symtab.lookup (Data.get (Context.Proof ctxt)) |
53906 | 153 |
#> Option.map (transfer_ctr_sugar ctxt); |
154 |
||
155 |
fun ctr_sugars_of ctxt = |
|
156 |
Symtab.fold (cons o transfer_ctr_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) []; |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
157 |
|
54403 | 158 |
fun ctr_sugar_of_case ctxt s = |
159 |
find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of ctxt); |
|
160 |
||
56345 | 161 |
structure Ctr_Sugar_Interpretation = Interpretation |
162 |
( |
|
163 |
type T = ctr_sugar; |
|
164 |
val eq: T * T -> bool = op = o pairself #ctrs; |
|
165 |
); |
|
166 |
||
56376
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents:
56345
diff
changeset
|
167 |
fun with_repaired_path f (ctr_sugar as {ctrs = ctr1 :: _, ...} : ctr_sugar) thy = |
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents:
56345
diff
changeset
|
168 |
thy |
56767 | 169 |
|> Sign.root_path |
56376
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents:
56345
diff
changeset
|
170 |
|> Sign.add_path (Long_Name.qualifier (fst (dest_Type (body_type (fastype_of ctr1))))) |
56767 | 171 |
|> (fn thy => f (morph_ctr_sugar (Morphism.transfer_morphism thy) ctr_sugar) thy) |
172 |
|> Sign.restore_naming thy; |
|
56376
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents:
56345
diff
changeset
|
173 |
|
56657 | 174 |
fun ctr_sugar_interpretation f = Ctr_Sugar_Interpretation.interpretation (with_repaired_path f); |
56345 | 175 |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
176 |
fun register_ctr_sugar key ctr_sugar = |
54285
578371ba74cc
reverted 3e1d230f1c00 -- pervasiveness is useful, cf. Coinductive_Nat in the AFP
blanchet
parents:
54265
diff
changeset
|
177 |
Local_Theory.declaration {syntax = false, pervasive = true} |
56345 | 178 |
(fn phi => Data.map (Symtab.update (key, morph_ctr_sugar phi ctr_sugar))) |
179 |
#> Local_Theory.background_theory (Ctr_Sugar_Interpretation.data ctr_sugar); |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
180 |
|
56345 | 181 |
fun default_register_ctr_sugar_global key ctr_sugar thy = |
182 |
let val tab = Data.get (Context.Theory thy) in |
|
183 |
if Symtab.defined tab key then |
|
184 |
thy |
|
185 |
else |
|
186 |
thy |
|
187 |
|> Context.theory_map (Data.put (Symtab.update_new (key, ctr_sugar) tab)) |
|
188 |
|> Ctr_Sugar_Interpretation.data ctr_sugar |
|
189 |
end; |
|
54400 | 190 |
|
49223 | 191 |
val isN = "is_"; |
192 |
val unN = "un_"; |
|
193 |
fun mk_unN 1 1 suf = unN ^ suf |
|
194 |
| mk_unN _ l suf = unN ^ suf ^ string_of_int l; |
|
49046
3c5eba97d93a
allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents:
49045
diff
changeset
|
195 |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49591
diff
changeset
|
196 |
val caseN = "case"; |
49054 | 197 |
val case_congN = "case_cong"; |
54491 | 198 |
val case_eq_ifN = "case_eq_if"; |
49118 | 199 |
val collapseN = "collapse"; |
49122 | 200 |
val disc_excludeN = "disc_exclude"; |
49054 | 201 |
val disc_exhaustN = "disc_exhaust"; |
53694 | 202 |
val discN = "disc"; |
53700 | 203 |
val discIN = "discI"; |
49054 | 204 |
val distinctN = "distinct"; |
49075 | 205 |
val exhaustN = "exhaust"; |
49486 | 206 |
val expandN = "expand"; |
49075 | 207 |
val injectN = "inject"; |
208 |
val nchotomyN = "nchotomy"; |
|
53694 | 209 |
val selN = "sel"; |
53916 | 210 |
val sel_exhaustN = "sel_exhaust"; |
53917 | 211 |
val sel_splitN = "sel_split"; |
212 |
val sel_split_asmN = "sel_split_asm"; |
|
49054 | 213 |
val splitN = "split"; |
49633 | 214 |
val splitsN = "splits"; |
49054 | 215 |
val split_asmN = "split_asm"; |
216 |
val weak_case_cong_thmsN = "weak_case_cong"; |
|
49019 | 217 |
|
53700 | 218 |
val cong_attrs = @{attributes [cong]}; |
53836 | 219 |
val dest_attrs = @{attributes [dest]}; |
53700 | 220 |
val safe_elim_attrs = @{attributes [elim!]}; |
221 |
val iff_attrs = @{attributes [iff]}; |
|
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
222 |
val inductsimp_attrs = @{attributes [induct_simp]}; |
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
223 |
val nitpicksimp_attrs = @{attributes [nitpick_simp]}; |
49300 | 224 |
val simp_attrs = @{attributes [simp]}; |
54151 | 225 |
val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs; |
226 |
val code_nitpicksimp_simp_attrs = code_nitpicksimp_attrs @ simp_attrs; |
|
49046
3c5eba97d93a
allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents:
49045
diff
changeset
|
227 |
|
55480
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
blanchet
parents:
55471
diff
changeset
|
228 |
fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); |
49258
84f13469d7f0
allow same selector name for several constructors
blanchet
parents:
49257
diff
changeset
|
229 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
230 |
fun mk_half_pairss' _ ([], []) = [] |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
231 |
| mk_half_pairss' indent (x :: xs, _ :: ys) = |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
232 |
indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); |
49486 | 233 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
234 |
fun mk_half_pairss p = mk_half_pairss' [[]] p; |
49027 | 235 |
|
49486 | 236 |
fun join_halves n half_xss other_half_xss = |
55342 | 237 |
(splice (flat half_xss) (flat other_half_xss), |
238 |
map2 (map2 append) (Library.chop_groups n half_xss) |
|
239 |
(transpose (Library.chop_groups n other_half_xss))); |
|
49027 | 240 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
241 |
fun mk_undefined T = Const (@{const_name undefined}, T); |
49055 | 242 |
|
49500 | 243 |
fun mk_ctr Ts t = |
244 |
let val Type (_, Ts0) = body_type (fastype_of t) in |
|
54435
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
blanchet
parents:
54422
diff
changeset
|
245 |
subst_nonatomic_types (Ts0 ~~ Ts) t |
49203 | 246 |
end; |
247 |
||
49536 | 248 |
fun mk_case Ts T t = |
249 |
let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in |
|
54435
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
blanchet
parents:
54422
diff
changeset
|
250 |
subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t |
49536 | 251 |
end; |
252 |
||
53864
a48d4bd3faaa
use case rather than sequence of ifs in expansion
blanchet
parents:
53857
diff
changeset
|
253 |
fun mk_disc_or_sel Ts t = |
54435
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
blanchet
parents:
54422
diff
changeset
|
254 |
subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; |
53864
a48d4bd3faaa
use case rather than sequence of ifs in expansion
blanchet
parents:
53857
diff
changeset
|
255 |
|
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
256 |
fun name_of_const what t = |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
257 |
(case head_of t of |
49622 | 258 |
Const (s, _) => s |
259 |
| Free (s, _) => s |
|
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
260 |
| _ => error ("Cannot extract name of " ^ what)); |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
261 |
|
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
262 |
val name_of_ctr = name_of_const "constructor"; |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
263 |
|
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
264 |
val notN = "not_"; |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
265 |
val eqN = "eq_"; |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
266 |
val neqN = "neq_"; |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
267 |
|
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
268 |
fun name_of_disc t = |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
269 |
(case head_of t of |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
270 |
Abs (_, _, @{const Not} $ (t' $ Bound 0)) => |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
271 |
Long_Name.map_base_name (prefix notN) (name_of_disc t') |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
272 |
| Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') => |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
273 |
Long_Name.map_base_name (prefix eqN) (name_of_disc t') |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
274 |
| Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) => |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
275 |
Long_Name.map_base_name (prefix neqN) (name_of_disc t') |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
276 |
| t' => name_of_const "destructor" t'); |
49622 | 277 |
|
278 |
val base_name_of_ctr = Long_Name.base_name o name_of_ctr; |
|
49046
3c5eba97d93a
allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents:
49045
diff
changeset
|
279 |
|
53888 | 280 |
fun dest_ctr ctxt s t = |
55342 | 281 |
let val (f, args) = Term.strip_comb t in |
53888 | 282 |
(case ctr_sugar_of ctxt s of |
283 |
SOME {ctrs, ...} => |
|
284 |
(case find_first (can (fo_match ctxt f)) ctrs of |
|
285 |
SOME f' => (f', args) |
|
286 |
| NONE => raise Fail "dest_ctr") |
|
287 |
| NONE => raise Fail "dest_ctr") |
|
288 |
end; |
|
289 |
||
53870 | 290 |
fun dest_case ctxt s Ts t = |
291 |
(case Term.strip_comb t of |
|
292 |
(Const (c, _), args as _ :: _) => |
|
293 |
(case ctr_sugar_of ctxt s of |
|
54970
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
blanchet
parents:
54900
diff
changeset
|
294 |
SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => |
53870 | 295 |
if case_name = c then |
53924 | 296 |
let val n = length discs0 in |
297 |
if n < length args then |
|
298 |
let |
|
299 |
val (branches, obj :: leftovers) = chop n args; |
|
300 |
val discs = map (mk_disc_or_sel Ts) discs0; |
|
301 |
val selss = map (map (mk_disc_or_sel Ts)) selss0; |
|
302 |
val conds = map (rapp obj) discs; |
|
303 |
val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; |
|
304 |
val branches' = map2 (curry Term.betapplys) branches branch_argss; |
|
305 |
in |
|
54970
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
blanchet
parents:
54900
diff
changeset
|
306 |
SOME (ctr_sugar, conds, branches') |
53924 | 307 |
end |
308 |
else |
|
309 |
NONE |
|
53870 | 310 |
end |
311 |
else |
|
312 |
NONE |
|
313 |
| _ => NONE) |
|
314 |
| _ => NONE); |
|
315 |
||
49437 | 316 |
fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs; |
317 |
||
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
318 |
type ('c, 'a, 'v) ctr_spec = ((binding * 'c) * 'a list) * (binding * 'v) list; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
319 |
|
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
320 |
fun disc_of_ctr_spec (((disc, _), _), _) = disc; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
321 |
fun ctr_of_ctr_spec (((_, ctr), _), _) = ctr; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
322 |
fun args_of_ctr_spec ((_, args), _) = args; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
323 |
fun sel_defaults_of_ctr_spec (_, ds) = ds; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
324 |
|
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
325 |
fun prepare_free_constructors prep_term (((no_discs_sels, no_code), raw_case_binding), ctr_specs) |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
326 |
no_defs_lthy = |
49017 | 327 |
let |
49019 | 328 |
(* TODO: sanity checks on arguments *) |
49025 | 329 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
330 |
val raw_ctrs = map ctr_of_ctr_spec ctr_specs; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
331 |
val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
332 |
val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
333 |
val raw_sel_defaultss = map sel_defaults_of_ctr_spec ctr_specs; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
334 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
335 |
val n = length raw_ctrs; |
49054 | 336 |
val ks = 1 upto n; |
337 |
||
49121 | 338 |
val _ = if n > 0 then () else error "No constructors specified"; |
339 |
||
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
340 |
val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
341 |
val sel_defaultss = map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss; |
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
342 |
|
53908 | 343 |
val Type (fcT_name, As0) = body_type (fastype_of (hd ctrs0)); |
344 |
val fc_b_name = Long_Name.base_name fcT_name; |
|
345 |
val fc_b = Binding.name fc_b_name; |
|
49055 | 346 |
|
55410 | 347 |
fun qualify mandatory = Binding.qualify mandatory fc_b_name; |
49633 | 348 |
|
54386 | 349 |
fun dest_TFree_or_TVar (TFree sS) = sS |
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
350 |
| dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S) |
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
351 |
| dest_TFree_or_TVar _ = error "Invalid type argument"; |
52965
0cd62cb233e0
handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents:
52963
diff
changeset
|
352 |
|
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
353 |
val (unsorted_As, B) = |
49055 | 354 |
no_defs_lthy |
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
355 |
|> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) |
49055 | 356 |
||> the_single o fst o mk_TFrees 1; |
357 |
||
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
358 |
val As = map2 (resort_tfree o snd o dest_TFree_or_TVar) As0 unsorted_As; |
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
359 |
|
53908 | 360 |
val fcT = Type (fcT_name, As); |
49055 | 361 |
val ctrs = map (mk_ctr As) ctrs0; |
362 |
val ctr_Tss = map (binder_types o fastype_of) ctrs; |
|
363 |
||
364 |
val ms = map length ctr_Tss; |
|
365 |
||
57091 | 366 |
fun can_definitely_rely_on_disc k = |
367 |
not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; |
|
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
368 |
fun can_rely_on_disc k = |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
369 |
can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); |
53925 | 370 |
fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
371 |
|
57090
0224caba67ca
use '%x. x = C' as default discriminator for nullary constructor C, instead of relying on odd '=:' syntax
blanchet
parents:
57038
diff
changeset
|
372 |
val equal_binding = @{binding "="}; |
0224caba67ca
use '%x. x = C' as default discriminator for nullary constructor C, instead of relying on odd '=:' syntax
blanchet
parents:
57038
diff
changeset
|
373 |
|
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
374 |
fun is_disc_binding_valid b = |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
375 |
not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); |
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
376 |
|
52322 | 377 |
val standard_disc_binding = Binding.name o prefix isN o base_name_of_ctr; |
49120
7f8e69fc6ac9
smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents:
49119
diff
changeset
|
378 |
|
49336 | 379 |
val disc_bindings = |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
380 |
raw_disc_bindings |
49120
7f8e69fc6ac9
smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents:
49119
diff
changeset
|
381 |
|> map4 (fn k => fn m => fn ctr => fn disc => |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
382 |
qualify false |
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
383 |
(if Binding.is_empty disc then |
57091 | 384 |
if m = 0 then equal_binding |
385 |
else if should_omit_disc_binding k then disc |
|
57090
0224caba67ca
use '%x. x = C' as default discriminator for nullary constructor C, instead of relying on odd '=:' syntax
blanchet
parents:
57038
diff
changeset
|
386 |
else standard_disc_binding ctr |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
387 |
else if Binding.eq_name (disc, standard_binding) then |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
388 |
standard_disc_binding ctr |
49302
f5bd87aac224
added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents:
49300
diff
changeset
|
389 |
else |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
390 |
disc)) ks ms ctrs0; |
49056 | 391 |
|
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
392 |
fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; |
49120
7f8e69fc6ac9
smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents:
49119
diff
changeset
|
393 |
|
49336 | 394 |
val sel_bindingss = |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
395 |
map3 (fn ctr => fn m => map2 (fn l => fn sel => |
49633 | 396 |
qualify false |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
397 |
(if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then |
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
398 |
standard_sel_binding m l ctr |
49302
f5bd87aac224
added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents:
49300
diff
changeset
|
399 |
else |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
400 |
sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
401 |
|
49201 | 402 |
val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; |
49043 | 403 |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
404 |
val (((((((([exh_y], (xss, xss')), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
405 |
no_defs_lthy |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
406 |
|> mk_Frees "y" [fcT] (* for compatibility with "datatype_realizer.ML" *) |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
407 |
||>> mk_Freess' "x" ctr_Tss |
49025 | 408 |
||>> mk_Freess "y" ctr_Tss |
49201 | 409 |
||>> mk_Frees "f" case_Ts |
410 |
||>> mk_Frees "g" case_Ts |
|
53908 | 411 |
||>> (apfst (map (rpair fcT)) oo Variable.variant_fixes) [fc_b_name, fc_b_name ^ "'"] |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
412 |
||>> mk_Frees "z" [B] |
49043 | 413 |
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; |
414 |
||
49498 | 415 |
val u = Free u'; |
416 |
val v = Free v'; |
|
49463 | 417 |
val q = Free (fst p', mk_pred1T B); |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
418 |
|
49025 | 419 |
val xctrs = map2 (curry Term.list_comb) ctrs xss; |
420 |
val yctrs = map2 (curry Term.list_comb) ctrs yss; |
|
49032 | 421 |
|
49043 | 422 |
val xfs = map2 (curry Term.list_comb) fs xss; |
423 |
val xgs = map2 (curry Term.list_comb) gs xss; |
|
424 |
||
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
425 |
(* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
426 |
nicer names). Consider removing. *) |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
427 |
val eta_fs = map2 eta_expand_arg xss xfs; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
428 |
val eta_gs = map2 eta_expand_arg xss xgs; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
429 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
430 |
val case_binding = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
431 |
qualify false |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
432 |
(if Binding.is_empty raw_case_binding orelse |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
433 |
Binding.eq_name (raw_case_binding, standard_binding) then |
54493
5b34a5b93ec2
use type suffixes instead of prefixes for 'case', '(un)fold', '(co)rec'
blanchet
parents:
54491
diff
changeset
|
434 |
Binding.prefix_name (caseN ^ "_") fc_b |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
435 |
else |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
436 |
raw_case_binding); |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
437 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
438 |
fun mk_case_disj xctr xf xs = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
439 |
list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
440 |
|
53925 | 441 |
val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] |
442 |
(Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $ |
|
443 |
Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss))); |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
444 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
445 |
val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy |
54155 | 446 |
|> Local_Theory.define ((case_binding, NoSyn), |
447 |
((Binding.conceal (Thm.def_binding case_binding), []), case_rhs)) |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
448 |
||> `Local_Theory.restore; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
449 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
450 |
val phi = Proof_Context.export_morphism lthy lthy'; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
451 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
452 |
val case_def = Morphism.thm phi raw_case_def; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
453 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
454 |
val case0 = Morphism.term phi raw_case; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
455 |
val casex = mk_case As B case0; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
456 |
|
51759
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
457 |
val fcase = Term.list_comb (casex, fs); |
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
458 |
|
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
459 |
val ufcase = fcase $ u; |
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
460 |
val vfcase = fcase $ v; |
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
461 |
|
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
462 |
val eta_fcase = Term.list_comb (casex, eta_fs); |
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
463 |
val eta_gcase = Term.list_comb (casex, eta_gs); |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
464 |
|
51759
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
465 |
val eta_ufcase = eta_fcase $ u; |
587bba425430
eta-contracted weak congruence rules (like in the old package)
blanchet
parents:
51757
diff
changeset
|
466 |
val eta_vgcase = eta_gcase $ v; |
49486 | 467 |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
468 |
fun mk_uu_eq () = HOLogic.mk_eq (u, u); |
49486 | 469 |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
470 |
val uv_eq = mk_Trueprop_eq (u, v); |
49486 | 471 |
|
472 |
val exist_xs_u_eq_ctrs = |
|
473 |
map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; |
|
49022 | 474 |
|
51743 | 475 |
val unique_disc_no_def = TrueI; (*arbitrary marker*) |
476 |
val alternate_disc_no_def = FalseE; (*arbitrary marker*) |
|
477 |
||
49486 | 478 |
fun alternate_disc_lhs get_udisc k = |
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
479 |
HOLogic.mk_not |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
480 |
(let val b = nth disc_bindings (k - 1) in |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
481 |
if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
482 |
end); |
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
483 |
|
53917 | 484 |
val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') = |
52969
f2df0730f8ac
clarified option name (since case/fold/rec are also destructors)
blanchet
parents:
52968
diff
changeset
|
485 |
if no_discs_sels then |
55407
81f7e60755c3
use right local theory -- shows up when 'no_discs_sels' is set
blanchet
parents:
55394
diff
changeset
|
486 |
(true, [], [], [], [], [], lthy') |
49278 | 487 |
else |
488 |
let |
|
53908 | 489 |
fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); |
49025 | 490 |
|
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
491 |
fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
492 |
|
49500 | 493 |
fun alternate_disc k = |
494 |
Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); |
|
49278 | 495 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
496 |
fun mk_sel_case_args b proto_sels T = |
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
497 |
map2 (fn Ts => fn k => |
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
498 |
(case AList.lookup (op =) proto_sels k of |
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
499 |
NONE => |
49364
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
500 |
(case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of |
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
501 |
NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
502 |
| SOME t => t |> Type.constraint (Ts ---> T) |> Syntax.check_term lthy) |
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
503 |
| SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks; |
49258
84f13469d7f0
allow same selector name for several constructors
blanchet
parents:
49257
diff
changeset
|
504 |
|
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
505 |
fun sel_spec b proto_sels = |
49278 | 506 |
let |
507 |
val _ = |
|
508 |
(case duplicates (op =) (map fst proto_sels) of |
|
509 |
k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ |
|
510 |
" for constructor " ^ |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
511 |
quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) |
49278 | 512 |
| [] => ()) |
513 |
val T = |
|
514 |
(case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of |
|
515 |
[T] => T |
|
516 |
| T :: T' :: _ => error ("Inconsistent range type for selector " ^ |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
517 |
quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
518 |
^ quote (Syntax.string_of_typ lthy T'))); |
49278 | 519 |
in |
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
520 |
mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
521 |
Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) |
49278 | 522 |
end; |
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
523 |
|
49336 | 524 |
val sel_bindings = flat sel_bindingss; |
525 |
val uniq_sel_bindings = distinct Binding.eq_name sel_bindings; |
|
526 |
val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings); |
|
49285
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
527 |
|
49336 | 528 |
val sel_binding_index = |
529 |
if all_sels_distinct then 1 upto length sel_bindings |
|
530 |
else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings; |
|
49285
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
531 |
|
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
532 |
val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss); |
49300 | 533 |
val sel_infos = |
49336 | 534 |
AList.group (op =) (sel_binding_index ~~ proto_sels) |
49285
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
535 |
|> sort (int_ord o pairself fst) |
49336 | 536 |
|> map snd |> curry (op ~~) uniq_sel_bindings; |
537 |
val sel_bindings = map fst sel_infos; |
|
49258
84f13469d7f0
allow same selector name for several constructors
blanchet
parents:
49257
diff
changeset
|
538 |
|
49336 | 539 |
fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; |
49258
84f13469d7f0
allow same selector name for several constructors
blanchet
parents:
49257
diff
changeset
|
540 |
|
49278 | 541 |
val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
542 |
lthy |
51809 | 543 |
|> apfst split_list o fold_map3 (fn k => fn exist_xs_u_eq_ctr => fn b => |
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
544 |
if Binding.is_empty b then |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
545 |
if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
546 |
else pair (alternate_disc k, alternate_disc_no_def) |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
547 |
else if Binding.eq_name (b, equal_binding) then |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
548 |
pair (Term.lambda u exist_xs_u_eq_ctr, refl) |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
549 |
else |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
550 |
Specification.definition (SOME (b, NONE, NoSyn), |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
551 |
((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd) |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
552 |
ks exist_xs_u_eq_ctrs disc_bindings |
49278 | 553 |
||>> apfst split_list o fold_map (fn (b, proto_sels) => |
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
554 |
Specification.definition (SOME (b, NONE, NoSyn), |
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
555 |
((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos |
49278 | 556 |
||> `Local_Theory.restore; |
49022 | 557 |
|
49278 | 558 |
val phi = Proof_Context.export_morphism lthy lthy'; |
49025 | 559 |
|
49278 | 560 |
val disc_defs = map (Morphism.thm phi) raw_disc_defs; |
49281 | 561 |
val sel_defs = map (Morphism.thm phi) raw_sel_defs; |
562 |
val sel_defss = unflat_selss sel_defs; |
|
49278 | 563 |
|
564 |
val discs0 = map (Morphism.term phi) raw_discs; |
|
565 |
val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); |
|
49028 | 566 |
|
49278 | 567 |
val discs = map (mk_disc_or_sel As) discs0; |
568 |
val selss = map (map (mk_disc_or_sel As)) selss0; |
|
569 |
in |
|
53917 | 570 |
(all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') |
49278 | 571 |
end; |
49025 | 572 |
|
49032 | 573 |
fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); |
49029 | 574 |
|
49458 | 575 |
val exhaust_goal = |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
576 |
let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
577 |
fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
578 |
end; |
49019 | 579 |
|
49484 | 580 |
val inject_goalss = |
49017 | 581 |
let |
49034
b77e1910af8a
make parallel list indexing possible for inject theorems
blanchet
parents:
49033
diff
changeset
|
582 |
fun mk_goal _ _ [] [] = [] |
49025 | 583 |
| mk_goal xctr yctr xs ys = |
49121 | 584 |
[fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), |
585 |
Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; |
|
49017 | 586 |
in |
49034
b77e1910af8a
make parallel list indexing possible for inject theorems
blanchet
parents:
49033
diff
changeset
|
587 |
map4 mk_goal xctrs yctrs xss yss |
49017 | 588 |
end; |
589 |
||
49484 | 590 |
val half_distinct_goalss = |
49121 | 591 |
let |
49203 | 592 |
fun mk_goal ((xs, xc), (xs', xc')) = |
49121 | 593 |
fold_rev Logic.all (xs @ xs') |
49203 | 594 |
(HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); |
49121 | 595 |
in |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
596 |
map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) |
49121 | 597 |
end; |
49019 | 598 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
599 |
val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; |
49019 | 600 |
|
601 |
fun after_qed thmss lthy = |
|
602 |
let |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
603 |
val ([exhaust_thm0], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss)); |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
604 |
(* for "datatype_realizer.ML": *) |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
605 |
val exhaust_thm = |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
606 |
Thm.name_derivation (fcT_name ^ Long_Name.separator ^ exhaustN) exhaust_thm0; |
49019 | 607 |
|
49438 | 608 |
val inject_thms = flat inject_thmss; |
609 |
||
53210
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
blanchet
parents:
53040
diff
changeset
|
610 |
val rho_As = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As); |
49486 | 611 |
|
612 |
fun inst_thm t thm = |
|
613 |
Drule.instantiate' [] [SOME (certify lthy t)] |
|
53210
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
blanchet
parents:
53040
diff
changeset
|
614 |
(Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); |
49486 | 615 |
|
616 |
val uexhaust_thm = inst_thm u exhaust_thm; |
|
49032 | 617 |
|
49300 | 618 |
val exhaust_cases = map base_name_of_ctr ctrs; |
619 |
||
49048
4e0f0f98be02
rationalized data structure for distinctness theorems
blanchet
parents:
49046
diff
changeset
|
620 |
val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; |
4e0f0f98be02
rationalized data structure for distinctness theorems
blanchet
parents:
49046
diff
changeset
|
621 |
|
49486 | 622 |
val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
623 |
join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; |
49019 | 624 |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
625 |
val nchotomy_thm = |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
626 |
let |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
627 |
val goal = |
49486 | 628 |
HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', |
629 |
Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
630 |
in |
51551 | 631 |
Goal.prove_sorry lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm) |
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
632 |
|> Thm.close_derivation |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
633 |
end; |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
634 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
635 |
val case_thms = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
636 |
let |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
637 |
val goals = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
638 |
map3 (fn xctr => fn xf => fn xs => |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
639 |
fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
640 |
in |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
641 |
map4 (fn k => fn goal => fn injects => fn distinctss => |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
642 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
643 |
mk_case_tac ctxt n k case_def injects distinctss) |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
644 |
|> Thm.close_derivation) |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
645 |
ks goals inject_thmss distinct_thmsss |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
646 |
end; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
647 |
|
53917 | 648 |
val (case_cong_thm, weak_case_cong_thm) = |
649 |
let |
|
650 |
fun mk_prem xctr xs xf xg = |
|
651 |
fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), |
|
652 |
mk_Trueprop_eq (xf, xg))); |
|
653 |
||
654 |
val goal = |
|
655 |
Logic.list_implies (uv_eq :: map4 mk_prem xctrs xss xfs xgs, |
|
656 |
mk_Trueprop_eq (eta_ufcase, eta_vgcase)); |
|
657 |
val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); |
|
658 |
in |
|
659 |
(Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms), |
|
660 |
Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1))) |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
661 |
|> pairself (singleton (Proof_Context.export names_lthy lthy) #> |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
662 |
Thm.close_derivation) |
53917 | 663 |
end; |
664 |
||
665 |
val split_lhs = q $ ufcase; |
|
666 |
||
667 |
fun mk_split_conjunct xctr xs f_xs = |
|
668 |
list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); |
|
669 |
fun mk_split_disjunct xctr xs f_xs = |
|
670 |
list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), |
|
671 |
HOLogic.mk_not (q $ f_xs))); |
|
672 |
||
673 |
fun mk_split_goal xctrs xss xfs = |
|
674 |
mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj |
|
675 |
(map3 mk_split_conjunct xctrs xss xfs)); |
|
676 |
fun mk_split_asm_goal xctrs xss xfs = |
|
677 |
mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj |
|
678 |
(map3 mk_split_disjunct xctrs xss xfs))); |
|
679 |
||
680 |
fun prove_split selss goal = |
|
681 |
Goal.prove_sorry lthy [] [] goal (fn _ => |
|
682 |
mk_split_tac lthy uexhaust_thm case_thms selss inject_thmss distinct_thmsss) |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
683 |
|> singleton (Proof_Context.export names_lthy lthy) |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
684 |
|> Thm.close_derivation; |
53917 | 685 |
|
686 |
fun prove_split_asm asm_goal split_thm = |
|
687 |
Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} => |
|
688 |
mk_split_asm_tac ctxt split_thm) |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
689 |
|> singleton (Proof_Context.export names_lthy lthy) |
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
690 |
|> Thm.close_derivation; |
53917 | 691 |
|
692 |
val (split_thm, split_asm_thm) = |
|
693 |
let |
|
694 |
val goal = mk_split_goal xctrs xss xfs; |
|
695 |
val asm_goal = mk_split_asm_goal xctrs xss xfs; |
|
696 |
||
697 |
val thm = prove_split (replicate n []) goal; |
|
698 |
val asm_thm = prove_split_asm asm_goal thm; |
|
699 |
in |
|
700 |
(thm, asm_thm) |
|
701 |
end; |
|
702 |
||
56858 | 703 |
val (sel_defs, all_sel_thms, sel_thmss, disc_defs, disc_thmss, nontriv_disc_thmss, |
704 |
discI_thms, nontriv_discI_thms, disc_exclude_thms, disc_exclude_thmsss, |
|
705 |
disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms, safe_collapse_thms, |
|
706 |
expand_thms, sel_split_thms, sel_split_asm_thms, case_eq_if_thms) = |
|
52969
f2df0730f8ac
clarified option name (since case/fold/rec are also destructors)
blanchet
parents:
52968
diff
changeset
|
707 |
if no_discs_sels then |
56858 | 708 |
([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) |
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
709 |
else |
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
710 |
let |
53917 | 711 |
val udiscs = map (rapp u) discs; |
712 |
val uselss = map (map (rapp u)) selss; |
|
713 |
val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; |
|
714 |
val usel_fs = map2 (curry Term.list_comb) fs uselss; |
|
715 |
||
716 |
val vdiscs = map (rapp v) discs; |
|
717 |
val vselss = map (map (rapp v)) selss; |
|
718 |
||
49364
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
719 |
fun make_sel_thm xs' case_thm sel_def = |
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
720 |
zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs') |
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
721 |
(Drule.forall_intr_vars (case_thm RS (sel_def RS trans))))); |
49281 | 722 |
|
53704 | 723 |
val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss; |
724 |
||
49281 | 725 |
fun has_undefined_rhs thm = |
726 |
(case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of |
|
727 |
Const (@{const_name undefined}, _) => true |
|
728 |
| _ => false); |
|
729 |
||
730 |
val all_sel_thms = |
|
49285
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
731 |
(if all_sels_distinct andalso forall null sel_defaultss then |
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
732 |
flat sel_thmss |
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
733 |
else |
49364
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
734 |
map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs |
838b5e8ede73
allow default values to refer to selector arguments -- this is useful, e.g. for tllist: ttl (TNil x) = TNil x (example by Andreas Lochbihler)
blanchet
parents:
49336
diff
changeset
|
735 |
(xss' ~~ case_thms)) |
49285
036b833b99aa
removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents:
49281
diff
changeset
|
736 |
|> filter_out has_undefined_rhs; |
49278 | 737 |
|
738 |
fun mk_unique_disc_def () = |
|
739 |
let |
|
740 |
val m = the_single ms; |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
741 |
val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); |
49278 | 742 |
in |
51551 | 743 |
Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm) |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
744 |
|> singleton (Proof_Context.export names_lthy lthy) |
49692 | 745 |
|> Thm.close_derivation |
49278 | 746 |
end; |
747 |
||
748 |
fun mk_alternate_disc_def k = |
|
749 |
let |
|
750 |
val goal = |
|
49486 | 751 |
mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), |
752 |
nth exist_xs_u_eq_ctrs (k - 1)); |
|
49278 | 753 |
in |
51551 | 754 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
755 |
mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) |
49486 | 756 |
(nth distinct_thms (2 - k)) uexhaust_thm) |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
757 |
|> singleton (Proof_Context.export names_lthy lthy) |
49692 | 758 |
|> Thm.close_derivation |
49278 | 759 |
end; |
49028 | 760 |
|
49278 | 761 |
val has_alternate_disc_def = |
762 |
exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; |
|
763 |
||
764 |
val disc_defs' = |
|
765 |
map2 (fn k => fn def => |
|
766 |
if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () |
|
767 |
else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k |
|
54634
366297517091
reverted 141cb34744de and e78e7df36690 -- better provide nicer "eta-expanded" definitions for discriminators and selectors, since users might want to unfold them
blanchet
parents:
54626
diff
changeset
|
768 |
else def) ks disc_defs; |
49278 | 769 |
|
770 |
val discD_thms = map (fn def => def RS iffD1) disc_defs'; |
|
771 |
val discI_thms = |
|
772 |
map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms |
|
773 |
disc_defs'; |
|
774 |
val not_discI_thms = |
|
775 |
map2 (fn m => fn def => funpow m (fn thm => allI RS thm) |
|
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49500
diff
changeset
|
776 |
(unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) |
49278 | 777 |
ms disc_defs'; |
778 |
||
779 |
val (disc_thmss', disc_thmss) = |
|
780 |
let |
|
781 |
fun mk_thm discI _ [] = refl RS discI |
|
782 |
| mk_thm _ not_discI [distinct] = distinct RS not_discI; |
|
783 |
fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; |
|
784 |
in |
|
785 |
map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose |
|
786 |
end; |
|
787 |
||
55464 | 788 |
val nontriv_disc_thmss = |
789 |
map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; |
|
53704 | 790 |
|
791 |
fun is_discI_boring b = |
|
792 |
(n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); |
|
793 |
||
794 |
val nontriv_discI_thms = |
|
795 |
flat (map2 (fn b => if is_discI_boring b then K [] else single) disc_bindings |
|
796 |
discI_thms); |
|
49028 | 797 |
|
49486 | 798 |
val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) = |
799 |
let |
|
800 |
fun mk_goal [] = [] |
|
801 |
| mk_goal [((_, udisc), (_, udisc'))] = |
|
802 |
[Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, |
|
803 |
HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; |
|
804 |
||
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
805 |
fun prove tac goal = |
51551 | 806 |
Goal.prove_sorry lthy [] [] goal (K tac) |
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
807 |
|> Thm.close_derivation; |
49486 | 808 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
809 |
val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); |
49486 | 810 |
|
811 |
val half_goalss = map mk_goal half_pairss; |
|
812 |
val half_thmss = |
|
813 |
map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => |
|
51798 | 814 |
fn disc_thm => [prove (mk_half_disc_exclude_tac lthy m discD disc_thm) goal]) |
49486 | 815 |
half_goalss half_pairss (flat disc_thmss'); |
49278 | 816 |
|
49486 | 817 |
val other_half_goalss = map (mk_goal o map swap) half_pairss; |
818 |
val other_half_thmss = |
|
819 |
map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss |
|
820 |
other_half_goalss; |
|
821 |
in |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
822 |
join_halves n half_thmss other_half_thmss ||> `transpose |
49486 | 823 |
|>> has_alternate_disc_def ? K [] |
824 |
end; |
|
49278 | 825 |
|
49486 | 826 |
val disc_exhaust_thm = |
827 |
let |
|
828 |
fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; |
|
829 |
val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); |
|
830 |
in |
|
51551 | 831 |
Goal.prove_sorry lthy [] [] goal (fn _ => |
49486 | 832 |
mk_disc_exhaust_tac n exhaust_thm discI_thms) |
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
833 |
|> Thm.close_derivation |
49486 | 834 |
end; |
49028 | 835 |
|
53740 | 836 |
val (safe_collapse_thms, all_collapse_thms) = |
49486 | 837 |
let |
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
838 |
fun mk_goal m udisc usel_ctr = |
49486 | 839 |
let |
840 |
val prem = HOLogic.mk_Trueprop udisc; |
|
53916 | 841 |
val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); |
49486 | 842 |
in |
53740 | 843 |
(prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) |
49486 | 844 |
end; |
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
845 |
val (trivs, goals) = map3 mk_goal ms udiscs usel_ctrs |> split_list; |
53740 | 846 |
val thms = |
847 |
map5 (fn m => fn discD => fn sel_thms => fn triv => fn goal => |
|
848 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
|
849 |
mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL atac) |
|
850 |
|> Thm.close_derivation |
|
851 |
|> not triv ? perhaps (try (fn thm => refl RS thm))) |
|
852 |
ms discD_thms sel_thmss trivs goals; |
|
49486 | 853 |
in |
53740 | 854 |
(map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), |
855 |
thms) |
|
49486 | 856 |
end; |
49025 | 857 |
|
53916 | 858 |
val swapped_all_collapse_thms = |
859 |
map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; |
|
860 |
||
861 |
val sel_exhaust_thm = |
|
862 |
let |
|
863 |
fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; |
|
864 |
val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); |
|
865 |
in |
|
866 |
Goal.prove_sorry lthy [] [] goal (fn _ => |
|
867 |
mk_sel_exhaust_tac n disc_exhaust_thm swapped_all_collapse_thms) |
|
868 |
|> Thm.close_derivation |
|
869 |
end; |
|
870 |
||
53919 | 871 |
val expand_thm = |
49486 | 872 |
let |
873 |
fun mk_prems k udisc usels vdisc vsels = |
|
874 |
(if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ |
|
875 |
(if null usels then |
|
876 |
[] |
|
877 |
else |
|
878 |
[Logic.list_implies |
|
879 |
(if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], |
|
880 |
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj |
|
881 |
(map2 (curry HOLogic.mk_eq) usels vsels)))]); |
|
882 |
||
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
883 |
val goal = |
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
884 |
Library.foldr Logic.list_implies |
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
885 |
(map5 mk_prems ks udiscs uselss vdiscs vselss, uv_eq); |
49486 | 886 |
val uncollapse_thms = |
53740 | 887 |
map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; |
49486 | 888 |
in |
53919 | 889 |
Goal.prove_sorry lthy [] [] goal (fn _ => |
890 |
mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm) |
|
891 |
(inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss |
|
892 |
disc_exclude_thmsss') |
|
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
893 |
|> singleton (Proof_Context.export names_lthy lthy) |
53919 | 894 |
|> Thm.close_derivation |
49486 | 895 |
end; |
49278 | 896 |
|
53917 | 897 |
val (sel_split_thm, sel_split_asm_thm) = |
898 |
let |
|
899 |
val zss = map (K []) xss; |
|
900 |
val goal = mk_split_goal usel_ctrs zss usel_fs; |
|
901 |
val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; |
|
902 |
||
903 |
val thm = prove_split sel_thmss goal; |
|
904 |
val asm_thm = prove_split_asm asm_goal thm; |
|
905 |
in |
|
906 |
(thm, asm_thm) |
|
907 |
end; |
|
908 |
||
54491 | 909 |
val case_eq_if_thm = |
49486 | 910 |
let |
53917 | 911 |
val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); |
49486 | 912 |
in |
53919 | 913 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
54491 | 914 |
mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
915 |
|> singleton (Proof_Context.export names_lthy lthy) |
53919 | 916 |
|> Thm.close_derivation |
49486 | 917 |
end; |
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
918 |
in |
56858 | 919 |
(sel_defs, all_sel_thms, sel_thmss, disc_defs, disc_thmss, nontriv_disc_thmss, |
920 |
discI_thms, nontriv_discI_thms, disc_exclude_thms, disc_exclude_thmsss, |
|
921 |
[disc_exhaust_thm], [sel_exhaust_thm], all_collapse_thms, safe_collapse_thms, |
|
922 |
[expand_thm], [sel_split_thm], [sel_split_asm_thm], [case_eq_if_thm]) |
|
49116
3d520eec2746
allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents:
49114
diff
changeset
|
923 |
end; |
49025 | 924 |
|
49437 | 925 |
val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); |
53908 | 926 |
val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); |
49300 | 927 |
|
55464 | 928 |
val nontriv_disc_eq_thmss = |
929 |
map (map (fn th => th RS @{thm eq_False[THEN iffD2]} |
|
930 |
handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; |
|
931 |
||
54151 | 932 |
val anonymous_notes = |
933 |
[(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), |
|
55464 | 934 |
(flat nontriv_disc_eq_thmss, code_nitpicksimp_attrs)] |
54151 | 935 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
936 |
||
57038 | 937 |
(* "exhaust" is deliberately put first to avoid apparent circular dependencies in the proof |
938 |
objects, which would confuse MaSh. *) |
|
49052 | 939 |
val notes = |
57038 | 940 |
[(exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), |
941 |
(caseN, case_thms, code_nitpicksimp_simp_attrs), |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49591
diff
changeset
|
942 |
(case_congN, [case_cong_thm], []), |
54491 | 943 |
(case_eq_ifN, case_eq_if_thms, []), |
53740 | 944 |
(collapseN, safe_collapse_thms, simp_attrs), |
55464 | 945 |
(discN, flat nontriv_disc_thmss, simp_attrs), |
53700 | 946 |
(discIN, nontriv_discI_thms, []), |
53836 | 947 |
(disc_excludeN, disc_exclude_thms, dest_attrs), |
49300 | 948 |
(disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]), |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
949 |
(distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), |
49486 | 950 |
(expandN, expand_thms, []), |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
951 |
(injectN, inject_thms, iff_attrs @ inductsimp_attrs), |
49300 | 952 |
(nchotomyN, [nchotomy_thm], []), |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
953 |
(selN, all_sel_thms, code_nitpicksimp_simp_attrs), |
53916 | 954 |
(sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]), |
53917 | 955 |
(sel_splitN, sel_split_thms, []), |
956 |
(sel_split_asmN, sel_split_asm_thms, []), |
|
49300 | 957 |
(splitN, [split_thm], []), |
958 |
(split_asmN, [split_asm_thm], []), |
|
49633 | 959 |
(splitsN, [split_thm, split_asm_thm], []), |
49300 | 960 |
(weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)] |
961 |
|> filter_out (null o #2) |
|
962 |
|> map (fn (thmN, thms, attrs) => |
|
49633 | 963 |
((qualify true (Binding.name thmN), attrs), [(thms, [])])); |
49300 | 964 |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
965 |
val ctr_sugar = |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
966 |
{ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
967 |
nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
968 |
case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm, |
56858 | 969 |
split = split_thm, split_asm = split_asm_thm, disc_defs = disc_defs, |
970 |
disc_thmss = disc_thmss, discIs = discI_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, |
|
971 |
disc_excludesss = disc_exclude_thmsss, disc_exhausts = disc_exhaust_thms, |
|
972 |
sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms, |
|
973 |
sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms, |
|
974 |
case_eq_ifs = case_eq_if_thms}; |
|
49019 | 975 |
in |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
976 |
(ctr_sugar, |
51819 | 977 |
lthy |
55464 | 978 |
|> Spec_Rules.add Spec_Rules.Equational ([casex], case_thms) |
55471 | 979 |
|> fold (Spec_Rules.add Spec_Rules.Equational) |
55535
10194808430d
name derivations in 'primrec' for code extraction from proof terms
blanchet
parents:
55480
diff
changeset
|
980 |
(AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |
55471 | 981 |
|> fold (Spec_Rules.add Spec_Rules.Equational) |
982 |
(filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) |
|
55410 | 983 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
54493
diff
changeset
|
984 |
(fn phi => Case_Translation.register |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
54493
diff
changeset
|
985 |
(Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |
54691
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
986 |
|> Local_Theory.background_theory (fold (fold Code.del_eqn) [disc_defs, sel_defs]) |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
987 |
|> not no_code ? |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
988 |
Local_Theory.declaration {syntax = false, pervasive = false} |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
989 |
(fn phi => Context.mapping |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
990 |
(add_ctr_code fcT_name (map (Morphism.typ phi) As) |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
991 |
(map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
992 |
(Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54635
diff
changeset
|
993 |
I) |
54635
30666a281ae3
proper code generation for discriminators/selectors
blanchet
parents:
54634
diff
changeset
|
994 |
|> Local_Theory.notes (anonymous_notes @ notes) |> snd |
30666a281ae3
proper code generation for discriminators/selectors
blanchet
parents:
54634
diff
changeset
|
995 |
|> register_ctr_sugar fcT_name ctr_sugar) |
49019 | 996 |
end; |
49017 | 997 |
in |
49121 | 998 |
(goalss, after_qed, lthy') |
49017 | 999 |
end; |
1000 |
||
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1001 |
fun free_constructors tacss = (fn (goalss, after_qed, lthy) => |
51551 | 1002 |
map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1003 |
|> (fn thms => after_qed thms lthy)) oo prepare_free_constructors (K I); |
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
1004 |
|
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1005 |
val free_constructors_cmd = (fn (goalss, after_qed, lthy) => |
49297 | 1006 |
Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1007 |
prepare_free_constructors Syntax.read_term; |
49297 | 1008 |
|
57091 | 1009 |
val parse_bound_term = Parse.binding --| @{keyword ":"} -- Parse.term; |
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
1010 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1011 |
val parse_ctr_options = |
54626 | 1012 |
Scan.optional (@{keyword "("} |-- Parse.list1 |
55410 | 1013 |
(Parse.reserved "no_discs_sels" >> K 0 || Parse.reserved "no_code" >> K 1) --| |
1014 |
@{keyword ")"} |
|
1015 |
>> (fn js => (member (op =) js 0, member (op =) js 1))) |
|
1016 |
(false, false); |
|
49278 | 1017 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1018 |
val parse_defaults = |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1019 |
@{keyword "("} |-- Parse.reserved "defaults" |-- Scan.repeat parse_bound_term --| @{keyword ")"}; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1020 |
|
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1021 |
fun parse_ctr_spec parse_ctr parse_arg = |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1022 |
parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg -- |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1023 |
Scan.optional parse_defaults []; |
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1024 |
|
57091 | 1025 |
val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1026 |
|
49017 | 1027 |
val _ = |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1028 |
Outer_Syntax.local_theory_to_proof @{command_spec "free_constructors"} |
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1029 |
"register an existing freely generated type's constructors" |
57091 | 1030 |
(parse_ctr_options -- Parse.binding --| @{keyword "for"} -- parse_ctr_specs |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1031 |
>> free_constructors_cmd); |
49017 | 1032 |
|
56522
f54003750e7d
don't forget to init Interpretation and transfer theorems in the interpretation hook
kuncar
parents:
56376
diff
changeset
|
1033 |
val _ = Context.>> (Context.map_theory Ctr_Sugar_Interpretation.init); |
f54003750e7d
don't forget to init Interpretation and transfer theorems in the interpretation hook
kuncar
parents:
56376
diff
changeset
|
1034 |
|
49017 | 1035 |
end; |