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