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