src/HOL/BNF/Tools/bnf_ctr_sugar.ML
author blanchet
Fri, 16 Aug 2013 18:56:33 +0200
changeset 53039 476db75906ae
parent 52988 d1bdc6a97460
child 53040 e6edd7abc4ce
permissions -rw-r--r--
eliminate quasi-duplicate function
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51790
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_ctr_sugar.ML
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     3
    Copyright   2012
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
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51790
diff changeset
     8
signature BNF_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,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    26
     sel_thmss: thm list list};
51809
d4c1abbb4095 code tuning
blanchet
parents: 51798
diff changeset
    27
51840
b304fb6c5ef5 renamed records
blanchet
parents: 51839
diff changeset
    28
  val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
51823
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    29
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
    30
  val rep_compat_prefix: string
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
    31
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
    32
  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
    33
  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
    34
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    35
  val mk_ctr: typ list -> term -> term
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
    36
  val mk_disc_or_sel: typ list -> term -> term
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
    37
49622
a93f976c3307 use a nicer scheme to indexify names
blanchet
parents: 49619
diff changeset
    38
  val name_of_ctr: term -> string
51777
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
    39
  val name_of_disc: term -> string
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
    40
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
    41
  val wrap_free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list ->
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
    42
    (((bool * bool) * term list) * binding) *
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    43
      (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
51840
b304fb6c5ef5 renamed records
blanchet
parents: 51839
diff changeset
    44
    ctr_sugar * local_theory
52823
blanchet
parents: 52375
diff changeset
    45
  val parse_wrap_free_constructors_options: (bool * bool) parser
49286
dde4967c9233 added "defaults" option
blanchet
parents: 49285
diff changeset
    46
  val parse_bound_term: (binding * string) parser
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    47
end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    48
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51790
diff changeset
    49
structure BNF_Ctr_Sugar : BNF_CTR_SUGAR =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    50
struct
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    51
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    52
open BNF_Util
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51790
diff changeset
    53
open BNF_Ctr_Sugar_Tactics
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    54
51840
b304fb6c5ef5 renamed records
blanchet
parents: 51839
diff changeset
    55
type ctr_sugar =
51839
5c552de1d8d1 added constructors to data structure
blanchet
parents: 51823
diff changeset
    56
  {ctrs: term list,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    57
   casex: term,
51839
5c552de1d8d1 added constructors to data structure
blanchet
parents: 51823
diff changeset
    58
   discs: term list,
51819
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    59
   selss: term list list,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    60
   exhaust: thm,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    61
   nchotomy: thm,
51819
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    62
   injects: thm list,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    63
   distincts: thm list,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    64
   case_thms: thm list,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    65
   case_cong: thm,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    66
   weak_case_cong: thm,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    67
   split: thm,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    68
   split_asm: thm,
51819
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    69
   disc_thmss: thm list list,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    70
   discIs: thm list,
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
    71
   sel_thmss: thm list list};
51809
d4c1abbb4095 code tuning
blanchet
parents: 51798
diff changeset
    72
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    73
fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    74
    case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss} =
51839
5c552de1d8d1 added constructors to data structure
blanchet
parents: 51823
diff changeset
    75
  {ctrs = map (Morphism.term phi) ctrs,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    76
   casex = Morphism.term phi casex,
51839
5c552de1d8d1 added constructors to data structure
blanchet
parents: 51823
diff changeset
    77
   discs = map (Morphism.term phi) discs,
51823
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    78
   selss = map (map (Morphism.term phi)) selss,
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    79
   exhaust = Morphism.thm phi exhaust,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    80
   nchotomy = Morphism.thm phi nchotomy,
51823
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    81
   injects = map (Morphism.thm phi) injects,
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    82
   distincts = map (Morphism.thm phi) distincts,
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    83
   case_thms = map (Morphism.thm phi) case_thms,
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    84
   case_cong = Morphism.thm phi case_cong,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    85
   weak_case_cong = Morphism.thm phi weak_case_cong,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    86
   split = Morphism.thm phi split,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
    87
   split_asm = Morphism.thm phi split_asm,
51823
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    88
   disc_thmss = map (map (Morphism.thm phi)) disc_thmss,
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    89
   discIs = map (Morphism.thm phi) discIs,
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    90
   sel_thmss = map (map (Morphism.thm phi)) sel_thmss};
38996458bc5c create data structure for storing (co)datatype information
blanchet
parents: 51819
diff changeset
    91
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
    92
val rep_compat_prefix = "new";
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
    93
49223
blanchet
parents: 49210
diff changeset
    94
val isN = "is_";
blanchet
parents: 49210
diff changeset
    95
val unN = "un_";
blanchet
parents: 49210
diff changeset
    96
