author | nipkow |
Wed, 10 Jan 2018 15:25:09 +0100 | |
changeset 67399 | eab6ce8368fa |
parent 66251 | cd935b7cb3fb |
child 67405 | e9ab4ad7bd15 |
permissions | -rw-r--r-- |
54701 | 1 |
(* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML |
49017 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
59266 | 3 |
Author: Martin Desharnais, TU Muenchen |
54397 | 4 |
Copyright 2012, 2013 |
49017 | 5 |
|
51797 | 6 |
Wrapping existing freely generated type's constructors. |
49017 | 7 |
*) |
8 |
||
54006 | 9 |
signature CTR_SUGAR = |
49017 | 10 |
sig |
58675 | 11 |
datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown |
12 |
||
51840 | 13 |
type ctr_sugar = |
58675 | 14 |
{kind: ctr_sugar_kind, |
15 |
T: typ, |
|
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
16 |
ctrs: term list, |
52375 | 17 |
casex: term, |
51839 | 18 |
discs: term list, |
51819 | 19 |
selss: term list list, |
20 |
exhaust: thm, |
|
52375 | 21 |
nchotomy: thm, |
51819 | 22 |
injects: thm list, |
23 |
distincts: thm list, |
|
24 |
case_thms: thm list, |
|
52375 | 25 |
case_cong: thm, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
26 |
case_cong_weak: thm, |
59266 | 27 |
case_distribs: thm list, |
52375 | 28 |
split: thm, |
29 |
split_asm: thm, |
|
56858 | 30 |
disc_defs: thm list, |
51819 | 31 |
disc_thmss: thm list list, |
32 |
discIs: thm list, |
|
59271 | 33 |
disc_eq_cases: thm list, |
56858 | 34 |
sel_defs: thm list, |
53475 | 35 |
sel_thmss: thm list list, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
36 |
distinct_discsss: thm list list list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
37 |
exhaust_discs: thm list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
38 |
exhaust_sels: thm list, |
53475 | 39 |
collapses: thm list, |
40 |
expands: thm list, |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
41 |
split_sels: thm list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
42 |
split_sel_asms: thm list, |
54491 | 43 |
case_eq_ifs: thm list}; |
51809 | 44 |
|
51840 | 45 |
val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar |
58115 | 46 |
val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
47 |
val ctr_sugar_of: Proof.context -> string -> ctr_sugar option |
58116 | 48 |
val ctr_sugar_of_global: theory -> string -> ctr_sugar option |
53906 | 49 |
val ctr_sugars_of: Proof.context -> ctr_sugar list |
58116 | 50 |
val ctr_sugars_of_global: theory -> ctr_sugar list |
54403 | 51 |
val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option |
58116 | 52 |
val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option |
61760 | 53 |
val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> |
54 |
theory |
|
58188 | 55 |
val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory |
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
56 |
val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory |
58188 | 57 |
val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory |
58 |
val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory |
|
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
59 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
60 |
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
|
61 |
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
|
62 |
|
49203 | 63 |
val mk_ctr: typ list -> term -> term |
53864
a48d4bd3faaa
use case rather than sequence of ifs in expansion
blanchet
parents:
53857
diff
changeset
|
64 |
val mk_case: typ list -> typ -> term -> term |
49484 | 65 |
val mk_disc_or_sel: typ list -> term -> term |
49622 | 66 |
val name_of_ctr: term -> string |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
67 |
val name_of_disc: term -> string |
53888 | 68 |
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
|
69 |
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
|
70 |
(ctr_sugar * term list * term list) option |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
71 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
72 |
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
|
73 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
74 |
val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
75 |
val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
76 |
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
|
77 |
|
59283 | 78 |
val code_plugin: string |
79 |
||
58191 | 80 |
type ctr_options = (string -> bool) * bool |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
81 |
type ctr_options_cmd = (Proof.context -> string -> bool) * bool |
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
82 |
|
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
83 |
val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context |
58675 | 84 |
val free_constructors: ctr_sugar_kind -> |
85 |
({prems: thm list, context: Proof.context} -> tactic) list list -> |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
86 |
((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> |
51840 | 87 |
ctr_sugar * local_theory |
61301 | 88 |
val free_constructors_cmd: ctr_sugar_kind -> |
89 |
((((Proof.context -> Plugin_Name.filter) * bool) * binding) |
|
90 |
* ((binding * string) * binding list) list) * string list -> |
|
91 |
Proof.context -> Proof.state |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
92 |
val default_ctr_options: ctr_options |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
93 |
val default_ctr_options_cmd: ctr_options_cmd |
49286 | 94 |
val parse_bound_term: (binding * string) parser |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
95 |
val parse_ctr_options: ctr_options_cmd parser |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
96 |
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
|
97 |
val parse_sel_default_eqs: string list parser |
49017 | 98 |
end; |
99 |
||
54006 | 100 |
structure Ctr_Sugar : CTR_SUGAR = |
49017 | 101 |
struct |
102 |
||
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
103 |
open Ctr_Sugar_Util |
54006 | 104 |
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
|
105 |
open Ctr_Sugar_Code |
49017 | 106 |
|
58675 | 107 |
datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; |
108 |
||
51840 | 109 |
type ctr_sugar = |
58675 | 110 |
{kind: ctr_sugar_kind, |
111 |
T: typ, |
|
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
112 |
ctrs: term list, |
52375 | 113 |
casex: term, |
51839 | 114 |
discs: term list, |
51819 | 115 |
selss: term list list, |
116 |
exhaust: thm, |
|
52375 | 117 |
nchotomy: thm, |
51819 | 118 |
injects: thm list, |
119 |
distincts: thm list, |
|
120 |
case_thms: thm list, |
|
52375 | 121 |
case_cong: thm, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
122 |
case_cong_weak: thm, |
59266 | 123 |
case_distribs: thm list, |
52375 | 124 |
split: thm, |
125 |
split_asm: thm, |
|
56858 | 126 |
disc_defs: thm list, |
51819 | 127 |
disc_thmss: thm list list, |
128 |
discIs: thm list, |
|
59271 | 129 |
disc_eq_cases: thm list, |
56858 | 130 |
sel_defs: thm list, |
53475 | 131 |
sel_thmss: thm list list, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
132 |
distinct_discsss: thm list list list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
133 |
exhaust_discs: thm list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
134 |
exhaust_sels: thm list, |
53475 | 135 |
collapses: thm list, |
136 |
expands: thm list, |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
137 |
split_sels: thm list, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
138 |
split_sel_asms: thm list, |
54491 | 139 |
case_eq_ifs: thm list}; |
51809 | 140 |
|
58685 | 141 |
fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, |
59266 | 142 |
case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, |
59271 | 143 |
discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, |
144 |
collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = |
|
58675 | 145 |
{kind = kind, |
146 |
T = Morphism.typ phi T, |
|
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
147 |
ctrs = map (Morphism.term phi) ctrs, |
52375 | 148 |
casex = Morphism.term phi casex, |
51839 | 149 |
discs = map (Morphism.term phi) discs, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
150 |
selss = map (map (Morphism.term phi)) selss, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
151 |
exhaust = Morphism.thm phi exhaust, |
52375 | 152 |
nchotomy = Morphism.thm phi nchotomy, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
153 |
injects = map (Morphism.thm phi) injects, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
154 |
distincts = map (Morphism.thm phi) distincts, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
155 |
case_thms = map (Morphism.thm phi) case_thms, |
52375 | 156 |
case_cong = Morphism.thm phi case_cong, |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
157 |
case_cong_weak = Morphism.thm phi case_cong_weak, |
59266 | 158 |
case_distribs = map (Morphism.thm phi) case_distribs, |
52375 | 159 |
split = Morphism.thm phi split, |
160 |
split_asm = Morphism.thm phi split_asm, |
|
56858 | 161 |
disc_defs = map (Morphism.thm phi) disc_defs, |
51823
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
162 |
disc_thmss = map (map (Morphism.thm phi)) disc_thmss, |
38996458bc5c
create data structure for storing (co)datatype information
blanchet
parents:
51819
diff
changeset
|
163 |
discIs = map (Morphism.thm phi) discIs, |
59271 | 164 |
disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, |
56858 | 165 |
sel_defs = map (Morphism.thm phi) sel_defs, |
53475 | 166 |
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
|
167 |
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
|
168 |
exhaust_discs = map (Morphism.thm phi) exhaust_discs, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
169 |
exhaust_sels = map (Morphism.thm phi) exhaust_sels, |
53475 | 170 |
collapses = map (Morphism.thm phi) collapses, |
171 |
expands = map (Morphism.thm phi) expands, |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
172 |
split_sels = map (Morphism.thm phi) split_sels, |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
173 |
split_sel_asms = map (Morphism.thm phi) split_sel_asms, |
54491 | 174 |
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
|
175 |
|
58115 | 176 |
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
|
177 |
|
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
178 |
structure Data = Generic_Data |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
179 |
( |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
180 |
type T = ctr_sugar Symtab.table; |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
181 |
val empty = Symtab.empty; |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
182 |
val extend = I; |
55394
d5ebe055b599
more liberal merging of BNFs and constructor sugar
blanchet
parents:
55356
diff
changeset
|
183 |
fun merge data : T = Symtab.merge (K true) data; |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
184 |
); |
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
185 |
|
58116 | 186 |
fun ctr_sugar_of_generic context = |
187 |
Option.map (transfer_ctr_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); |
|
188 |
||
189 |
fun ctr_sugars_of_generic context = |
|
190 |
Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o snd) (Data.get context) []; |
|
53906 | 191 |
|
58116 | 192 |
fun ctr_sugar_of_case_generic context s = |
193 |
find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) |
|
194 |
(ctr_sugars_of_generic context); |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
195 |
|
58116 | 196 |
val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; |
197 |
val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; |
|
198 |
||
199 |
val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; |
|
200 |
val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; |
|
201 |
||
202 |
val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; |
|
203 |
val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; |
|
54403 | 204 |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
205 |
structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); |
56345 | 206 |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
207 |
fun ctr_sugar_interpretation name f = |
62322 | 208 |
Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => |
209 |
f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); |
|
56376
5a93b8f928a2
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
blanchet
parents:
56345
diff
changeset
|
210 |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
211 |
val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; |
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
212 |
|
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
213 |
fun register_ctr_sugar_raw (ctr_sugar as {T = Type (s, _), ...}) = |
54285
578371ba74cc
reverted 3e1d230f1c00 -- pervasiveness is useful, cf. Coinductive_Nat in the AFP
blanchet
parents:
54265
diff
changeset
|
214 |
Local_Theory.declaration {syntax = false, pervasive = true} |
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
215 |
(fn phi => Data.map (Symtab.update (s, morph_ctr_sugar phi ctr_sugar))); |
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
216 |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
217 |
fun register_ctr_sugar plugins ctr_sugar = |
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
218 |
register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; |
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
219 |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
220 |
fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (s, _), ...}) thy = |
56345 | 221 |
let val tab = Data.get (Context.Theory thy) in |
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
222 |
if Symtab.defined tab s then |
56345 | 223 |
thy |
224 |
else |
|
225 |
thy |
|
58187
d2ddd401d74d
fixed infinite loops in 'register' functions + more uniform API
blanchet
parents:
58186
diff
changeset
|
226 |
|> Context.theory_map (Data.put (Symtab.update_new (s, ctr_sugar) tab)) |
58665 | 227 |
|> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) |
56345 | 228 |
end; |
54400 | 229 |
|
62322 | 230 |
val is_prefix = "is_"; |
231 |
val un_prefix = "un_"; |
|
232 |
val not_prefix = "not_"; |
|
57629
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
blanchet
parents:
57260
diff
changeset
|
233 |
|
62322 | 234 |
fun mk_unN 1 1 suf = un_prefix ^ suf |
235 |
| mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; |
|
49046
3c5eba97d93a
allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents:
49045
diff
changeset
|
236 |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49591
diff
changeset
|
237 |
val caseN = "case"; |
49054 | 238 |
val case_congN = "case_cong"; |
54491 | 239 |
val case_eq_ifN = "case_eq_if"; |
49118 | 240 |
val collapseN = "collapse"; |
53694 | 241 |
val discN = "disc"; |
59272 | 242 |
val disc_eq_caseN = "disc_eq_case"; |
53700 | 243 |
val discIN = "discI"; |
49054 | 244 |
val distinctN = "distinct"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
245 |
val distinct_discN = "distinct_disc"; |
49075 | 246 |
val exhaustN = "exhaust"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
247 |
val exhaust_discN = "exhaust_disc"; |
49486 | 248 |
val expandN = "expand"; |
49075 | 249 |
val injectN = "inject"; |
250 |
val nchotomyN = "nchotomy"; |
|
53694 | 251 |
val selN = "sel"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
252 |
val exhaust_selN = "exhaust_sel"; |
49054 | 253 |
val splitN = "split"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
254 |
val split_asmN = "split_asm"; |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
255 |
val split_selN = "split_sel"; |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
256 |
val split_sel_asmN = "split_sel_asm"; |
49633 | 257 |
val splitsN = "splits"; |
57984
cbe9a16f8e11
added collection theorem for consistency and convenience
blanchet
parents:
57983
diff
changeset
|
258 |
val split_selsN = "split_sels"; |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
259 |
val case_cong_weak_thmsN = "case_cong_weak"; |
59266 | 260 |
val case_distribN = "case_distrib"; |
49019 | 261 |
|
53700 | 262 |
val cong_attrs = @{attributes [cong]}; |
53836 | 263 |
val dest_attrs = @{attributes [dest]}; |
53700 | 264 |
val safe_elim_attrs = @{attributes [elim!]}; |
265 |
val iff_attrs = @{attributes [iff]}; |
|
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
266 |
val inductsimp_attrs = @{attributes [induct_simp]}; |
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
267 |
val nitpicksimp_attrs = @{attributes [nitpick_simp]}; |
49300 | 268 |
val simp_attrs = @{attributes [simp]}; |
49046
3c5eba97d93a
allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents:
49045
diff
changeset
|
269 |
|
55480
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
blanchet
parents:
55471
diff
changeset
|
270 |
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
|
271 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
272 |
fun mk_half_pairss' _ ([], []) = [] |
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
273 |
| mk_half_pairss' indent (x :: xs, _ :: ys) = |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
274 |
indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); |
49486 | 275 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
276 |
fun mk_half_pairss p = mk_half_pairss' [[]] p; |
49027 | 277 |
|
49486 | 278 |
fun join_halves n half_xss other_half_xss = |
55342 | 279 |
(splice (flat half_xss) (flat other_half_xss), |
280 |
map2 (map2 append) (Library.chop_groups n half_xss) |
|
281 |
(transpose (Library.chop_groups n other_half_xss))); |
|
49027 | 282 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
283 |
fun mk_undefined T = Const (@{const_name undefined}, T); |
49055 | 284 |
|
49500 | 285 |
fun mk_ctr Ts t = |
286 |
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
|
287 |
subst_nonatomic_types (Ts0 ~~ Ts) t |
49203 | 288 |
end; |
289 |
||
49536 | 290 |
fun mk_case Ts T t = |
291 |
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
|
292 |
subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t |
49536 | 293 |
end; |
294 |
||
53864
a48d4bd3faaa
use case rather than sequence of ifs in expansion
blanchet
parents:
53857
diff
changeset
|
295 |
fun mk_disc_or_sel Ts t = |
54435
4a655e62ad34
fixed type variable confusion in 'datatype_new_compat'
blanchet
parents:
54422
diff
changeset
|
296 |
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
|
297 |
|
58284
f9b6af3017fd
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
blanchet
parents:
58256
diff
changeset
|
298 |
val name_of_ctr = name_of_const "constructor" body_type; |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
299 |
|
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
300 |
fun name_of_disc t = |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
301 |
(case head_of t of |
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
302 |
Abs (_, _, @{const Not} $ (t' $ Bound 0)) => |
62322 | 303 |
Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
304 |
| Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') => |
62322 | 305 |
Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') |
51777
48a0ae342ea0
generate proper attributes for coinduction rules
blanchet
parents:
51771
diff
changeset
|
306 |
| Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) => |
62322 | 307 |
Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') |
58289 | 308 |
| t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); |
49622 | 309 |
|
310 |
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
|
311 |
|
53888 | 312 |
fun dest_ctr ctxt s t = |
55342 | 313 |
let val (f, args) = Term.strip_comb t in |
53888 | 314 |
(case ctr_sugar_of ctxt s of |
315 |
SOME {ctrs, ...} => |
|
316 |
(case find_first (can (fo_match ctxt f)) ctrs of |
|
317 |
SOME f' => (f', args) |
|
318 |
| NONE => raise Fail "dest_ctr") |
|
319 |
| NONE => raise Fail "dest_ctr") |
|
320 |
end; |
|
321 |
||
53870 | 322 |
fun dest_case ctxt s Ts t = |
323 |
(case Term.strip_comb t of |
|
324 |
(Const (c, _), args as _ :: _) => |
|
325 |
(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
|
326 |
SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => |
53870 | 327 |
if case_name = c then |
53924 | 328 |
let val n = length discs0 in |
329 |
if n < length args then |
|
330 |
let |
|
331 |
val (branches, obj :: leftovers) = chop n args; |
|
332 |
val discs = map (mk_disc_or_sel Ts) discs0; |
|
333 |
val selss = map (map (mk_disc_or_sel Ts)) selss0; |
|
334 |
val conds = map (rapp obj) discs; |
|
335 |
val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; |
|
336 |
val branches' = map2 (curry Term.betapplys) branches branch_argss; |
|
337 |
in |
|
54970
891141de5672
only destruct cases equipped with the right stuff (in particular, 'sel_split')
blanchet
parents:
54900
diff
changeset
|
338 |
SOME (ctr_sugar, conds, branches') |
53924 | 339 |
end |
340 |
else |
|
341 |
NONE |
|
53870 | 342 |
end |
343 |
else |
|
344 |
NONE |
|
345 |
| _ => NONE) |
|
346 |
| _ => NONE); |
|
347 |
||
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
348 |
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
|
349 |
| const_or_free_name (Free (s, _)) = s |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
350 |
| const_or_free_name t = raise TERM ("const_or_free_name", [t]) |
49437 | 351 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
352 |
fun extract_sel_default ctxt t = |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
353 |
let |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
354 |
fun malformed () = |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
355 |
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
|
356 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
357 |
val ((sel, (ctr, vars)), rhs) = |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
358 |
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
|
359 |
|> HOLogic.dest_eq |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
360 |
|>> (Term.dest_comb |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
361 |
#>> const_or_free_name |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
362 |
##> (Term.strip_comb #>> (Term.dest_Const #> fst))) |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
363 |
handle TERM _ => malformed (); |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
364 |
in |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
365 |
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
|
366 |
((ctr, sel), fold_rev Term.lambda vars rhs) |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
367 |
else |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
368 |
malformed () |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
369 |
end; |
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
370 |
|
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
371 |
(* 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
|
372 |
fun fake_local_theory_for_sel_defaults sel_bTs = |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
373 |
Proof_Context.allow_dummies |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
374 |
#> 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
|
375 |
#> snd; |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
376 |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
377 |
type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
378 |
|
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
379 |
fun disc_of_ctr_spec ((disc, _), _) = disc; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
380 |
fun ctr_of_ctr_spec ((_, ctr), _) = ctr; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
381 |
fun args_of_ctr_spec (_, args) = args; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
382 |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
383 |
val code_plugin = Plugin_Name.declare_setup @{binding code}; |
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
384 |
|
58675 | 385 |
fun prepare_free_constructors kind prep_plugins prep_term |
59281 | 386 |
((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = |
49017 | 387 |
let |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
388 |
val plugins = prep_plugins no_defs_lthy raw_plugins; |
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
389 |
|
49019 | 390 |
(* TODO: sanity checks on arguments *) |
49025 | 391 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
392 |
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
|
393 |
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
|
394 |
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
|
395 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
396 |
val n = length raw_ctrs; |
49054 | 397 |
val ks = 1 upto n; |
398 |
||
59281 | 399 |
val _ = n > 0 orelse error "No constructors specified"; |
49121 | 400 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
401 |
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
|
402 |
|
60279
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
403 |
val (fcT_name, As0) = |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
404 |
(case body_type (fastype_of (hd ctrs0)) of |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
405 |
Type T' => T' |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
406 |
| _ => error "Expected type constructor in body type of constructor"); |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
407 |
val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
408 |
o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; |
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
409 |
|
53908 | 410 |
val fc_b_name = Long_Name.base_name fcT_name; |
411 |
val fc_b = Binding.name fc_b_name; |
|
49055 | 412 |
|
55410 | 413 |
fun qualify mandatory = Binding.qualify mandatory fc_b_name; |
49633 | 414 |
|
59266 | 415 |
val (unsorted_As, [B, C]) = |
49055 | 416 |
no_defs_lthy |
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
417 |
|> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) |
59266 | 418 |
||> fst o mk_TFrees 2; |
49055 | 419 |
|
58234 | 420 |
val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; |
53268
de1dc709f9d4
handle type class annotations on (co)datatype parameters gracefully
blanchet
parents:
53210
diff
changeset
|
421 |
|
53908 | 422 |
val fcT = Type (fcT_name, As); |
49055 | 423 |
val ctrs = map (mk_ctr As) ctrs0; |
424 |
val ctr_Tss = map (binder_types o fastype_of) ctrs; |
|
425 |
||
426 |
val ms = map length ctr_Tss; |
|
427 |
||
57091 | 428 |
fun can_definitely_rely_on_disc k = |
429 |
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
|
430 |
fun can_rely_on_disc k = |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
431 |
can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); |
53925 | 432 |
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
|
433 |
|
57090
0224caba67ca
use '%x. x = C' as default discriminator for nullary constructor C, instead of relying on odd '=:' syntax
blanchet
parents:
57038
diff
changeset
|
434 |
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
|
435 |
|
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
436 |
fun is_disc_binding_valid b = |
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
437 |
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
|
438 |
|
62322 | 439 |
val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; |
49120
7f8e69fc6ac9
smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents:
49119
diff
changeset
|
440 |
|
49336 | 441 |
val disc_bindings = |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
442 |
raw_disc_bindings |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
443 |
|> @{map 4} (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
|
444 |
qualify false |
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
445 |
(if Binding.is_empty disc then |
57091 | 446 |
if m = 0 then equal_binding |
447 |
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
|
448 |
else standard_disc_binding ctr |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
449 |
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
|
450 |
standard_disc_binding ctr |
49302
f5bd87aac224
added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents:
49300
diff
changeset
|
451 |
else |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
452 |
disc)) ks ms ctrs0; |
49056 | 453 |
|
51787
1267c28c7bdd
changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents:
51781
diff
changeset
|
454 |
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
|
455 |
|
49336 | 456 |
val sel_bindingss = |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
457 |
@{map 3} (fn ctr => fn m => map2 (fn l => fn sel => |
49633 | 458 |
qualify false |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
459 |
(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
|
460 |
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
|
461 |
else |
55470
46e6e1d91056
removed needless robustness (no longer needed thanks to new syntax)
blanchet
parents:
55469
diff
changeset
|
462 |
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
|
463 |
|
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
464 |
val add_bindings = |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
465 |
Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
466 |
(map Binding.name_of (disc_bindings @ flat sel_bindingss)))) |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
467 |
#> snd; |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
468 |
|
49201 | 469 |
val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; |
49043 | 470 |
|
62320 | 471 |
val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
472 |
|> add_bindings |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
473 |
|> yield_singleton (mk_Frees fc_b_name) fcT |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
474 |
||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) |
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
475 |
||>> mk_Freess "x" ctr_Tss |
49025 | 476 |
||>> mk_Freess "y" ctr_Tss |
49201 | 477 |
||>> mk_Frees "f" case_Ts |
478 |
||>> mk_Frees "g" case_Ts |
|
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
479 |
||>> yield_singleton (mk_Frees "z") B |
67399 | 480 |
||>> yield_singleton (apfst (~~) oo mk_Frees' "P") HOLogic.boolT; |
49043 | 481 |
|
49463 | 482 |
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
|
483 |
|
49025 | 484 |
val xctrs = map2 (curry Term.list_comb) ctrs xss; |
485 |
val yctrs = map2 (curry Term.list_comb) ctrs yss; |
|
49032 | 486 |
|
49043 | 487 |
val xfs = map2 (curry Term.list_comb) fs xss; |
488 |
val xgs = map2 (curry Term.list_comb) gs xss; |
|
489 |
||
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
490 |
(* 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
|
491 |
nicer names). Consider removing. *) |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
492 |
val eta_fs = map2 (fold_rev Term.lambda) xss xfs; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
493 |
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
|
494 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
495 |
val case_binding = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
496 |
qualify false |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
497 |
(if Binding.is_empty raw_case_binding orelse |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
498 |
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
|
499 |
Binding.prefix_name (caseN ^ "_") fc_b |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
500 |
else |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
501 |
raw_case_binding); |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
502 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
503 |
fun mk_case_disj xctr xf xs = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
504 |
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
|
505 |
|
53925 | 506 |
val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] |
507 |
(Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $ |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
508 |
Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
509 |
|
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
510 |
val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
511 |
|> Local_Theory.open_target |> snd |
54155 | 512 |
|> Local_Theory.define ((case_binding, NoSyn), |
59859 | 513 |
((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) |
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
514 |
||> `Local_Theory.close_target; |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
515 |
|
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
516 |
val phi = Proof_Context.export_morphism lthy_old lthy; |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
517 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
518 |
val case_def = Morphism.thm phi raw_case_def; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
519 |
|
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
520 |
val case0 = Morphism.term phi raw_case; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
521 |
val casex = mk_case As B case0; |
59266 | 522 |
val casexC = mk_case As C case0; |
59271 | 523 |
val casexBool = mk_case As HOLogic.boolT case0; |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
524 |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
525 |
fun mk_uu_eq () = HOLogic.mk_eq (u, u); |
49486 | 526 |
|
527 |
val exist_xs_u_eq_ctrs = |
|
528 |
map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; |
|
49022 | 529 |
|
51743 | 530 |
val unique_disc_no_def = TrueI; (*arbitrary marker*) |
531 |
val alternate_disc_no_def = FalseE; (*arbitrary marker*) |
|
532 |
||
49486 | 533 |
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
|
534 |
HOLogic.mk_not |
51790
22517d04d20b
more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents:
51787
diff
changeset
|
535 |
(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
|
536 |
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
|
537 |
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
|
538 |
|
57094
589ec121ce1a
don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents:
57091
diff
changeset
|
539 |
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
|
540 |
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
|
541 |
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
|
542 |
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
|
543 |
|
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
544 |
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
|
545 |
if no_discs_sels then |
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
546 |
(true, [], [], [], [], [], lthy) |
49278 | 547 |
else |
548 |
let |
|
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
549 |
val all_sel_bindings = flat sel_bindingss; |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
550 |
val num_all_sel_bindings = length all_sel_bindings; |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
551 |
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
|
552 |
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
|
553 |
|
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
554 |
val sel_binding_index = |
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
555 |
if all_sels_distinct then |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
556 |
1 upto num_all_sel_bindings |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
557 |
else |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
558 |
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
|
559 |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
560 |
val all_proto_sels = flat (@{map 3} (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
|
561 |
val sel_infos = |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
562 |
AList.group (op =) (sel_binding_index ~~ all_proto_sels) |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58998
diff
changeset
|
563 |
|> sort (int_ord o apply2 fst) |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
564 |
|> map snd |> curry (op ~~) uniq_sel_bindings; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
565 |
val sel_bindings = map fst sel_infos; |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
566 |
|
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
567 |
val sel_defaults = |
57830
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
568 |
if null sel_default_eqs then |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
569 |
[] |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
570 |
else |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
571 |
let |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
572 |
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
|
573 |
val fake_lthy = |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
574 |
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
|
575 |
in |
2d0f0d6fdf3e
correctly resolve selector names in default value equations
blanchet
parents:
57633
diff
changeset
|
576 |
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
|
577 |
end; |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
578 |
|
53908 | 579 |
fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); |
49025 | 580 |
|
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
|
581 |
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
|
582 |
|
49500 | 583 |
fun alternate_disc k = |
584 |
Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); |
|
49278 | 585 |
|
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
586 |
fun mk_sel_case_args b proto_sels T = |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
587 |
@{map 3} (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
|
588 |
(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
|
589 |
NONE => |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
590 |
(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
|
591 |
[] => 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
|
592 |
| [(_, t)] => t |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
593 |
| _ => error "Multiple default values for selector/constructor pair") |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
594 |
| 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
|
595 |
|
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
|
596 |
fun sel_spec b proto_sels = |
49278 | 597 |
let |
598 |
val _ = |
|
599 |
(case duplicates (op =) (map fst proto_sels) of |
|
600 |
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
|
601 |
" for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) |
49278 | 602 |
| [] => ()) |
603 |
val T = |
|
604 |
(case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of |
|
605 |
[T] => T |
|
606 |
| T :: T' :: _ => error ("Inconsistent range type for selector " ^ |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
607 |
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
|
608 |
" vs. " ^ quote (Syntax.string_of_typ lthy T'))); |
49278 | 609 |
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
|
610 |
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
|
611 |
Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) |
49278 | 612 |
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
|
613 |
|
49336 | 614 |
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
|
615 |
|
49278 | 616 |
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
|
617 |
lthy |
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
618 |
|> Local_Theory.open_target |> snd |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
619 |
|> apfst split_list o @{fold_map 3} (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
|
620 |
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
|
621 |
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
|
622 |
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
|
623 |
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
|
624 |
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
|
625 |
else |
63180 | 626 |
Specification.definition (SOME (b, NONE, NoSyn)) [] [] |
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
62326
diff
changeset
|
627 |
((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) |
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
|
628 |
ks exist_xs_u_eq_ctrs disc_bindings |
49278 | 629 |
||>> apfst split_list o fold_map (fn (b, proto_sels) => |
63180 | 630 |
Specification.definition (SOME (b, NONE, NoSyn)) [] [] |
63064
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
wenzelm
parents:
62326
diff
changeset
|
631 |
((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos |
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
632 |
||> `Local_Theory.close_target; |
49022 | 633 |
|
49278 | 634 |
val phi = Proof_Context.export_morphism lthy lthy'; |
49025 | 635 |
|
49278 | 636 |
val disc_defs = map (Morphism.thm phi) raw_disc_defs; |
49281 | 637 |
val sel_defs = map (Morphism.thm phi) raw_sel_defs; |
638 |
val sel_defss = unflat_selss sel_defs; |
|
49278 | 639 |
|
640 |
val discs0 = map (Morphism.term phi) raw_discs; |
|
641 |
val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); |
|
49028 | 642 |
|
49278 | 643 |
val discs = map (mk_disc_or_sel As) discs0; |
644 |
val selss = map (map (mk_disc_or_sel As)) selss0; |
|
645 |
in |
|
53917 | 646 |
(all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') |
49278 | 647 |
end; |
49025 | 648 |
|
49032 | 649 |
fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); |
49029 | 650 |
|
49458 | 651 |
val exhaust_goal = |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
652 |
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
|
653 |
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
|
654 |
end; |
49019 | 655 |
|
49484 | 656 |
val inject_goalss = |
49017 | 657 |
let |
49034
b77e1910af8a
make parallel list indexing possible for inject theorems
blanchet
parents:
49033
diff
changeset
|
658 |
fun mk_goal _ _ [] [] = [] |
49025 | 659 |
| mk_goal xctr yctr xs ys = |
49121 | 660 |
[fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), |
661 |
Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; |
|
49017 | 662 |
in |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
663 |
@{map 4} mk_goal xctrs yctrs xss yss |
49017 | 664 |
end; |
665 |
||
49484 | 666 |
val half_distinct_goalss = |
49121 | 667 |
let |
49203 | 668 |
fun mk_goal ((xs, xc), (xs', xc')) = |
49121 | 669 |
fold_rev Logic.all (xs @ xs') |
49203 | 670 |
(HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); |
49121 | 671 |
in |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
672 |
map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) |
49121 | 673 |
end; |
49019 | 674 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
675 |
val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; |
49019 | 676 |
|
60279
351536745704
took out unreliable 'blast' from tactic altogether
blanchet
parents:
60251
diff
changeset
|
677 |
fun after_qed ([exhaust_thm] :: thmss) lthy = |
49019 | 678 |
let |
62320 | 679 |
val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
680 |
|> add_bindings |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
681 |
|> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT |
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
682 |
||>> mk_Freess' "x" ctr_Tss |
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
683 |
||>> mk_Frees "f" case_Ts |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
684 |
||>> mk_Frees "g" case_Ts |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
685 |
||>> yield_singleton (mk_Frees "h") (B --> C) |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
686 |
||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
687 |
||>> yield_singleton (mk_Frees "P") HOLogic.boolT; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
688 |
|
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
689 |
val xfs = map2 (curry Term.list_comb) fs xss; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
690 |
val xgs = map2 (curry Term.list_comb) gs xss; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
691 |
|
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
692 |
val fcase = Term.list_comb (casex, fs); |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
693 |
|
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
694 |
val ufcase = fcase $ u; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
695 |
val vfcase = fcase $ v; |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
696 |
|
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
697 |
val eta_fcase = Term.list_comb (casex, eta_fs); |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
698 |
val eta_gcase = Term.list_comb (casex, eta_gs); |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
699 |
|
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
700 |
val eta_ufcase = eta_fcase $ u; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
701 |
val eta_vgcase = eta_gcase $ v; |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
702 |
|
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
703 |
fun mk_uu_eq () = HOLogic.mk_eq (u, u); |
61550
0b39a1f26604
allow selectors and discriminators with same name as type
blanchet
parents:
61334
diff
changeset
|
704 |
|
61272
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
705 |
val uv_eq = mk_Trueprop_eq (u, v); |
f49644098959
restructure fresh variable generation to make exports more wellformed
traytel
parents:
61271
diff
changeset
|
706 |
|
57633 | 707 |
val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; |
49438 | 708 |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60279
diff
changeset
|
709 |
val rho_As = |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60279
diff
changeset
|
710 |
map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60279
diff
changeset
|
711 |
(map Logic.varifyT_global As ~~ As); |
49486 | 712 |
|
713 |
fun inst_thm t thm = |
|
60801 | 714 |
Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] |
53210
7219c61796c0
simplify code (now that ctr_sugar uses the same type variables as fp_sugar)
blanchet
parents:
53040
diff
changeset
|
715 |
(Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); |
49486 | 716 |
|
717 |
val uexhaust_thm = inst_thm u exhaust_thm; |
|
49032 | 718 |
|
49300 | 719 |
val exhaust_cases = map base_name_of_ctr ctrs; |
720 |
||
49048
4e0f0f98be02
rationalized data structure for distinctness theorems
blanchet
parents:
49046
diff
changeset
|
721 |
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
|
722 |
|
49486 | 723 |
val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
724 |
join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; |
49019 | 725 |
|
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
726 |
val nchotomy_thm = |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
727 |
let |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
728 |
val goal = |
49486 | 729 |
HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', |
730 |
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
|
731 |
in |
60728 | 732 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
733 |
mk_nchotomy_tac ctxt n exhaust_thm) |
|
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
734 |
|> Thm.close_derivation |
49020
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
735 |
end; |
f379cf5d71bd
more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents:
49019
diff
changeset
|
736 |
|
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
737 |
val case_thms = |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
738 |
let |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
739 |
val goals = |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
740 |
@{map 3} (fn xctr => fn xf => fn xs => |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
741 |
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
|
742 |
in |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
743 |
@{map 4} (fn k => fn goal => fn injects => fn distinctss => |
52968
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
744 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
745 |
mk_case_tac ctxt n k case_def injects distinctss) |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
746 |
|> Thm.close_derivation) |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
747 |
ks goals inject_thmss distinct_thmsss |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
748 |
end; |
2b430bbb5a1a
define case constant from other 'free constructor' axioms
blanchet
parents:
52965
diff
changeset
|
749 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
750 |
val (case_cong_thm, case_cong_weak_thm) = |
53917 | 751 |
let |
752 |
fun mk_prem xctr xs xf xg = |
|
753 |
fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), |
|
754 |
mk_Trueprop_eq (xf, xg))); |
|
755 |
||
756 |
val goal = |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
757 |
Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, |
53917 | 758 |
mk_Trueprop_eq (eta_ufcase, eta_vgcase)); |
759 |
val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
760 |
val vars = Variable.add_free_names lthy goal []; |
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
761 |
val weak_vars = Variable.add_free_names lthy weak_goal []; |
53917 | 762 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
763 |
(Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => |
61271 | 764 |
mk_case_cong_tac ctxt uexhaust_thm case_thms), |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
765 |
Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => |
60728 | 766 |
etac ctxt arg_cong 1)) |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
767 |
|> apply2 (Thm.close_derivation) |
53917 | 768 |
end; |
769 |
||
770 |
val split_lhs = q $ ufcase; |
|
771 |
||
772 |
fun mk_split_conjunct xctr xs f_xs = |
|
773 |
list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); |
|
774 |
fun mk_split_disjunct xctr xs f_xs = |
|
775 |
list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), |
|
776 |
HOLogic.mk_not (q $ f_xs))); |
|
777 |
||
778 |
fun mk_split_goal xctrs xss xfs = |
|
779 |
mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
780 |
(@{map 3} mk_split_conjunct xctrs xss xfs)); |
53917 | 781 |
fun mk_split_asm_goal xctrs xss xfs = |
782 |
mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
783 |
(@{map 3} mk_split_disjunct xctrs xss xfs))); |
53917 | 784 |
|
785 |
fun prove_split selss goal = |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
786 |
Variable.add_free_names lthy goal [] |
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
787 |
|> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => |
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
788 |
mk_split_tac ctxt 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
|
789 |
|> Thm.close_derivation; |
53917 | 790 |
|
791 |
fun prove_split_asm asm_goal split_thm = |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
792 |
Variable.add_free_names lthy asm_goal [] |
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
793 |
|> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => |
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
794 |
mk_split_asm_tac ctxt split_thm)) |
55409
b74248961819
made 'ctr_sugar' more friendly to the 'datatype_realizer'
blanchet
parents:
55407
diff
changeset
|
795 |
|> Thm.close_derivation; |
53917 | 796 |
|
797 |
val (split_thm, split_asm_thm) = |
|
798 |
let |
|
799 |
val goal = mk_split_goal xctrs xss xfs; |
|
800 |
val asm_goal = mk_split_asm_goal xctrs xss xfs; |
|
801 |
||
802 |
val thm = prove_split (replicate n []) goal; |
|
803 |
val asm_thm = prove_split_asm asm_goal thm; |
|
804 |
in |
|
805 |
(thm, asm_thm) |
|
806 |
end; |
|
807 |
||
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
808 |
val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
809 |
discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
810 |
exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
811 |
expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = |
52969
f2df0730f8ac
clarified option name (since case/fold/rec are also destructors)
blanchet
parents:
52968
diff
changeset
|
812 |
if no_discs_sels then |
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
813 |
([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) |
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
|
814 |
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
|
815 |
let |
53917 | 816 |
val udiscs = map (rapp u) discs; |
817 |
val uselss = map (map (rapp u)) selss; |
|
818 |
val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; |
|
819 |
val usel_fs = map2 (curry Term.list_comb) fs uselss; |
|
820 |
||
821 |
val vdiscs = map (rapp v) discs; |
|
822 |
val vselss = map (map (rapp v)) selss; |
|
823 |
||
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
|
824 |
fun make_sel_thm xs' case_thm sel_def = |
59647
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
wenzelm
parents:
59621
diff
changeset
|
825 |
zero_var_indexes |
60822 | 826 |
(Variable.gen_all lthy |
59647
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
wenzelm
parents:
59621
diff
changeset
|
827 |
(Drule.rename_bvars' |
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
wenzelm
parents:
59621
diff
changeset
|
828 |
(map (SOME o fst) xs') |
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
wenzelm
parents:
59621
diff
changeset
|
829 |
(Drule.forall_intr_vars (case_thm RS (sel_def RS trans))))); |
49281 | 830 |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
831 |
val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; |
53704 | 832 |
|
49281 | 833 |
fun has_undefined_rhs thm = |
59582 | 834 |
(case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of |
49281 | 835 |
Const (@{const_name undefined}, _) => true |
836 |
| _ => false); |
|
837 |
||
838 |
val all_sel_thms = |
|
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
839 |
(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
|
840 |
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
|
841 |
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
|
842 |
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
|
843 |
(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
|
844 |
|> filter_out has_undefined_rhs; |
49278 | 845 |
|
846 |
fun mk_unique_disc_def () = |
|
847 |
let |
|
848 |
val m = the_single ms; |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
849 |
val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
850 |
val vars = Variable.add_free_names lthy goal []; |
49278 | 851 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
852 |
Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => |
60728 | 853 |
mk_unique_disc_def_tac ctxt m uexhaust_thm) |
49692 | 854 |
|> Thm.close_derivation |
49278 | 855 |
end; |
856 |
||
857 |
fun mk_alternate_disc_def k = |
|
858 |
let |
|
859 |
val goal = |
|
49486 | 860 |
mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), |
861 |
nth exist_xs_u_eq_ctrs (k - 1)); |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
862 |
val vars = Variable.add_free_names lthy goal []; |
49278 | 863 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
864 |
Goal.prove_sorry lthy vars [] 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
|
865 |
mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) |
49486 | 866 |
(nth distinct_thms (2 - k)) uexhaust_thm) |
49692 | 867 |
|> Thm.close_derivation |
49278 | 868 |
end; |
49028 | 869 |
|
49278 | 870 |
val has_alternate_disc_def = |
871 |
exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; |
|
872 |
||
63313
0c956a9a0ac0
avoid runtime warning with discriminators due to 'Code.del_eqn'
blanchet
parents:
63239
diff
changeset
|
873 |
val nontriv_disc_defs = disc_defs |
63853 | 874 |
|> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, |
875 |
refl]); |
|
63313
0c956a9a0ac0
avoid runtime warning with discriminators due to 'Code.del_eqn'
blanchet
parents:
63239
diff
changeset
|
876 |
|
49278 | 877 |
val disc_defs' = |
878 |
map2 (fn k => fn def => |
|
879 |
if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () |
|
880 |
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
|
881 |
else def) ks disc_defs; |
49278 | 882 |
|
883 |
val discD_thms = map (fn def => def RS iffD1) disc_defs'; |
|
884 |
val discI_thms = |
|
885 |
map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms |
|
886 |
disc_defs'; |
|
887 |
val not_discI_thms = |
|
888 |
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
|
889 |
(unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) |
49278 | 890 |
ms disc_defs'; |
891 |
||
892 |
val (disc_thmss', disc_thmss) = |
|
893 |
let |
|
894 |
fun mk_thm discI _ [] = refl RS discI |
|
895 |
| mk_thm _ not_discI [distinct] = distinct RS not_discI; |
|
896 |
fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; |
|
897 |
in |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
898 |
@{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose |
49278 | 899 |
end; |
900 |
||
55464 | 901 |
val nontriv_disc_thmss = |
902 |
map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; |
|
53704 | 903 |
|
62326
3cf7a067599c
allow predicator instead of map function in 'primrec'
blanchet
parents:
62322
diff
changeset
|
904 |
fun is_discI_triv b = |
53704 | 905 |
(n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); |
906 |
||
907 |
val nontriv_discI_thms = |
|
62326
3cf7a067599c
allow predicator instead of map function in 'primrec'
blanchet
parents:
62322
diff
changeset
|
908 |
flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings |
53704 | 909 |
discI_thms); |
49028 | 910 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
911 |
val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = |
49486 | 912 |
let |
913 |
fun mk_goal [] = [] |
|
914 |
| mk_goal [((_, udisc), (_, udisc'))] = |
|
915 |
[Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, |
|
916 |
HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; |
|
917 |
||
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
918 |
fun prove tac goal = |
60728 | 919 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |
49667
44d85dc8ca08
use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents:
49633
diff
changeset
|
920 |
|> Thm.close_derivation; |
49486 | 921 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
922 |
val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); |
49486 | 923 |
|
924 |
val half_goalss = map mk_goal half_pairss; |
|
925 |
val half_thmss = |
|
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
926 |
@{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => |
60728 | 927 |
fn disc_thm => [prove (fn ctxt => |
928 |
mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) |
|
49486 | 929 |
half_goalss half_pairss (flat disc_thmss'); |
49278 | 930 |
|
49486 | 931 |
val other_half_goalss = map (mk_goal o map swap) half_pairss; |
932 |
val other_half_thmss = |
|
60728 | 933 |
map2 (map2 (fn thm => prove (fn ctxt => |
934 |
mk_other_half_distinct_disc_tac ctxt thm))) half_thmss |
|
49486 | 935 |
other_half_goalss; |
936 |
in |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49536
diff
changeset
|
937 |
join_halves n half_thmss other_half_thmss ||> `transpose |
49486 | 938 |
|>> has_alternate_disc_def ? K [] |
939 |
end; |
|
49278 | 940 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
941 |
val exhaust_disc_thm = |
49486 | 942 |
let |
943 |
fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; |
|
944 |
val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); |
|
945 |
in |
|
60728 | 946 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
947 |
mk_exhaust_disc_tac ctxt 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
|
948 |
|> Thm.close_derivation |
49486 | 949 |
end; |
49028 | 950 |
|
53740 | 951 |
val (safe_collapse_thms, all_collapse_thms) = |
49486 | 952 |
let |
54008
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
blanchet
parents:
54007
diff
changeset
|
953 |
fun mk_goal m udisc usel_ctr = |
49486 | 954 |
let |
955 |
val prem = HOLogic.mk_Trueprop udisc; |
|
53916 | 956 |
val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); |
49486 | 957 |
in |
53740 | 958 |
(prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) |
49486 | 959 |
end; |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
960 |
val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; |
53740 | 961 |
val thms = |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
962 |
@{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => |
53740 | 963 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58685
diff
changeset
|
964 |
mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |
53740 | 965 |
|> Thm.close_derivation |
966 |
|> not triv ? perhaps (try (fn thm => refl RS thm))) |
|
967 |
ms discD_thms sel_thmss trivs goals; |
|
49486 | 968 |
in |
53740 | 969 |
(map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), |
970 |
thms) |
|
49486 | 971 |
end; |
49025 | 972 |
|
53916 | 973 |
val swapped_all_collapse_thms = |
974 |
map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; |
|
975 |
||
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
976 |
val exhaust_sel_thm = |
53916 | 977 |
let |
978 |
fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; |
|
979 |
val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); |
|
980 |
in |
|
60728 | 981 |
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => |
982 |
mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |
|
53916 | 983 |
|> Thm.close_derivation |
984 |
end; |
|
985 |
||
53919 | 986 |
val expand_thm = |
49486 | 987 |
let |
988 |
fun mk_prems k udisc usels vdisc vsels = |
|
989 |
(if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ |
|
990 |
(if null usels then |
|
991 |
[] |
|
992 |
else |
|
993 |
[Logic.list_implies |
|
994 |
(if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], |
|
995 |
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj |
|
996 |
(map2 (curry HOLogic.mk_eq) usels vsels)))]); |
|
997 |
||
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
998 |
val goal = |
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49586
diff
changeset
|
999 |
Library.foldr Logic.list_implies |
58634
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents:
58335
diff
changeset
|
1000 |
(@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); |
49486 | 1001 |
val uncollapse_thms = |
53740 | 1002 |
map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1003 |
val vars = Variable.add_free_names lthy goal []; |
49486 | 1004 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1005 |
Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => |
61271 | 1006 |
mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1007 |
(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
|
1008 |
distinct_disc_thmsss') |
53919 | 1009 |
|> Thm.close_derivation |
49486 | 1010 |
end; |
49278 | 1011 |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1012 |
val (split_sel_thm, split_sel_asm_thm) = |
53917 | 1013 |
let |
1014 |
val zss = map (K []) xss; |
|
1015 |
val goal = mk_split_goal usel_ctrs zss usel_fs; |
|
1016 |
val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; |
|
1017 |
||
1018 |
val thm = prove_split sel_thmss goal; |
|
1019 |
val asm_thm = prove_split_asm asm_goal thm; |
|
1020 |
in |
|
1021 |
(thm, asm_thm) |
|
1022 |
end; |
|
1023 |
||
54491 | 1024 |
val case_eq_if_thm = |
49486 | 1025 |
let |
53917 | 1026 |
val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1027 |
val vars = Variable.add_free_names lthy goal []; |
49486 | 1028 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1029 |
Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => |
54491 | 1030 |
mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |
53919 | 1031 |
|> Thm.close_derivation |
49486 | 1032 |
end; |
59271 | 1033 |
|
1034 |
val disc_eq_case_thms = |
|
1035 |
let |
|
59272 | 1036 |
fun const_of_bool b = if b then @{const True} else @{const False}; |
59271 | 1037 |
fun mk_case_args n = map_index (fn (k, argTs) => |
59272 | 1038 |
fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; |
59271 | 1039 |
val goals = map_index (fn (n, udisc) => |
1040 |
mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1041 |
val goal = Logic.mk_conjunction_balanced goals; |
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1042 |
val vars = Variable.add_free_names lthy goal []; |
59271 | 1043 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1044 |
Goal.prove_sorry lthy vars [] goal |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
1045 |
(fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) |
59271 | 1046 |
exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |
61116
6189d179c2b5
close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
wenzelm
parents:
61101
diff
changeset
|
1047 |
|> Thm.close_derivation |
59271 | 1048 |
|> Conjunction.elim_balanced (length goals) |
1049 |
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
|
1050 |
in |
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1051 |
(sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1052 |
discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1053 |
[exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1054 |
[expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1055 |
disc_eq_case_thms) |
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
|
1056 |
end; |
49025 | 1057 |
|
59266 | 1058 |
val case_distrib_thm = |
1059 |
let |
|
1060 |
val args = @{map 2} (fn f => fn argTs => |
|
1061 |
let val (args, _) = mk_Frees "x" argTs lthy in |
|
1062 |
fold_rev Term.lambda args (h $ list_comb (f, args)) |
|
1063 |
end) fs ctr_Tss; |
|
1064 |
val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); |
|
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1065 |
val vars = Variable.add_free_names lthy goal []; |
59266 | 1066 |
in |
61334
8d40ddaa427f
collect the names from goals in favor of fragile exports
traytel
parents:
61301
diff
changeset
|
1067 |
Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
1068 |
mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |
59266 | 1069 |
|> Thm.close_derivation |
1070 |
end; |
|
1071 |
||
49437 | 1072 |
val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); |
53908 | 1073 |
val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); |
49300 | 1074 |
|
55464 | 1075 |
val nontriv_disc_eq_thmss = |
1076 |
map (map (fn th => th RS @{thm eq_False[THEN iffD2]} |
|
1077 |
handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; |
|
1078 |
||
54151 | 1079 |
val anonymous_notes = |
1080 |
[(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1081 |
(flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |
54151 | 1082 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
1083 |
||
49052 | 1084 |
val notes = |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1085 |
[(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), |
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49591
diff
changeset
|
1086 |
(case_congN, [case_cong_thm], []), |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1087 |
(case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), |
59266 | 1088 |
(case_distribN, [case_distrib_thm], []), |
54491 | 1089 |
(case_eq_ifN, case_eq_if_thms, []), |
58151
414deb2ef328
take out 'x = C' of the simplifier for unit types
blanchet
parents:
58116
diff
changeset
|
1090 |
(collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), |
55464 | 1091 |
(discN, flat nontriv_disc_thmss, simp_attrs), |
59272 | 1092 |
(disc_eq_caseN, disc_eq_case_thms, []), |
53700 | 1093 |
(discIN, nontriv_discI_thms, []), |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
1094 |
(distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), |
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1095 |
(distinct_discN, distinct_disc_thms, dest_attrs), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1096 |
(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
|
1097 |
(exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1098 |
(exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), |
49486 | 1099 |
(expandN, expand_thms, []), |
54145
297d1c603999
make sure that registered code equations are actually equations
blanchet
parents:
54008
diff
changeset
|
1100 |
(injectN, inject_thms, iff_attrs @ inductsimp_attrs), |
49300 | 1101 |
(nchotomyN, [nchotomy_thm], []), |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1102 |
(selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), |
57985 | 1103 |
(splitN, [split_thm], []), |
1104 |
(split_asmN, [split_asm_thm], []), |
|
57983
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1105 |
(split_selN, split_sel_thms, []), |
6edc3529bb4e
reordered some (co)datatype property names for more consistency
blanchet
parents:
57882
diff
changeset
|
1106 |
(split_sel_asmN, split_sel_asm_thms, []), |
57985 | 1107 |
(split_selsN, split_sel_thms @ split_sel_asm_thms, []), |
1108 |
(splitsN, [split_thm, split_asm_thm], [])] |
|
49300 | 1109 |
|> filter_out (null o #2) |
1110 |
|> map (fn (thmN, thms, attrs) => |
|
49633 | 1111 |
((qualify true (Binding.name thmN), attrs), [(thms, [])])); |
49300 | 1112 |
|
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1113 |
val (noted, lthy') = lthy |
57629
e88b5f59cade
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
blanchet
parents:
57260
diff
changeset
|
1114 |
|> 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
|
1115 |
|> 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
|
1116 |
(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
|
1117 |
|> 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
|
1118 |
(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
|
1119 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
58335
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58317
diff
changeset
|
1120 |
(fn phi => Case_Translation.register |
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
blanchet
parents:
58317
diff
changeset
|
1121 |
(Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |
58191 | 1122 |
|> plugins code_plugin ? |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1123 |
(Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1124 |
#> Local_Theory.declaration {syntax = false, pervasive = false} |
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1125 |
(fn phi => Context.mapping |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1126 |
(add_ctr_code fcT_name (map (Morphism.typ phi) As) |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1127 |
(map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) |
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63853
diff
changeset
|
1128 |
(Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
64430
diff
changeset
|
1129 |
I)) |
57633 | 1130 |
|> Local_Theory.notes (anonymous_notes @ notes) |
58317 | 1131 |
(* for "datatype_realizer.ML": *) |
57633 | 1132 |
|>> 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
|
1133 |
|
53867
8ad44ecc0d15
keep a database of free constructor type information
blanchet
parents:
53864
diff
changeset
|
1134 |
val ctr_sugar = |
58675 | 1135 |
{kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, |
1136 |
exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, |
|
1137 |
distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, |
|
59267 | 1138 |
case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], |
63313
0c956a9a0ac0
avoid runtime warning with discriminators due to 'Code.del_eqn'
blanchet
parents:
63239
diff
changeset
|
1139 |
split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, |
59271 | 1140 |
disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, |
1141 |
sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, |
|
1142 |
exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, |
|
1143 |
collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, |
|
1144 |
split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_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
|
1145 |
|> morph_ctr_sugar (substitute_noted_thm noted); |
49019 | 1146 |
in |
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
1147 |
(ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) |
49019 | 1148 |
end; |
49017 | 1149 |
in |
61101
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
traytel
parents:
60822
diff
changeset
|
1150 |
(goalss, after_qed, lthy) |
49017 | 1151 |
end; |
1152 |
||
58675 | 1153 |
fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => |
51551 | 1154 |
map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss |
58675 | 1155 |
|> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); |
49280
52413dc96326
allow default values for selectors in low-level "wrap_data" command
blanchet
parents:
49278
diff
changeset
|
1156 |
|
58675 | 1157 |
fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => |
49297 | 1158 |
Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo |
58675 | 1159 |
prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; |
49297 | 1160 |
|
57091 | 1161 |
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
|
1162 |
|
58660 | 1163 |
type ctr_options = Plugin_Name.filter * bool; |
1164 |
type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
1165 |
|
58660 | 1166 |
val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); |
1167 |
val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); |
|
58189
9d714be4f028
added 'plugins' option to control which hooks are enabled
blanchet
parents:
58188
diff
changeset
|
1168 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1169 |
val parse_ctr_options = |
54626 | 1170 |
Scan.optional (@{keyword "("} |-- Parse.list1 |
59281 | 1171 |
(Plugin_Name.parse_filter >> (apfst o K) |
1172 |
|| Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| |
|
55410 | 1173 |
@{keyword ")"} |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
1174 |
>> (fn fs => fold I fs default_ctr_options_cmd)) |
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
58634
diff
changeset
|
1175 |
default_ctr_options_cmd; |
49278 | 1176 |
|
55469
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents:
55468
diff
changeset
|
1177 |
fun parse_ctr_spec parse_ctr parse_arg = |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57094
diff
changeset
|
1178 |
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
|
1179 |
|
57091 | 1180 |
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
|
1181 |
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
|
1182 |
|
49017 | 1183 |
val _ = |
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59859
diff
changeset
|
1184 |
Outer_Syntax.local_theory_to_proof @{command_keyword free_constructors} |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55464
diff
changeset
|
1185 |
"register an existing freely generated type's constructors" |
57091 | 1186 |
(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
|
1187 |
-- parse_sel_default_eqs |
58675 | 1188 |
>> free_constructors_cmd Unknown); |
49017 | 1189 |
|
1190 |
end; |