src/HOL/Codatatype/Tools/bnf_wrap.ML
author blanchet
Thu, 20 Sep 2012 02:42:48 +0200
changeset 49458 9321a9465021
parent 49438 5bc80d96241e
child 49463 83ac281bcdc2
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
     1
(*  Title:      HOL/Codatatype/Tools/bnf_wrap.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
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
     5
Wrapping existing datatypes.
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     6
*)
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     7
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
     8
signature BNF_WRAP =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     9
sig
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    10
  val mk_half_pairss: 'a list -> ('a * 'a) list list
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    11
  val mk_ctr: typ list -> term -> term
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
    12
  val base_name_of_ctr: term -> string
49199
7c9a3c67c55d added high-level recursor, not yet curried
blanchet
parents: 49174
diff changeset
    13
  val wrap_datatype: ({prems: thm list, context: Proof.context} -> tactic) list list ->
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    14
    ((bool * term list) * term) *
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    15
      (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
    16
    (term list list * thm list * thm list * thm list * thm list * thm list list) * local_theory
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
    17
  val parse_wrap_options: bool parser
49286
dde4967c9233 added "defaults" option
blanchet
parents: 49285
diff changeset
    18
  val parse_bound_term: (binding * string) parser
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    19
end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    20
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
    21
structure BNF_Wrap : BNF_WRAP =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    22
struct
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    23
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    24
open BNF_Util
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
    25
open BNF_Wrap_Tactics
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    26
49223
blanchet
parents: 49210
diff changeset
    27
val isN = "is_";
blanchet
parents: 49210
diff changeset
    28
val unN = "un_";
blanchet
parents: 49210
diff changeset
    29
fun mk_unN 1 1 suf = unN ^ suf
blanchet
parents: 49210
diff changeset
    30
  | 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
    31
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    32
val case_congN = "case_cong";
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
    33
val case_eqN = "case_eq";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    34
val casesN = "cases";
49118
b815fa776b91 renamed theorem
blanchet
parents: 49117
diff changeset
    35
val collapseN = "collapse";
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
    36
val disc_excludeN = "disc_exclude";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    37
val disc_exhaustN = "disc_exhaust";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    38
val discsN = "discs";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    39
val distinctN = "distinct";
49075
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    40
val exhaustN = "exhaust";
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    41
val injectN = "inject";
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    42
val nchotomyN = "nchotomy";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    43
val selsN = "sels";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    44
val splitN = "split";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    45
val split_asmN = "split_asm";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    46
val weak_case_cong_thmsN = "weak_case_cong";
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    47
49336
blanchet
parents: 49311
diff changeset
    48
val std_binding = @{binding _};
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    49
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    50
val induct_simp_attrs = @{attributes [induct_simp]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    51
val cong_attrs = @{attributes [cong]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    52
val iff_attrs = @{attributes [iff]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    53
val safe_elim_attrs = @{attributes [elim!]};
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    54
val simp_attrs = @{attributes [simp]};
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    55
49056
blanchet
parents: 49055
diff changeset
    56
fun pad_list x n xs = xs @ replicate (n - length xs) x;
blanchet
parents: 49055
diff changeset
    57
49259
b21c03c7a097 minor optimization
blanchet
parents: 49258
diff changeset
    58
fun unflat_lookup eq ys zs = map (map (fn x => nth zs (find_index (curry eq x) ys)));
49258
84f13469d7f0 allow same selector name for several constructors
blanchet
parents: 49257
diff changeset
    59
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
    60
fun mk_half_pairss' _ [] = []
49056
blanchet
parents: 49055
diff changeset
    61
  | mk_half_pairss' indent (y :: ys) =
blanchet
parents: 49055
diff changeset
    62
    indent @ fold_rev (cons o single o pair y) ys (mk_half_pairss' ([] :: indent) ys);
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    63
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
    64
fun mk_half_pairss ys = mk_half_pairss' [[]] ys;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    65
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    66
fun mk_undefined T = Const (@{const_name undefined}, T);
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    67
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    68
fun mk_ctr Ts ctr =
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    69
  let val Type (_, Ts0) = body_type (fastype_of ctr) in
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    70
    Term.subst_atomic_types (Ts0 ~~ Ts) ctr
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    71
  end;
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
    72
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    73
fun base_name_of_ctr c =
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    74
  Long_Name.base_name (case head_of c of
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    75
      Const (s, _) => s
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    76
    | Free (s, _) => s
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    77
    | _ => error "Cannot extract name of constructor");
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    78
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
    79
fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
    80
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
    81
fun prepare_wrap_datatype prep_term (((no_dests, raw_ctrs), raw_case),
49336
blanchet
parents: 49311
diff changeset
    82
    (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    83
  let
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    84
    (* TODO: sanity checks on arguments *)
49113
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    85
    (* TODO: attributes (simp, case_names, etc.) *)
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    86
    (* TODO: case syntax *)
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    87
    (* TODO: integration with function package ("size") *)
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    88
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    89
    val n = length raw_ctrs;
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    90
    val ks = 1 upto n;
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    91
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    92
    val _ = if n > 0 then () else error "No constructors specified";
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    93
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    94
    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
    95
    val case0 = prep_term no_defs_lthy raw_case;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    96
    val sel_defaultss =
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
    97
      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
    98
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
    99
    val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   100
    val data_b = Binding.qualified_name dataT_name;
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   101
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   102
    val (As, B) =
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   103
      no_defs_lthy
49298
36e551d3af3b support for sort constraints in new (co)data commands
blanchet
parents: 49297
diff changeset
   104
      |> mk_TFrees' (map Type.sort_of_atyp As0)
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   105
      ||> the_single o fst o mk_TFrees 1;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   106
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   107
    val dataT = Type (dataT_name, As);
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   108
    val ctrs = map (mk_ctr As) ctrs0;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   109
    val ctr_Tss = map (binder_types o fastype_of) ctrs;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   110
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   111
    val ms = map length ctr_Tss;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
   112
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
   113
    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
   114
49174
41790d616f63 by default, only generate one discriminator for a two-value datatype
blanchet
parents: 49157
diff changeset
   115
    fun can_really_rely_on_disc k =
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
   116
      not (Binding.eq_name (nth raw_disc_bindings' (k - 1), Binding.empty)) orelse
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
   117
      nth ms (k - 1) = 0;
49174
41790d616f63 by default, only generate one discriminator for a two-value datatype
blanchet
parents: 49157
diff changeset
   118
    fun can_rely_on_disc k =
41790d616f63 by default, only generate one discriminator for a two-value datatype
blanchet
parents: 49157
diff changeset
   119
      can_really_rely_on_disc k orelse (k = 1 andalso not (can_really_rely_on_disc 2));
49336
blanchet
parents: 49311
diff changeset
   120
    fun can_omit_disc_binding k m =
49174
41790d616f63 by default, only generate one discriminator for a two-value datatype
blanchet
parents: 49157
diff changeset
   121
      n = 1 orelse m = 0 orelse (n = 2 andalso can_rely_on_disc (3 - k));
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   122
49336
blanchet
parents: 49311
diff changeset
   123
    val std_disc_binding =
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   124
      Binding.qualify false (Binding.name_of data_b) o Binding.name o prefix isN o base_name_of_ctr;
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   125
49336
blanchet
parents: 49311
diff changeset
   126
    val disc_bindings =
blanchet
parents: 49311
diff changeset
   127
      raw_disc_bindings'
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   128
      |> map4 (fn k => fn m => fn ctr => fn disc =>
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   129
        Option.map (Binding.qualify false (Binding.name_of data_b))
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
   130
          (if Binding.eq_name (disc, Binding.empty) then
49336
blanchet
parents: 49311
diff changeset
   131
             if can_omit_disc_binding k m then NONE else SOME (std_disc_binding ctr)
blanchet
parents: 49311
diff changeset
   132
           else if Binding.eq_name (disc, std_binding) then
blanchet
parents: 49311
diff changeset
   133
             SOME (std_disc_binding ctr)
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   134
           else
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   135
             SOME disc)) ks ms ctrs0;
49056
blanchet
parents: 49055
diff changeset
   136
49336
blanchet
parents: 49311
diff changeset
   137
    val no_discs = map is_none disc_bindings;
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   138
    val no_discs_at_all = forall I no_discs;
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
   139
49336
blanchet
parents: 49311
diff changeset
   140
    fun std_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
   141
49336
blanchet
parents: 49311
diff changeset
   142
    val sel_bindingss =
blanchet
parents: 49311
diff changeset
   143
      pad_list [] n raw_sel_bindingss
49056
blanchet
parents: 49055
diff changeset
   144
      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   145
        Binding.qualify false (Binding.name_of data_b)
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
   146
          (if Binding.eq_name (sel, Binding.empty) orelse Binding.eq_name (sel, std_binding) then
49336
blanchet
parents: 49311
diff changeset
   147
            std_sel_binding m l ctr
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   148
          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
   149
            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
   150
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   151
    fun mk_case Ts T =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   152
      let
49336
blanchet
parents: 49311
diff changeset
   153
        val (bindings, body) = strip_type (fastype_of case0)
blanchet
parents: 49311
diff changeset
   154
        val Type (_, Ts0) = List.last bindings
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   155
      in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   156
49201
blanchet
parents: 49199
diff changeset
   157
    val casex = mk_case As B;
blanchet
parents: 49199
diff changeset
   158
    val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   159
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
   160
    val ((((((((xss, xss'), yss), fs), gs), (v, v')), w), (p, p')), names_lthy) = no_defs_lthy |>
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
   161
      mk_Freess' "x" ctr_Tss
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   162
      ||>> mk_Freess "y" ctr_Tss
49201
blanchet
parents: 49199
diff changeset
   163
      ||>> mk_Frees "f" case_Ts
blanchet
parents: 49199
diff changeset
   164
      ||>> mk_Frees "g" case_Ts
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   165
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "v") dataT
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   166
      ||>> yield_singleton (mk_Frees "w") dataT
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   167
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   168
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   169
    val q = Free (fst p', B --> HOLogic.boolT);
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   170
49210
656fb50d33f0 define coiterators
blanchet
parents: 49203
diff changeset
   171
    fun ap_v t = t $ v;
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   172
    fun mk_v_eq_v () = HOLogic.mk_eq (v, v);
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   173
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   174
    val xctrs = map2 (curry Term.list_comb) ctrs xss;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   175
    val yctrs = map2 (curry Term.list_comb) ctrs yss;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   176
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   177
    val xfs = map2 (curry Term.list_comb) fs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   178
    val xgs = map2 (curry Term.list_comb) gs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   179
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   180
    val eta_fs = map2 eta_expand_arg xss xfs;
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   181
    val eta_gs = map2 eta_expand_arg xss xgs;
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   182
49201
blanchet
parents: 49199
diff changeset
   183
    val fcase = Term.list_comb (casex, eta_fs);
blanchet
parents: 49199
diff changeset
   184
    val gcase = Term.list_comb (casex, eta_gs);
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   185
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   186
    val exist_xs_v_eq_ctrs =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   187
      map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (v, xctr))) xctrs xss;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   188
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   189
    val unique_disc_no_def = TrueI; (*arbitrary marker*)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   190
    val alternate_disc_no_def = FalseE; (*arbitrary marker*)
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
   191
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   192
    fun alternate_disc_lhs get_disc 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
   193
      HOLogic.mk_not
49336
blanchet
parents: 49311
diff changeset
   194
        (case nth disc_bindings (k - 1) of
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   195
          NONE => nth exist_xs_v_eq_ctrs (k - 1)
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   196
        | SOME b => get_disc b (k - 1) $ v);
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
   197
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
   198
    val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   199
      if no_dests then
49285
036b833b99aa removed wrong "transpose" and ensure "sel" theorems are put in the right order (grouped per selector, in the order in which the selectors appear)
blanchet
parents: 49281
diff changeset
   200
        (true, [], [], [], [], [], no_defs_lthy)
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   201
      else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   202
        let
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   203
          fun disc_free b = Free (Binding.name_of b, dataT --> HOLogic.boolT);
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   204
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   205
          fun disc_spec b exist_xs_v_eq_ctr = mk_Trueprop_eq (disc_free b $ v, exist_xs_v_eq_ctr);
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   206
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   207
          fun alternate_disc k = Term.lambda v (alternate_disc_lhs (K o disc_free) (3 - k));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   208
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
   209
          fun mk_default T t =
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
   210
            let
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
   211
              val Ts0 = map TFree (Term.add_tfreesT (fastype_of t) []);
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
   212
              val Ts = map TFree (Term.add_tfreesT T []);
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
   213
            in Term.subst_atomic_types (Ts0 ~~ Ts) t end;
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
   214
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   215
          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
   216
            map2 (fn Ts => fn k =>
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   217
              (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
   218
                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
   219
                (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
   220
                  NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
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
   221
                | SOME t => mk_default (Ts ---> T) t)
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   222
              | 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
   223
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   224
          fun sel_spec b proto_sels =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   225
            let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   226
              val _ =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   227
                (case duplicates (op =) (map fst proto_sels) of
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   228
                   k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   229
                     " for constructor " ^
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   230
                     quote (Syntax.string_of_term no_defs_lthy (nth ctrs (k - 1))))
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   231
                 | [] => ())
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   232
              val T =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   233
                (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   234
                  [T] => T
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   235
                | T :: T' :: _ => error ("Inconsistent range type for selector " ^
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   236
                    quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   237
                    " vs. " ^ quote (Syntax.string_of_typ no_defs_lthy T')));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   238
            in
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   239
              mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ v,
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   240
                Term.list_comb (mk_case As T, mk_sel_case_args b proto_sels T) $ v)
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   241
            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
   242
49336
blanchet
parents: 49311
diff changeset
   243
          val sel_bindings = flat sel_bindingss;
blanchet
parents: 49311
diff changeset
   244
          val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
blanchet
parents: 49311
diff changeset
   245
          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
   246
49336
blanchet
parents: 49311
diff changeset
   247
          val sel_binding_index =
blanchet
parents: 49311
diff changeset
   248
            if all_sels_distinct then 1 upto length sel_bindings
blanchet
parents: 49311
diff changeset
   249
            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
   250
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
   251
          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
   252
          val sel_infos =
49336
blanchet
parents: 49311
diff changeset
   253
            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
   254
            |> sort (int_ord o pairself fst)
49336
blanchet
parents: 49311
diff changeset
   255
            |> map snd |> curry (op ~~) uniq_sel_bindings;
blanchet
parents: 49311
diff changeset
   256
          val sel_bindings = map fst sel_infos;
49258
84f13469d7f0 allow same selector name for several constructors
blanchet
parents: 49257
diff changeset
   257
49336
blanchet
parents: 49311
diff changeset
   258
          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
   259
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   260
          val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   261
            no_defs_lthy
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   262
            |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_v_eq_ctr =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   263
              fn NONE =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   264
                 if n = 1 then pair (Term.lambda v (mk_v_eq_v ()), unique_disc_no_def)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   265
                 else if m = 0 then pair (Term.lambda v exist_xs_v_eq_ctr, refl)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   266
                 else pair (alternate_disc k, alternate_disc_no_def)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   267
               | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   268
                   ((Thm.def_binding b, []), disc_spec b exist_xs_v_eq_ctr)) #>> apsnd snd)
49336
blanchet
parents: 49311
diff changeset
   269
              ks ms exist_xs_v_eq_ctrs disc_bindings
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   270
            ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   271
              Specification.definition (SOME (b, NONE, NoSyn),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   272
                ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   273
            ||> `Local_Theory.restore;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   274
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   275
          val phi = Proof_Context.export_morphism lthy lthy';
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   276
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   277
          val disc_defs = map (Morphism.thm phi) raw_disc_defs;
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   278
          val sel_defs = map (Morphism.thm phi) raw_sel_defs;
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   279
          val sel_defss = unflat_selss sel_defs;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   280
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   281
          val discs0 = map (Morphism.term phi) raw_discs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   282
          val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   283
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   284
          fun mk_disc_or_sel Ts c =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   285
            Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of c))) ~~ Ts) c;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   286
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   287
          val discs = map (mk_disc_or_sel As) discs0;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   288
          val selss = map (map (mk_disc_or_sel As)) selss0;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   289
        in
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
   290
          (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   291
        end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   292
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   293
    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   294
49458
blanchet
parents: 49438
diff changeset
   295
    val exhaust_goal =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   296
      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (v, xctr)]) in
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   297
        fold_rev Logic.all [p, v] (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
   298
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   299
49458
blanchet
parents: 49438
diff changeset
   300
    val injectss_goal =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   301
      let
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   302
        fun mk_goal _ _ [] [] = []
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   303
          | mk_goal xctr yctr xs ys =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   304
            [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
   305
              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   306
      in
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   307
        map4 mk_goal xctrs yctrs xss yss
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   308
      end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   309
49458
blanchet
parents: 49438
diff changeset
   310
    val half_distinctss_goal =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   311
      let
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   312
        fun mk_goal ((xs, xc), (xs', xc')) =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   313
          fold_rev Logic.all (xs @ xs')
49203
262ab1ac38b9 repaired constant types
blanchet
parents: 49201
diff changeset
   314
            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   315
      in
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   316
        map (map mk_goal) (mk_half_pairss (xss ~~ xctrs))
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   317
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   318
49458
blanchet
parents: 49438
diff changeset
   319
    val cases_goal =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   320
      map3 (fn xs => fn xctr => fn xf =>
49201
blanchet
parents: 49199
diff changeset
   321
        fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xss xctrs xfs;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   322
49458
blanchet
parents: 49438
diff changeset
   323
    val goalss = [exhaust_goal] :: injectss_goal @ half_distinctss_goal @ [cases_goal];
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   324
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   325
    fun after_qed thmss lthy =
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   326
      let
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   327
        val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   328
          (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   329
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   330
        val inject_thms = flat inject_thmss;
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   331
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   332
        val exhaust_thm' =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   333
          let val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As) in
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   334
            Drule.instantiate' [] [SOME (certify lthy v)]
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   335
              (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes exhaust_thm))
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   336
          end;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   337
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   338
        val exhaust_cases = map base_name_of_ctr ctrs;
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   339
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   340
        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
   341
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   342
        val (distinct_thmsss', distinct_thmsss) =
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   343
          map2 (map2 append) (Library.chop_groups n half_distinct_thmss)
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   344
            (transpose (Library.chop_groups n other_half_distinct_thmss))
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   345
          |> `transpose;
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   346
        val distinct_thms = interleave (flat half_distinct_thmss) (flat other_half_distinct_thmss);
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   347
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   348
        val nchotomy_thm =
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   349
          let
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   350
            val goal =
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   351
              HOLogic.mk_Trueprop (HOLogic.mk_all (fst v', snd v',
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   352
                Library.foldr1 HOLogic.mk_disj exist_xs_v_eq_ctrs));
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   353
          in
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   354
            Skip_Proof.prove lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   355
          end;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   356
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   357
        val (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   358
             collapse_thms, case_eq_thms) =
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   359
          if no_dests then
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   360
            ([], [], [], [], [], [], [], [])
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
   361
          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
   362
            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
   363
              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
   364
                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
   365
                    (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   366
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   367
              fun has_undefined_rhs thm =
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   368
                (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   369
                  Const (@{const_name undefined}, _) => true
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   370
                | _ => false);
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   371
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
   372
              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
   373
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   374
              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
   375
                (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
   376
                   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
   377
                 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
   378
                   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
   379
                     (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
   380
                |> filter_out has_undefined_rhs;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   381
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   382
              fun mk_unique_disc_def () =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   383
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   384
                  val m = the_single ms;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   385
                  val goal = mk_Trueprop_eq (mk_v_eq_v (), the_single exist_xs_v_eq_ctrs);
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   386
                in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   387
                  Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m exhaust_thm')
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   388
                  |> singleton (Proof_Context.export names_lthy lthy)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   389
                  |> Thm.close_derivation
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   390
                end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   391
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   392
              fun mk_alternate_disc_def k =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   393
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   394
                  val goal =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   395
                    mk_Trueprop_eq (alternate_disc_lhs (K (nth discs)) (3 - k),
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   396
                      nth exist_xs_v_eq_ctrs (k - 1));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   397
                in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   398
                  Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   399
                    mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   400
                      (nth distinct_thms (2 - k)) exhaust_thm')
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   401
                  |> singleton (Proof_Context.export names_lthy lthy)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   402
                  |> Thm.close_derivation
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   403
                end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   404
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   405
              val has_alternate_disc_def =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   406
                exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   407
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   408
              val disc_defs' =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   409
                map2 (fn k => fn def =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   410
                  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
   411
                  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
   412
                  else def) ks disc_defs;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   413
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   414
              val discD_thms = map (fn def => def RS iffD1) disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   415
              val discI_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   416
                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
   417
                  disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   418
              val not_discI_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   419
                map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   420
                    (Local_Defs.unfold lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   421
                  ms disc_defs';
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   422
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   423
              val (disc_thmss', disc_thmss) =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   424
                let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   425
                  fun mk_thm discI _ [] = refl RS discI
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   426
                    | mk_thm _ not_discI [distinct] = distinct RS not_discI;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   427
                  fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   428
                in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   429
                  map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   430
                end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   431
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   432
              val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   433
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   434
              val disc_exclude_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   435
                if has_alternate_disc_def then
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   436
                  []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   437
                else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   438
                  let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   439
                    fun mk_goal [] = []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   440
                      | mk_goal [((_, true), (_, true))] = []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   441
                      | mk_goal [(((_, disc), _), ((_, disc'), _))] =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   442
                        [Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (betapply (disc, v)),
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   443
                           HOLogic.mk_Trueprop (HOLogic.mk_not (betapply (disc', v)))))];
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   444
                    fun prove tac goal = Skip_Proof.prove lthy [] [] goal (K tac);
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   445
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   446
                    val infos = ms ~~ discD_thms ~~ discs ~~ no_discs;
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   447
                    val half_pairss = mk_half_pairss infos;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   448
49458
blanchet
parents: 49438
diff changeset
   449
                    val halvess_goal = map mk_goal half_pairss;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   450
                    val half_thmss =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   451
                      map3 (fn [] => K (K []) | [goal] => fn [((((m, discD), _), _), _)] =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   452
                          fn disc_thm => [prove (mk_half_disc_exclude_tac m discD disc_thm) goal])
49458
blanchet
parents: 49438
diff changeset
   453
                        halvess_goal half_pairss (flat disc_thmss');
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   454
49458
blanchet
parents: 49438
diff changeset
   455
                    val other_halvess_goal = map (mk_goal o map swap) half_pairss;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   456
                    val other_half_thmss =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   457
                      map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
49458
blanchet
parents: 49438
diff changeset
   458
                        other_halvess_goal;
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   459
                  in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   460
                    interleave (flat half_thmss) (flat other_half_thmss)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   461
                  end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   462
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   463
              val disc_exhaust_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   464
                if has_alternate_disc_def orelse no_discs_at_all then
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   465
                  []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   466
                else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   467
                  let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   468
                    fun mk_prem disc = mk_imp_p [HOLogic.mk_Trueprop (betapply (disc, v))];
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   469
                    val goal = fold_rev Logic.all [p, v] (mk_imp_p (map mk_prem discs));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   470
                  in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   471
                    [Skip_Proof.prove lthy [] [] goal (fn _ =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   472
                       mk_disc_exhaust_tac n exhaust_thm discI_thms)]
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   473
                  end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   474
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   475
              val collapse_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   476
                if no_dests then
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   477
                  []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   478
                else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   479
                  let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   480
                    fun mk_goal ctr disc sels =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   481
                      let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   482
                        val prem = HOLogic.mk_Trueprop (betapply (disc, v));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   483
                        val concl =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   484
                          mk_Trueprop_eq ((null sels ? swap)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   485
                            (Term.list_comb (ctr, map ap_v sels), v));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   486
                      in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   487
                        if prem aconv concl then NONE
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   488
                        else SOME (Logic.all v (Logic.mk_implies (prem, concl)))
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   489
                      end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   490
                    val goals = map3 mk_goal ctrs discs selss;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   491
                  in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   492
                    map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   493
                      Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   494
                        mk_collapse_tac ctxt m discD sel_thms)
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   495
                      |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   496
                    |> map_filter I
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   497
                  end;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   498
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   499
              val case_eq_thms =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   500
                if no_dests then
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   501
                  []
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   502
                else
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   503
                  let
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   504
                    fun mk_body f sels = Term.list_comb (f, map ap_v sels);
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   505
                    val goal =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   506
                      mk_Trueprop_eq (fcase $ v, mk_IfN B (map ap_v discs) (map2 mk_body fs selss));
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   507
                  in
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   508
                    [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   509
                      mk_case_eq_tac ctxt n exhaust_thm' case_thms disc_thmss' sel_thmss)]
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   510
                    |> Proof_Context.export names_lthy lthy
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   511
                  end;
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   512
            in
49281
3d87f4fd0d50 generate all sel theorems
blanchet
parents: 49280
diff changeset
   513
              (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   514
               collapse_thms, case_eq_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
   515
            end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   516
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   517
        val (case_cong_thm, weak_case_cong_thm) =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   518
          let
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   519
            fun mk_prem xctr xs f g =
49045
7d9631754bba minor fixes (for compatibility with existing datatype package)
blanchet
parents: 49044
diff changeset
   520
              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (w, xctr),
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   521
                mk_Trueprop_eq (f, g)));
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   522
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   523
            val v_eq_w = mk_Trueprop_eq (v, w);
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   524
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   525
            val goal =
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   526
              Logic.list_implies (v_eq_w :: map4 mk_prem xctrs xss fs gs,
49201
blanchet
parents: 49199
diff changeset
   527
                 mk_Trueprop_eq (fcase $ v, gcase $ w));
49458
blanchet
parents: 49438
diff changeset
   528
            val weak_goal = Logic.mk_implies (v_eq_w, mk_Trueprop_eq (fcase $ v, fcase $ w));
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   529
          in
49049
blanchet
parents: 49048
diff changeset
   530
            (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac exhaust_thm' case_thms),
49458
blanchet
parents: 49438
diff changeset
   531
             Skip_Proof.prove lthy [] [] weak_goal (K (etac arg_cong 1)))
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   532
            |> pairself (singleton (Proof_Context.export names_lthy lthy))
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   533
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   534
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   535
        val (split_thm, split_asm_thm) =
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   536
          let
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   537
            fun mk_conjunct xctr xs f_xs =
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   538
              list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (v, xctr), q $ f_xs));
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   539
            fun mk_disjunct xctr xs f_xs =
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   540
              list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (v, xctr),
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   541
                HOLogic.mk_not (q $ f_xs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   542
49201
blanchet
parents: 49199
diff changeset
   543
            val lhs = q $ (fcase $ v);
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   544
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   545
            val goal =
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   546
              mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
49458
blanchet
parents: 49438
diff changeset
   547
            val asm_goal =
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   548
              mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   549
                (map3 mk_disjunct xctrs xss xfs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   550
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   551
            val split_thm =
49049
blanchet
parents: 49048
diff changeset
   552
              Skip_Proof.prove lthy [] [] goal
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   553
                (fn _ => mk_split_tac exhaust_thm' case_thms inject_thmss distinct_thmsss)
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   554
              |> singleton (Proof_Context.export names_lthy lthy)
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   555
            val split_asm_thm =
49458
blanchet
parents: 49438
diff changeset
   556
              Skip_Proof.prove lthy [] [] asm_goal (fn {context = ctxt, ...} =>
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   557
                mk_split_asm_tac ctxt split_thm)
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   558
              |> singleton (Proof_Context.export names_lthy lthy)
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   559
          in
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   560
            (split_thm, split_asm_thm)
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   561
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   562
49437
c139da00fb4a register induct attributes
blanchet
parents: 49434
diff changeset
   563
        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
   564
        val cases_type_attr = Attrib.internal (K (Induct.cases_type dataT_name));
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   565
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   566
        val notes =
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   567
          [(case_congN, [case_cong_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   568
           (case_eqN, case_eq_thms, []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   569
           (casesN, case_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   570
           (collapseN, collapse_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   571
           (discsN, disc_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   572
           (disc_excludeN, disc_exclude_thms, []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   573
           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   574
           (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   575
           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   576
           (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   577
           (nchotomyN, [nchotomy_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   578
           (selsN, all_sel_thms, simp_attrs),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   579
           (splitN, [split_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   580
           (split_asmN, [split_asm_thm], []),
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   581
           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   582
          |> filter_out (null o #2)
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   583
          |> map (fn (thmN, thms, attrs) =>
49302
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   584
            ((Binding.qualify true (Binding.name_of data_b) (Binding.name thmN), attrs),
f5bd87aac224 added optional qualifiers for constructors and destructors, similarly to the old package
blanchet
parents: 49300
diff changeset
   585
             [(thms, [])]));
49300
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   586
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   587
        val notes' =
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   588
          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
c707df2e2083 added attributes to theorems
blanchet
parents: 49298
diff changeset
   589
          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   590
      in
49438
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   591
        ((selss, inject_thms, distinct_thms, case_thms, discI_thms, sel_thmss),
5bc80d96241e group "simps" together
blanchet
parents: 49437
diff changeset
   592
         lthy |> Local_Theory.notes (notes' @ notes) |> snd)
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   593
      end;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   594
  in
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   595
    (goalss, after_qed, lthy')
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   596
  end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   597
49199
7c9a3c67c55d added high-level recursor, not yet curried
blanchet
parents: 49174
diff changeset
   598
fun wrap_datatype tacss = (fn (goalss, after_qed, lthy) =>
49111
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   599
  map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   600
  |> (fn thms => after_qed thms lthy)) oo prepare_wrap_datatype (K I);
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   601
49297
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   602
val wrap_datatype_cmd = (fn (goalss, after_qed, lthy) =>
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   603
  Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   604
  prepare_wrap_datatype Syntax.read_term;
47fbf2e3e89c provide a programmatic interface for FP sugar
blanchet
parents: 49286
diff changeset
   605
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   606
fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
49111
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   607
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   608
val parse_bindings = parse_bracket_list Parse.binding;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   609
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
   610
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   611
val parse_bound_term = (Parse.binding --| @{keyword ":"}) -- Parse.term;
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   612
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
   613
val parse_bound_termss = parse_bracket_list parse_bound_terms;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   614
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   615
val parse_wrap_options =
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   616
  Scan.optional (@{keyword "("} |-- (@{keyword "no_dests"} >> K true) --| @{keyword ")"}) false;
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   617
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   618
val _ =
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
   619
  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
49278
718e4ad1517e added no_dests option
blanchet
parents: 49277
diff changeset
   620
    ((parse_wrap_options -- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) --
49280
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   621
      Parse.term -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
52413dc96326 allow default values for selectors in low-level "wrap_data" command
blanchet
parents: 49278
diff changeset
   622
        Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
49199
7c9a3c67c55d added high-level recursor, not yet curried
blanchet
parents: 49174
diff changeset
   623
     >> wrap_datatype_cmd);
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   624
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   625
end;