fun mk_unN 1 1 suf = unN ^ suf
blanchet
parents: 49210
diff changeset
    97
  | 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
    98
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
    99
val caseN = "case";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   100
val case_congN = "case_cong";
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   101
val case_convN = "case_conv";
49118
b815fa776b91 renamed theorem
blanchet
parents: 49117
diff changeset
   102
val collapseN = "collapse";
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
   103
val disc_excludeN = "disc_exclude";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   104
val disc_exhaustN = "disc_exhaust";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   105
val discsN = "discs";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   106
val distinctN = "distinct";
49075
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
   107
val exhaustN = "exhaust";
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   108
val expandN = "expand";
49075
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
   109
val injectN = "inject";
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
   110
val nchotomyN = "nchotomy";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   111
val selsN = "sels";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   112
val splitN = "split";
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   113
val splitsN = "splits";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   114
val split_asmN = "split_asm";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   115
val weak_case_cong_thmsN = "weak_case_cong";
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   116
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   117
val induct_simp_attrs = @{attributes [induct_simp]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   118
val cong_attrs = @{attributes [cong]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   119
val iff_attrs = @{attributes [iff]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   120
val safe_elim_attrs = @{attributes [elim!]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   121
val simp_attrs = @{attributes [simp]};
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
   122
53039
476db75906ae eliminate quasi-duplicate function
blanchet
parents: 52988
diff changeset
   123
fun unflat_lookup eq xs ys = map (fn xs' => mk_permute eq xs xs' ys);
49258
84f13469d7f0 allow same selector name for several constructors
blanchet
parents: 49257
diff changeset
   124
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   125
fun mk_half_pairss' _ ([], []) = []
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   126
  | mk_half_pairss' indent (x :: xs, _ :: ys) =
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   127
    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
   128
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   129
fun mk_half_pairss p = mk_half_pairss' [[]] p;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   130
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   131
fun join_halves n half_xss other_half_xss =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   132
  let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   133
    val xsss =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   134
      map2 (map2 append) (Library.chop_groups n half_xss)
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   135
        (transpose (Library.chop_groups n other_half_xss))
49668
blanchet
parents: 49667
diff changeset
   136
    val xs = splice (flat half_xss) (flat other_half_xss);
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   137
  in (xs, xsss) end;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   138
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   139
fun mk_undefined T = Const (@{const_name undefined}, T);
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   140
49500
blanchet
parents: 49498
diff changeset
   141
fun mk_ctr Ts t =
blanchet
parents: 49498
diff changeset
   142
  let val Type (_, Ts0) = body_type (fastype_of t) in
blanchet
parents: 49498
diff changeset
   143
    Term.subst_atomic_types (Ts0 ~~ Ts) t
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   144
  end;
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   145
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   146
fun mk_disc_or_sel Ts t =
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   147
  Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   148
49536
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   149
fun mk_case Ts T t =
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   150
  let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   151
    Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) t
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   152
  end;
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   153
51777
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   154
fun name_of_const what t =
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   155
  (case head_of t of
49622
a93f976c3307 use a nicer scheme to indexify names
blanchet
parents: 49619
diff changeset
   156
    Const (s, _) => s
a93f976c3307 use a nicer scheme to indexify names
blanchet
parents: 49619
diff changeset
   157
  | Free (s, _) => s
51777
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   158
  | _ => error ("Cannot extract name of " ^ what));
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   159
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   160
val name_of_ctr = name_of_const "constructor";
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   161
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   162
val notN = "not_";
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   163
val eqN = "eq_";
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   164
val neqN = "neq_";
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   165
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   166
fun name_of_disc t =
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   167
  (case head_of t of
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   168
    Abs (_, _, @{const Not} $ (t' $ Bound 0)) =>
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   169
    Long_Name.map_base_name (prefix notN) (name_of_disc t')
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   170
  | Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') =>
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   171
    Long_Name.map_base_name (prefix eqN) (name_of_disc t')
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   172
  | Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) =>
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   173
    Long_Name.map_base_name (prefix neqN) (name_of_disc t')
48a0ae342ea0 generate proper attributes for coinduction rules
blanchet
parents: 51771
diff changeset
   174
  | t' => name_of_const "destructor" t');
49622
a93f976c3307 use a nicer scheme to indexify names
blanchet
parents: 49619
diff changeset
   175
a93f976c3307 use a nicer scheme to indexify names
blanchet
parents: 49619
diff changeset
   176
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
   177
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   178
fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   179
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52968
diff changeset
   180
fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, rep_compat), raw_ctrs),
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   181
    raw_case_binding), (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   182
  let
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   183
    (* TODO: sanity checks on arguments *)
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   184
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   185
    val n = length raw_ctrs;
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   186
    val ks = 1 upto n;
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
   187
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   188
    val _ = if n > 0 then () else error "No constructors specified";
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   189
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   190
    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   191
    val sel_defaultss =
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   192
      pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   193
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   194
    val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
52988
d1bdc6a97460 avoid double qualification for case constants
blanchet
parents: 52969
diff changeset
   195
    val data_b_name = Long_Name.base_name dataT_name;
d1bdc6a97460 avoid double qualification for case constants
blanchet
parents: 52969
diff changeset
   196
    val data_b = Binding.name data_b_name;
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   197
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   198
    fun qualify mandatory =
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   199
      Binding.qualify mandatory data_b_name o
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   200
      (rep_compat ? Binding.qualify false rep_compat_prefix);
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   201
52965
0cd62cb233e0 handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents: 52963
diff changeset
   202
    fun tfree_name_of (TFree (s, _)) = s
0cd62cb233e0 handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents: 52963
diff changeset
   203
      | tfree_name_of (TVar ((s, _), _)) = s
0cd62cb233e0 handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents: 52963
diff changeset
   204
      | tfree_name_of (Type (s, _)) = Long_Name.base_name s;
0cd62cb233e0 handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents: 52963
diff changeset
   205
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   206
    val (As, B) =
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   207
      no_defs_lthy
52965
0cd62cb233e0 handle both TVars and TFrees -- necessary for 'wrap_free_constructors'
blanchet
parents: 52963
diff changeset
   208
      |> variant_tfrees (map tfree_name_of As0)
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   209
      ||> the_single o fst o mk_TFrees 1;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   210
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   211
    val dataT = Type (dataT_name, As);
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   212
    val ctrs = map (mk_ctr As) ctrs0;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   213
    val ctr_Tss = map (binder_types o fastype_of) ctrs;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   214
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   215
    val ms = map length ctr_Tss;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   216
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   217
    val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   218
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   219
    fun can_definitely_rely_on_disc k =
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   220
      not (Binding.is_empty (nth raw_disc_bindings' (k - 1)));
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   221
    fun can_rely_on_disc k =
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   222
      can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2));
51787
1267c28c7bdd changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents: 51781
diff changeset
   223
    fun should_omit_disc_binding k =
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   224
      n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k));
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   225
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   226
    fun is_disc_binding_valid b =
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   227
      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
   228
52322
74315fe137ba too much qualification is like too little
blanchet
parents: 52280
diff changeset
   229
    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
   230
49336
blanchet
parents: 49311
diff changeset
   231
    val disc_bindings =
blanchet
parents: 49311
diff changeset
   232
      raw_disc_bindings'
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   233
      |> 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
   234
        qualify false
51787
1267c28c7bdd changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents: 51781
diff changeset
   235
          (if Binding.is_empty disc then
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   236
             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
   237
           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
   238
             if m = 0 then disc
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   239
             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
   240
           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
   241
             standard_disc_binding ctr
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   242
           else
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   243
             disc)) ks ms ctrs0;
49056
blanchet
parents: 49055
diff changeset
   244
51787
1267c28c7bdd changed discriminator default: avoid mixing ctor and dtor views
blanchet
parents: 51781
diff changeset
   245
    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
   246
49336
blanchet
parents: 49311
diff changeset
   247
    val sel_bindingss =
blanchet
parents: 49311
diff changeset
   248
      pad_list [] n raw_sel_bindingss
49056
blanchet
parents: 49055
diff changeset
   249
      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   250
        qualify false
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   251
          (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
   252
            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
   253
          else
49434
433dc7e028c8 separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
traytel
parents: 49364
diff changeset
   254
            sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   255
49201
blanchet
parents: 49199
diff changeset
   256
    val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   257
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   258
    val ((((((((xss, xss'), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = no_defs_lthy |>
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
   259
      mk_Freess' "x" ctr_Tss
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   260
      ||>> mk_Freess "y" ctr_Tss
49201
blanchet
parents: 49199
diff changeset
   261
      ||>> mk_Frees "f" case_Ts
blanchet
parents: 49199
diff changeset
   262
      ||>> mk_Frees "g" case_Ts
49498
acc583e14167 tuned variable names
blanchet
parents: 49486
diff changeset
   263
      ||>> (apfst (map (rpair dataT)) oo Variable.variant_fixes) [data_b_name, data_b_name ^ "'"]
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   264
      ||>> mk_Frees "z" [B]
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   265
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   266
49498
acc583e14167 tuned variable names
blanchet
parents: 49486
diff changeset
   267
    val u = Free u';
acc583e14167 tuned variable names
blanchet
parents: 49486
diff changeset
   268
    val v = Free v';
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49458
diff changeset
   269
    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
   270
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   271
    val xctrs = map2 (curry Term.list_comb) ctrs xss;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   272
    val yctrs = map2 (curry Term.list_comb) ctrs yss;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   273
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   274
    val xfs = map2 (curry Term.list_comb) fs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   275
    val xgs = map2 (curry Term.list_comb) gs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   276
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   277
    (* 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
   278
       nicer names). Consider removing. *)
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   279
    val eta_fs = map2 eta_expand_arg xss xfs;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   280
    val eta_gs = map2 eta_expand_arg xss xgs;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   281
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   282
    val case_binding =
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   283
      qualify false
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   284
        (if Binding.is_empty raw_case_binding orelse
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   285
            Binding.eq_name (raw_case_binding, standard_binding) then
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   286
           Binding.suffix_name ("_" ^ caseN) data_b
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   287
         else
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   288
           raw_case_binding);
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   289
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   290
    fun mk_case_disj xctr xf xs =
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   291
      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
   292
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   293
    val case_rhs =
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   294
      fold_rev (fold_rev Term.lambda) [fs, [u]]
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   295
        (Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   296
           Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   297
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   298
    val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   299
      |> Local_Theory.define ((case_binding, NoSyn), ((Thm.def_binding case_binding, []), case_rhs))
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   300
      ||> `Local_Theory.restore;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   301
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   302
    val phi = Proof_Context.export_morphism lthy lthy';
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   303
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   304
    val case_def = Morphism.thm phi raw_case_def;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   305
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   306
    val case0 = Morphism.term phi raw_case;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   307
    val casex = mk_case As B case0;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   308
51759
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   309
    val fcase = Term.list_comb (casex, fs);
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   310
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   311
    val ufcase = fcase $ u;
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   312
    val vfcase = fcase $ v;
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   313
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   314
    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
   315
    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
   316
51759
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   317
    val eta_ufcase = eta_fcase $ u;
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   318
    val eta_vgcase = eta_gcase $ v;
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   319
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   320
    fun mk_uu_eq () = HOLogic.mk_eq (u, u);
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   321
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   322
    val uv_eq = mk_Trueprop_eq (u, v);
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   323
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   324
    val exist_xs_u_eq_ctrs =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   325
      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
   326
51743
blanchet
parents: 51742
diff changeset
   327
    val unique_disc_no_def = TrueI; (*arbitrary marker*)
blanchet
parents: 51742
diff changeset
   328
    val alternate_disc_no_def = FalseE; (*arbitrary marker*)
blanchet
parents: 51742
diff changeset
   329
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   330
    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
   331
      HOLogic.mk_not
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   332
        (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
   333
           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
   334
         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
   335
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   336
    val (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, disc_defs, sel_defs,
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   337
         sel_defss, lthy') =
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52968
diff changeset
   338
      if no_discs_sels then
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   339
        (true, [], [], [], [], [], [], [], [], [], lthy)
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   340
      else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   341
        let
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49458
diff changeset
   342
          fun disc_free b = Free (Binding.name_of b, mk_pred1T dataT);
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   343
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   344
          fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr);
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   345
49500
blanchet
parents: 49498
diff changeset
   346
          fun alternate_disc k =
blanchet
parents: 49498
diff changeset
   347
            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
   348
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   349
          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
   350
            map2 (fn Ts => fn k =>
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   351
              (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
   352
                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
   353
                (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
   354
                  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
   355
                | 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
   356
              | 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
   357
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   358
          fun sel_spec b proto_sels =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   359
            let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   360
              val _ =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   361
                (case duplicates (op =) (map fst proto_sels) of
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   362
                   k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   363
                     " for constructor " ^
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   364
                     quote (Syntax.string_of_term lthy (nth ctrs (k - 1))))
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   365
                 | [] => ())
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   366
              val T =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   367
                (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   368
                  [T] => T
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   369
                | T :: T' :: _ => error ("Inconsistent range type for selector " ^
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   370
                    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
   371
                    ^ quote (Syntax.string_of_typ lthy T')));
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   372
            in
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   373
              mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ u,
49536
898aea2e7a94 started work on generation of "rel" theorems
blanchet
parents: 49510
diff changeset
   374
                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
   375
            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
   376
49336
blanchet
parents: 49311
diff changeset
   377
          val sel_bindings = flat sel_bindingss;
blanchet
parents: 49311
diff changeset
   378
          val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
blanchet
parents: 49311
diff changeset
   379
          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
   380
49336
blanchet
parents: 49311
diff changeset
   381
          val sel_binding_index =
blanchet
parents: 49311
diff changeset
   382
            if all_sels_distinct then 1 upto length sel_bindings
blanchet
parents: 49311
diff changeset
   383
            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
   384
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
   385
          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
   386
          val sel_infos =
49336
blanchet
parents: 49311
diff changeset
   387
            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
   388
            |> sort (int_ord o pairself fst)
49336
blanchet
parents: 49311
diff changeset
   389
            |> map snd |> curry (op ~~) uniq_sel_bindings;
blanchet
parents: 49311
diff changeset
   390
          val sel_bindings = map fst sel_infos;
49258
84f13469d7f0 allow same selector name for several constructors
blanchet
parents: 49257
diff changeset
   391
49336
blanchet
parents: 49311
diff changeset
   392
          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
   393
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   394
          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
   395
            lthy
51809
d4c1abbb4095 code tuning
blanchet
parents: 51798
diff changeset
   396
            |> apfst split_list o fold_map3 (fn k => fn exist_xs_u_eq_ctr => fn b =>
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   397
                if Binding.is_empty b then
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   398
                  if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def)
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   399
                  else pair (alternate_disc k, alternate_disc_no_def)
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   400
                else if Binding.eq_name (b, equal_binding) then
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   401
                  pair (Term.lambda u exist_xs_u_eq_ctr, refl)
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   402
                else
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   403
                  Specification.definition (SOME (b, NONE, NoSyn),
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   404
                    ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
51809
d4c1abbb4095 code tuning
blanchet
parents: 51798
diff changeset
   405
              ks exist_xs_u_eq_ctrs disc_bindings
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   406
            ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   407
              Specification.definition (SOME (b, NONE, NoSyn),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   408
                ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   409
            ||> `Local_Theory.restore;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   410
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   411
          val phi = Proof_Context.export_morphism lthy lthy';
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   412
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   413
          val disc_defs = map (Morphism.thm phi) raw_disc_defs;
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   414
          val sel_defs = map (Morphism.thm phi) raw_sel_defs;
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   415
          val sel_defss = unflat_selss sel_defs;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   416
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   417
          val discs0 = map (Morphism.term phi) raw_discs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   418
          val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   419
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   420
          val discs = map (mk_disc_or_sel As) discs0;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   421
          val selss = map (map (mk_disc_or_sel As)) selss0;
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   422
49500
blanchet
parents: 49498
diff changeset
   423
          val udiscs = map (rapp u) discs;
blanchet
parents: 49498
diff changeset
   424
          val uselss = map (map (rapp u)) selss;
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   425
49500
blanchet
parents: 49498
diff changeset
   426
          val vdiscs = map (rapp v) discs;
blanchet
parents: 49498
diff changeset
   427
          val vselss = map (map (rapp v)) selss;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   428
        in
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   429
          (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, disc_defs, sel_defs,
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   430
           sel_defss, lthy')
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   431
        end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   432
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   433
    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   434
49458
blanchet
parents: 49438
diff changeset
   435
    val exhaust_goal =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   436
      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (u, xctr)]) in
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   437
        fold_rev Logic.all [p, u] (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
   438
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   439
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   440
    val inject_goalss =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   441
      let
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   442
        fun mk_goal _ _ [] [] = []
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   443
          | mk_goal xctr yctr xs ys =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   444
            [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
   445
              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   446
      in
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   447
        map4 mk_goal xctrs yctrs xss yss
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   448
      end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   449
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   450
    val half_distinct_goalss =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   451
      let
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   452
        fun mk_goal ((xs, xc), (xs', xc')) =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   453
          fold_rev Logic.all (xs @ xs')
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   454
            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   455
      in
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   456
        map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs)))
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   457
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   458
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   459
    val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   460
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   461
    fun after_qed thmss lthy =
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   462
      let
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   463
        val ([exhaust_thm], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   464
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   465
        val inject_thms = flat inject_thmss;
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   466
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   467
        val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   468
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   469
        fun inst_thm t thm =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   470
          Drule.instantiate' [] [SOME (certify lthy t)]
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   471
            (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes thm));
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   472
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   473
        val uexhaust_thm = inst_thm u exhaust_thm;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   474
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   475
        val exhaust_cases = map base_name_of_ctr ctrs;
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   476
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   477
        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
   478
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   479
        val (distinct_thms, (distinct_thmsss', distinct_thmsss)) =
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   480
          join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   481
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   482
        val nchotomy_thm =
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   483
          let
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   484
            val goal =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   485
              HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u',
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   486
                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
   487
          in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   488
            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
   489
            |> Thm.close_derivation
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   490
          end;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   491
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   492
        val case_thms =
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   493
          let
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   494
            val goals =
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   495
              map3 (fn xctr => fn xf => fn xs =>
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   496
                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
   497
          in
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   498
            map4 (fn k => fn goal => fn injects => fn distinctss =>
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   499
                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   500
                  mk_case_tac ctxt n k case_def injects distinctss)
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   501
                |> Thm.close_derivation)
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   502
              ks goals inject_thmss distinct_thmsss
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   503
          end;
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   504
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   505
        val (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   506
             disc_exhaust_thms, collapse_thms, expand_thms, case_conv_thms) =
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52968
diff changeset
   507
          if no_discs_sels then
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   508
            ([], [], [], [], [], [], [], [], [], [])
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
   509
          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
   510
            let
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
   511
              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
   512
                zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
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
   513
                    (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   514
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   515
              fun has_undefined_rhs thm =
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   516
                (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   517
                  Const (@{const_name undefined}, _) => true
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   518
                | _ => false);
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   519
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
   520
              val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   521
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   522
              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
   523
                (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
   524
                   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
   525
                 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
   526
                   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
   527
                     (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
   528
                |> filter_out has_undefined_rhs;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   529
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   530
              fun mk_unique_disc_def () =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   531
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   532
                  val m = the_single ms;
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   533
                  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
   534
                in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   535
                  Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   536
                  |> Thm.close_derivation
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   537
                  |> singleton (Proof_Context.export names_lthy lthy)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   538
                end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   539
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   540
              fun mk_alternate_disc_def k =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   541
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   542
                  val goal =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   543
                    mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k),
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   544
                      nth exist_xs_u_eq_ctrs (k - 1));
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   545
                in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   546
                  Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   547
                    mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   548
                      (nth distinct_thms (2 - k)) uexhaust_thm)
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   549
                  |> Thm.close_derivation
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   550
                  |> singleton (Proof_Context.export names_lthy lthy)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   551
                end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   552
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   553
              val has_alternate_disc_def =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   554
                exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   555
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   556
              val disc_defs' =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   557
                map2 (fn k => fn def =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   558
                  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
   559
                  else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   560
                  else def) ks disc_defs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   561
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   562
              val discD_thms = map (fn def => def RS iffD1) disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   563
              val discI_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   564
                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
   565
                  disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   566
              val not_discI_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   567
                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
   568
                    (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   569
                  ms disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   570
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   571
              val (disc_thmss', disc_thmss) =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   572
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   573
                  fun mk_thm discI _ [] = refl RS discI
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   574
                    | mk_thm _ not_discI [distinct] = distinct RS not_discI;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   575
                  fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   576
                in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   577
                  map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   578
                end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   579
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   580
              val disc_thms = flat (map2 (fn b => if is_disc_binding_valid b then I else K [])
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   581
                disc_bindings disc_thmss);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   582
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   583
              val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   584
                let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   585
                  fun mk_goal [] = []
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   586
                    | mk_goal [((_, udisc), (_, udisc'))] =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   587
                      [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc,
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   588
                         HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))];
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   589
49667
44d85dc8ca08 use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents: 49633
diff changeset
   590
                  fun prove tac goal =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   591
                    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
   592
                    |> Thm.close_derivation;
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   593
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   594
                  val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs));
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   595
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   596
                  val half_goalss = map mk_goal half_pairss;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   597
                  val half_thmss =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   598
                    map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51797
diff changeset
   599
                        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
   600
                      half_goalss half_pairss (flat disc_thmss');
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   601
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   602
                  val other_half_goalss = map (mk_goal o map swap) half_pairss;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   603
                  val other_half_thmss =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   604
                    map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   605
                      other_half_goalss;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   606
                in
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49536
diff changeset
   607
                  join_halves n half_thmss other_half_thmss ||> `transpose
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   608
                  |>> has_alternate_disc_def ? K []
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   609
                end;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   610
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   611
              val disc_exhaust_thm =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   612
                let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   613
                  fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc];
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   614
                  val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs));
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   615
                in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   616
                  Goal.prove_sorry lthy [] [] goal (fn _ =>
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   617
                    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
   618
                  |> Thm.close_derivation
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   619
                end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   620
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   621
              val (collapse_thms, collapse_thm_opts) =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   622
                let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   623
                  fun mk_goal ctr udisc usels =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   624
                    let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   625
                      val prem = HOLogic.mk_Trueprop udisc;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   626
                      val concl =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   627
                        mk_Trueprop_eq ((null usels ? swap) (Term.list_comb (ctr, usels), u));
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   628
                    in
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   629
                      if prem aconv concl then NONE
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   630
                      else SOME (Logic.all u (Logic.mk_implies (prem, concl)))
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   631
                    end;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   632
                  val goals = map3 mk_goal ctrs udiscs uselss;
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   633
                in
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   634
                  map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   635
                    Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   636
                      mk_collapse_tac ctxt m discD sel_thms)
49667
44d85dc8ca08 use Thm.close_derivation in theorems proved using Skip_Proof.prove; tuned signature;
traytel
parents: 49633
diff changeset
   637
                    |> Thm.close_derivation
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   638
                    |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   639
                  |> `(map_filter I)
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   640
                end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   641
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   642
              val expand_thms =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   643
                let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   644
                  fun mk_prems k udisc usels vdisc vsels =
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   645
                    (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   646
                    (if null usels then
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   647
                       []
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   648
                     else
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   649
                       [Logic.list_implies
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   650
                          (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc],
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   651
                             HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   652
                               (map2 (curry HOLogic.mk_eq) usels vsels)))]);
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   653
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   654
                  val goal =
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   655
                    Library.foldr Logic.list_implies
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   656
                      (map5 mk_prems ks udiscs uselss vdiscs vselss, uv_eq);
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   657
                  val uncollapse_thms =
51742
b5ff7393642d fix bugs in expand tactic w.r.t. datatypes with "needless" discriminators (e.g. lists with is_Nil instead of ~= Nil)
blanchet
parents: 51717
diff changeset
   658
                    map2 (fn NONE => K asm_rl | SOME thm => fn [] => thm | _ => thm RS sym)
b5ff7393642d fix bugs in expand tactic w.r.t. datatypes with "needless" discriminators (e.g. lists with is_Nil instead of ~= Nil)
blanchet
parents: 51717
diff changeset
   659
                      collapse_thm_opts uselss;
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   660
                in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   661
                  [Goal.prove_sorry lthy [] [] goal (fn _ =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51697
diff changeset
   662
                     mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   663
                       (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   664
                       disc_exclude_thmsss')]
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   665
                  |> map Thm.close_derivation
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   666
                  |> Proof_Context.export names_lthy lthy
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   667
                end;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   668
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   669
              val case_conv_thms =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   670
                let
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   671
                  fun mk_body f usels = Term.list_comb (f, usels);
51770
78162d9e977d no eta-expansion for case in split rules and case_conv
blanchet
parents: 51763
diff changeset
   672
                  val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs (map2 mk_body fs uselss));
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   673
                in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   674
                  [Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   675
                     mk_case_conv_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)]
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   676
                  |> map Thm.close_derivation
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   677
                  |> Proof_Context.export names_lthy lthy
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   678
                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
   679
            in
49484
0194a18f80cf finished "disc_coiter_iff" etc. generation
blanchet
parents: 49478
diff changeset
   680
              (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   681
               [disc_exhaust_thm], collapse_thms, expand_thms, case_conv_thms)
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   682
            end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   683
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   684
        val (case_cong_thm, weak_case_cong_thm) =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   685
          let
51763
0051318acc9d apply arguments to f and g in "case_cong"
blanchet
parents: 51759
diff changeset
   686
            fun mk_prem xctr xs xf xg =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   687
              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
51763
0051318acc9d apply arguments to f and g in "case_cong"
blanchet
parents: 51759
diff changeset
   688
                mk_Trueprop_eq (xf, xg)));
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   689
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   690
            val goal =
51763
0051318acc9d apply arguments to f and g in "case_cong"
blanchet
parents: 51759
diff changeset
   691
              Logic.list_implies (uv_eq :: map4 mk_prem xctrs xss xfs xgs,
51759
587bba425430 eta-contracted weak congruence rules (like in the old package)
blanchet
parents: 51757
diff changeset
   692
                 mk_Trueprop_eq (eta_ufcase, eta_vgcase));
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49586
diff changeset
   693
            val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase));
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   694
          in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51697
diff changeset
   695
            (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms),
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   696
             Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1)))
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   697
            |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy))
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   698
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   699
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   700
        val (split_thm, split_asm_thm) =
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   701
          let
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   702
            fun mk_conjunct xctr xs f_xs =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   703
              list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   704
            fun mk_disjunct xctr xs f_xs =
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   705
              list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   706
                HOLogic.mk_not (q $ f_xs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   707
51770
78162d9e977d no eta-expansion for case in split rules and case_conv
blanchet
parents: 51763
diff changeset
   708
            val lhs = q $ ufcase;
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   709
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   710
            val goal =
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   711
              mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
49458
blanchet
parents: 49438
diff changeset
   712
            val asm_goal =
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   713
              mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   714
                (map3 mk_disjunct xctrs xss xfs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   715
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   716
            val split_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   717
              Goal.prove_sorry lthy [] [] goal
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51697
diff changeset
   718
                (fn _ => mk_split_tac lthy uexhaust_thm case_thms inject_thmss distinct_thmsss)
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   719
              |> Thm.close_derivation
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   720
              |> singleton (Proof_Context.export names_lthy lthy);
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   721
            val split_asm_thm =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   722
              Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} =>
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   723
                mk_split_asm_tac ctxt split_thm)
49692
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   724
              |> Thm.close_derivation
a8a3b82b37f8 made code more conform to rest of BNF package
blanchet
parents: 49668
diff changeset
   725
              |> singleton (Proof_Context.export names_lthy lthy);
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   726
          in
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   727
            (split_thm, split_asm_thm)
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   728
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   729
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   730
        val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   731
        val cases_type_attr = Attrib.internal (K (Induct.cases_type dataT_name));
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   732
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   733
        val notes =
49594
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   734
          [(caseN, case_thms, simp_attrs),
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   735
           (case_congN, [case_cong_thm], []),
55e798614c45 tweaked theorem names (in particular, dropped s's)
blanchet
parents: 49591
diff changeset
   736
           (case_convN, case_conv_thms, []),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   737
           (collapseN, collapse_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   738
           (discsN, disc_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   739
           (disc_excludeN, disc_exclude_thms, []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   740
           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   741
           (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   742
           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
49486
64cc57c0d0fe generate "expand" property
blanchet
parents: 49484
diff changeset
   743
           (expandN, expand_thms, []),
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   744
           (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   745
           (nchotomyN, [nchotomy_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   746
           (selsN, all_sel_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   747
           (splitN, [split_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   748
           (split_asmN, [split_asm_thm], []),
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   749
           (splitsN, [split_thm, split_asm_thm], []),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   750
           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   751
          |> filter_out (null o #2)
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   752
          |> map (fn (thmN, thms, attrs) =>
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   753
            ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   754
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   755
        val notes' =
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   756
          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   757
          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   758
      in
52375
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
   759
        ({ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
   760
          nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
   761
          case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm,
bae65fd74633 store more theorems in data structure
blanchet
parents: 52322
diff changeset
   762
          split = split_thm, split_asm = split_asm_thm, disc_thmss = disc_thmss,
51819
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   763
          discIs = discI_thms, sel_thmss = sel_thmss},
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   764
         lthy
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   765
         |> not rep_compat ?
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   766
            (Local_Theory.declaration {syntax = false, pervasive = true}
51697
1ce319118d59 do not add case translation syntax in rep_datatype compatibility mode
traytel
parents: 51696
diff changeset
   767
               (fn phi => Case_Translation.register
51819
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   768
                  (Morphism.term phi casex) (map (Morphism.term phi) ctrs)))
9df935196be9 use record instead of big tuple
blanchet
parents: 51809
diff changeset
   769
         |> Local_Theory.notes (notes' @ notes) |> snd)
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   770
      end;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   771
  in
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   772
    (goalss, after_qed, lthy')
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   773
  end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   774
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   775
fun wrap_free_constructors tacss = (fn (goalss, after_qed, lthy) =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 51380
diff changeset
   776
  map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   777
  |> (fn thms => after_qed thms lthy)) oo prepare_wrap_free_constructors (K I);
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   778
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   779
val wrap_free_constructors_cmd = (fn (goalss, after_qed, lthy) =>
49297
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   780
  Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   781
  prepare_wrap_free_constructors Syntax.read_term;
49297
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   782
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   783
fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
49111
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   784
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   785
val parse_bindings = parse_bracket_list parse_binding;
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   786
val parse_bindingss = parse_bracket_list parse_bindings;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   787
51790
22517d04d20b more intuitive syntax for equality-style discriminators of nullary constructors
blanchet
parents: 51787
diff changeset
   788
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
   789
val parse_bound_terms = parse_bracket_list parse_bound_term;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   790
val parse_bound_termss = parse_bracket_list parse_bound_terms;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   791
52823
blanchet
parents: 52375
diff changeset
   792
val parse_wrap_free_constructors_options =
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52968
diff changeset
   793
  Scan.optional (@{keyword "("} |-- Parse.list1 ((@{keyword "no_discs_sels"} >> K (true, false)) ||
49633
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   794
      (@{keyword "rep_compat"} >> K (false, true))) --| @{keyword ")"}
5b5450bc544c compatibility option to use "rep_datatype"
blanchet
parents: 49622
diff changeset
   795
    >> (pairself (exists I) o split_list)) (false, false);
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   796
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   797
val _ =
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   798
  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_free_constructors"}
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51790
diff changeset
   799
    "wrap an existing freely generated type's constructors"
52823
blanchet
parents: 52375
diff changeset
   800
    ((parse_wrap_free_constructors_options -- (@{keyword "["} |-- Parse.list Parse.term --|
blanchet
parents: 52375
diff changeset
   801
        @{keyword "]"}) --
52968
2b430bbb5a1a define case constant from other 'free constructor' axioms
blanchet
parents: 52965
diff changeset
   802
      parse_binding -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   803
        Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
51781
0504e286d66d renamed "wrap_data" to "wrap_free_constructors"
blanchet
parents: 51777
diff changeset
   804
     >> wrap_free_constructors_cmd);
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   805
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   806
end;