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