src/HOL/Codatatype/Tools/bnf_wrap.ML
author blanchet
Wed, 05 Sep 2012 15:40:13 +0200
changeset 49157 6407346b74c7
parent 49153 c15a7123605c
child 49174 41790d616f63
permissions -rw-r--r--
ported "Misc_Data" to new syntax
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
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
    10
  val no_binder: binding
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    11
  val mk_half_pairss: 'a list -> ('a * 'a) list list
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    12
  val wrap_data: ({prems: thm list, context: Proof.context} -> tactic) list list ->
49119
1f605c36869c more work on FP sugar
blanchet
parents: 49118
diff changeset
    13
    (term list * term) * (binding list * binding list list) -> local_theory -> local_theory
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    14
end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    15
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
    16
structure BNF_Wrap : BNF_WRAP =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    17
struct
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    18
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    19
open BNF_Util
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
    20
open BNF_Wrap_Tactics
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    21
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    22
val is_N = "is_";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    23
val un_N = "un_";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    24
fun mk_un_N 1 1 suf = un_N ^ suf
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    25
  | mk_un_N _ l suf = un_N ^ suf ^ string_of_int l;
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    26
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    27
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
    28
val case_eqN = "case_eq";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    29
val casesN = "cases";
49118
b815fa776b91 renamed theorem
blanchet
parents: 49117
diff changeset
    30
val collapseN = "collapse";
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
    31
val disc_excludeN = "disc_exclude";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    32
val disc_exhaustN = "disc_exhaust";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    33
val discsN = "discs";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    34
val distinctN = "distinct";
49075
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    35
val exhaustN = "exhaust";
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    36
val injectN = "inject";
ed769978dc8d rearrange dependencies
blanchet
parents: 49074
diff changeset
    37
val nchotomyN = "nchotomy";
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    38
val selsN = "sels";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    39
val splitN = "split";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    40
val split_asmN = "split_asm";
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    41
val weak_case_cong_thmsN = "weak_case_cong";
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    42
49139
e36ce78add78 use empty binding rather than "*" for default
blanchet
parents: 49137
diff changeset
    43
val no_binder = @{binding ""};
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
    44
val fallback_binder = @{binding _};
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    45
49056
blanchet
parents: 49055
diff changeset
    46
fun pad_list x n xs = xs @ replicate (n - length xs) x;
blanchet
parents: 49055
diff changeset
    47
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
    48
fun mk_half_pairss' _ [] = []
49056
blanchet
parents: 49055
diff changeset
    49
  | mk_half_pairss' indent (y :: ys) =
blanchet
parents: 49055
diff changeset
    50
    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
    51
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
    52
fun mk_half_pairss ys = mk_half_pairss' [[]] ys;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    53
49150
881e573a619e added TODO
blanchet
parents: 49148
diff changeset
    54
(* TODO: provide a way to have a different default value, e.g. "tl Nil = Nil" *)
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    55
fun mk_undef T Ts = Const (@{const_name undefined}, Ts ---> T);
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    56
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
    57
fun eta_expand_case_arg xs f_xs = fold_rev Term.lambda xs f_xs;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    58
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    59
fun name_of_ctr t =
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    60
  case head_of t of
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    61
    Const (s, _) => s
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    62
  | Free (s, _) => s
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    63
  | _ => error "Cannot extract name of constructor";
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
    64
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
    65
fun prepare_wrap_data prep_term ((raw_ctrs, raw_case), (raw_disc_binders, raw_sel_binderss))
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    66
  no_defs_lthy =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    67
  let
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    68
    (* TODO: sanity checks on arguments *)
49113
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    69
    (* TODO: attributes (simp, case_names, etc.) *)
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    70
    (* TODO: case syntax *)
ef3eea7ae251 tuned TODOs
blanchet
parents: 49111
diff changeset
    71
    (* TODO: integration with function package ("size") *)
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    72
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    73
    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
    74
    val case0 = prep_term no_defs_lthy raw_case;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    75
49054
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    76
    val n = length ctrs0;
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    77
    val ks = 1 upto n;
ee0a1d449f89 generate default names for selectors
blanchet
parents: 49053
diff changeset
    78
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    79
    val _ = if n > 0 then () else error "No constructors specified";
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    80
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    81
    val Type (T_name, As0) = body_type (fastype_of (hd ctrs0));
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    82
    val b = Binding.qualified_name T_name;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    83
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    84
    val (As, B) =
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    85
      no_defs_lthy
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    86
      |> mk_TFrees (length As0)
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    87
      ||> the_single o fst o mk_TFrees 1;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    88
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    89
    fun mk_ctr Ts ctr =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
    90
      let val Type (_, Ts0) = body_type (fastype_of ctr) in
49055
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    91
        Term.subst_atomic_types (Ts0 ~~ Ts) ctr
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    92
      end;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    93
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    94
    val T = Type (T_name, As);
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    95
    val ctrs = map (mk_ctr As) ctrs0;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    96
    val ctr_Tss = map (binder_types o fastype_of) ctrs;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    97
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    98
    val ms = map length ctr_Tss;
631512830082 pad the selectors' list with default names
blanchet
parents: 49054
diff changeset
    99
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   100
    val raw_disc_binders' = pad_list no_binder n raw_disc_binders;
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   101
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   102
    fun can_rely_on_disc k =
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   103
      not (Binding.eq_name (nth raw_disc_binders' (k - 1), no_binder)) orelse nth ms (k - 1) = 0;
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   104
    fun can_omit_disc_binder k m =
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   105
      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
   106
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   107
    val fallback_disc_binder = Binding.name o prefix is_N o Long_Name.base_name o name_of_ctr;
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   108
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   109
    val disc_binders =
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   110
      raw_disc_binders'
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   111
      |> map4 (fn k => fn m => fn ctr => fn disc =>
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   112
        if Binding.eq_name (disc, no_binder) then
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   113
          if can_omit_disc_binder k m then NONE else SOME (fallback_disc_binder ctr)
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   114
        else if Binding.eq_name (disc, fallback_binder) then
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   115
          SOME (fallback_disc_binder ctr)
49046
3c5eba97d93a allow default names for selectors via wildcard (_) + fix wrong index (k)
blanchet
parents: 49045
diff changeset
   116
        else
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   117
          SOME disc) ks ms ctrs0;
49056
blanchet
parents: 49055
diff changeset
   118
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   119
    val no_discs = map is_none disc_binders;
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   120
    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
   121
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   122
    fun fallback_sel_binder m l = Binding.name o mk_un_N m l o Long_Name.base_name o name_of_ctr;
49120
7f8e69fc6ac9 smarter "*" syntax -- fallback on "_" if "*" is impossible
blanchet
parents: 49119
diff changeset
   123
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   124
    val sel_binderss =
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   125
      pad_list [] n raw_sel_binderss
49056
blanchet
parents: 49055
diff changeset
   126
      |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   127
        if Binding.eq_name (sel, no_binder) orelse Binding.eq_name (sel, fallback_binder) then
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   128
          fallback_sel_binder m l ctr
49056
blanchet
parents: 49055
diff changeset
   129
        else
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   130
          sel) (1 upto m) o pad_list no_binder m) ctrs0 ms;
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   131
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   132
    fun mk_case Ts T =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   133
      let
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   134
        val (binders, body) = strip_type (fastype_of case0)
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   135
        val Type (_, Ts0) = List.last binders
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   136
      in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   137
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   138
    val caseB = mk_case As B;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   139
    val caseB_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   140
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   141
    fun mk_caseB_term eta_fs = Term.list_comb (caseB, eta_fs);
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   142
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   143
    val (((((((xss, yss), fs), gs), (v, v')), w), (p, p')), names_lthy) = no_defs_lthy |>
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   144
      mk_Freess "x" ctr_Tss
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   145
      ||>> mk_Freess "y" ctr_Tss
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   146
      ||>> mk_Frees "f" caseB_Ts
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   147
      ||>> mk_Frees "g" caseB_Ts
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   148
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "v") T
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   149
      ||>> yield_singleton (mk_Frees "w") T
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   150
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   151
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   152
    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
   153
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   154
    fun mk_v_eq_v () = HOLogic.mk_eq (v, v);
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   155
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   156
    val xctrs = map2 (curry Term.list_comb) ctrs xss;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   157
    val yctrs = map2 (curry Term.list_comb) ctrs yss;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   158
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   159
    val xfs = map2 (curry Term.list_comb) fs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   160
    val xgs = map2 (curry Term.list_comb) gs xss;
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   161
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   162
    val eta_fs = map2 eta_expand_case_arg xss xfs;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   163
    val eta_gs = map2 eta_expand_case_arg xss xgs;
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   164
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   165
    val caseB_fs = Term.list_comb (caseB, eta_fs);
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   166
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   167
    val exist_xs_v_eq_ctrs =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   168
      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
   169
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   170
    fun mk_sel_case_args k xs x T =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   171
      map2 (fn Ts => fn i => if i = k then fold_rev Term.lambda xs x else mk_undef T Ts) ctr_Tss ks;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   172
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
   173
    fun disc_free b = Free (Binding.name_of b, T --> HOLogic.boolT);
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   174
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   175
    fun disc_spec b exist_xs_v_eq_ctr = mk_Trueprop_eq (disc_free b $ v, exist_xs_v_eq_ctr);
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   176
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   177
    fun alternate_disc_lhs 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
   178
      HOLogic.mk_not
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   179
        (case nth disc_binders (k - 1) of
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   180
          NONE => nth exist_xs_v_eq_ctrs (k - 1)
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   181
        | SOME b => disc_free b $ 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
   182
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   183
    fun alternate_disc k =
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   184
      if n = 2 then Term.lambda v (alternate_disc_lhs (3 - k)) else error "Cannot use \"*\" here"
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   185
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   186
    fun sel_spec b x xs k =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   187
      let val T' = fastype_of x in
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   188
        mk_Trueprop_eq (Free (Binding.name_of b, T --> T') $ v,
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   189
          Term.list_comb (mk_case As T', mk_sel_case_args k xs x T') $ v)
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   190
      end;
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   191
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   192
    val missing_unique_disc_def = TrueI; (*arbitrary marker*)
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   193
    val missing_alternate_disc_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
   194
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49153
diff changeset
   195
    (* TODO: Allow use of same selector for several constructors *)
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49153
diff changeset
   196
    (* TODO: Allow use of same name for datatype and for constructor, e.g. "data L = L" *)
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49153
diff changeset
   197
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   198
    val (((raw_discs, raw_disc_defs), (raw_selss, raw_sel_defss)), (lthy', lthy)) =
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   199
      no_defs_lthy
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
   200
      |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_v_eq_ctr =>
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   201
        fn NONE =>
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   202
           if n = 1 then pair (Term.lambda v (mk_v_eq_v ()), missing_unique_disc_def)
49136
56a50871e25d smoothly handle one-constructor types
blanchet
parents: 49130
diff changeset
   203
           else if m = 0 then pair (Term.lambda v exist_xs_v_eq_ctr, refl)
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   204
           else pair (alternate_disc k, missing_alternate_disc_def)
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   205
         | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   206
             ((Thm.def_binding b, []), disc_spec b exist_xs_v_eq_ctr)) #>> apsnd snd)
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   207
        ks ms exist_xs_v_eq_ctrs disc_binders
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   208
      ||>> apfst split_list o fold_map3 (fn bs => fn xs => fn k => apfst split_list o
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   209
          fold_map2 (fn b => fn x => Specification.definition (SOME (b, NONE, NoSyn),
49129
b5413cb7d860 define "case" constant
blanchet
parents: 49125
diff changeset
   210
            ((Thm.def_binding b, []), sel_spec b x xs k)) #>> apsnd snd) bs xs) sel_binderss xss ks
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   211
      ||> `Local_Theory.restore;
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   212
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   213
    (*transforms defined frees into consts (and more)*)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   214
    val phi = Proof_Context.export_morphism lthy lthy';
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   215
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   216
    val disc_defs = map (Morphism.thm phi) raw_disc_defs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   217
    val sel_defss = map (map (Morphism.thm phi)) raw_sel_defss;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   218
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   219
    val discs0 = map (Morphism.term phi) raw_discs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   220
    val selss0 = map (map (Morphism.term phi)) raw_selss;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   221
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   222
    fun mk_disc_or_sel Ts t =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   223
      Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   224
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   225
    val discs = map (mk_disc_or_sel As) discs0;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   226
    val selss = map (map (mk_disc_or_sel As)) selss0;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   227
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   228
    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   229
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   230
    val goal_exhaust =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   231
      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
   232
        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
   233
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   234
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   235
    val goal_injectss =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   236
      let
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   237
        fun mk_goal _ _ [] [] = []
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   238
          | mk_goal xctr yctr xs ys =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   239
            [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
   240
              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   241
      in
49034
b77e1910af8a make parallel list indexing possible for inject theorems
blanchet
parents: 49033
diff changeset
   242
        map4 mk_goal xctrs yctrs xss yss
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   243
      end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   244
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   245
    val goal_half_distinctss =
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   246
      let
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   247
        fun mk_goal ((xs, t), (xs', t')) =
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   248
          fold_rev Logic.all (xs @ xs')
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   249
            (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (t, t'))));
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   250
      in
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   251
        map (map mk_goal) (mk_half_pairss (xss ~~ xctrs))
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   252
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   253
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   254
    val goal_cases =
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   255
      map3 (fn xs => fn xctr => fn xf =>
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   256
        fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (caseB_fs $ xctr, xf))) xss xctrs xfs;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   257
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   258
    val goalss = [goal_exhaust] :: goal_injectss @ goal_half_distinctss @ [goal_cases];
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   259
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   260
    fun after_qed thmss lthy =
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   261
      let
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   262
        val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   263
          (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   264
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   265
        val exhaust_thm' =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   266
          let val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As) in
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   267
            Drule.instantiate' [] [SOME (certify lthy v)]
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   268
              (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes exhaust_thm))
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   269
          end;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   270
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   271
        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
   272
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   273
        val (distinct_thmsss', distinct_thmsss) =
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   274
          map2 (map2 append) (Library.chop_groups n half_distinct_thmss)
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   275
            (transpose (Library.chop_groups n other_half_distinct_thmss))
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   276
          |> `transpose;
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   277
        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
   278
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   279
        val nchotomy_thm =
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   280
          let
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   281
            val goal =
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   282
              HOLogic.mk_Trueprop (HOLogic.mk_all (fst v', snd v',
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   283
                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
   284
          in
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   285
            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
   286
          end;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   287
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   288
        val sel_thmss =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   289
          let
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   290
            fun mk_thm k xs goal_case case_thm x sel_def =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   291
              let
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   292
                val T = fastype_of x;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   293
                val cTs =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   294
                  map ((fn T' => certifyT lthy (if T' = B then T else T')) o TFree)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   295
                    (rev (Term.add_tfrees goal_case []));
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   296
                val cxs = map (certify lthy) (mk_sel_case_args k xs x T);
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   297
              in
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   298
                Local_Defs.fold lthy [sel_def]
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   299
                  (Drule.instantiate' (map SOME cTs) (map SOME cxs) case_thm)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   300
              end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   301
            fun mk_thms k xs goal_case case_thm sel_defs =
49140
cb80b6404f8e fixed bug in type instantiation of case theorem
blanchet
parents: 49139
diff changeset
   302
              map2 (mk_thm k xs (strip_all_body goal_case) case_thm) xs sel_defs;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   303
          in
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   304
            map5 mk_thms ks xss goal_cases case_thms sel_defss
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   305
          end;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   306
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49153
diff changeset
   307
        fun mk_unique_disc_def () =
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   308
          let
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   309
            val m = the_single ms;
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   310
            val goal = mk_Trueprop_eq (mk_v_eq_v (), the_single exist_xs_v_eq_ctrs);
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   311
          in
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   312
            Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m exhaust_thm')
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   313
            |> singleton (Proof_Context.export names_lthy lthy)
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   314
            |> Thm.close_derivation
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   315
          end;
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   316
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   317
        fun mk_alternate_disc_def 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
   318
          let
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   319
            val goal =
49152
feb984727eec tuning (systematic 1-based indices)
blanchet
parents: 49150
diff changeset
   320
              mk_Trueprop_eq (Morphism.term phi (alternate_disc_lhs (3 - 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
   321
                nth exist_xs_v_eq_ctrs (k - 1));
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   322
          in
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   323
            Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
49148
93f281430e77 fixed "mk_alternate_disc_def_tac" in the case where the constructors are swapped compared with the common Nil/Cons case
blanchet
parents: 49140
diff changeset
   324
              mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - 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
   325
                exhaust_thm')
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   326
            |> singleton (Proof_Context.export names_lthy lthy)
49125
5fc5211cf104 implemented "mk_exhaust_tac"
blanchet
parents: 49123
diff changeset
   327
            |> Thm.close_derivation
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
   328
          end;
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   329
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   330
        val has_alternate_disc_def =
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   331
          exists (fn def => Thm.eq_thm_prop (def, missing_alternate_disc_def)) disc_defs;
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
   332
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   333
        val disc_defs' =
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   334
          map2 (fn k => fn def =>
49157
6407346b74c7 ported "Misc_Data" to new syntax
blanchet
parents: 49153
diff changeset
   335
            if Thm.eq_thm_prop (def, missing_unique_disc_def) then mk_unique_disc_def ()
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   336
            else if Thm.eq_thm_prop (def, missing_alternate_disc_def) then mk_alternate_disc_def k
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   337
            else def)
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
   338
          ks disc_defs;
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   339
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   340
        val discD_thms = map (fn def => def RS iffD1) disc_defs';
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   341
        val discI_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
   342
          map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs';
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   343
        val not_discI_thms =
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   344
          map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   345
            (Local_Defs.unfold lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   346
          ms disc_defs';
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   347
49050
9f4a7ed344b4 optimized "mk_case_disc_tac"
blanchet
parents: 49049
diff changeset
   348
        val (disc_thmss', disc_thmss) =
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   349
          let
49048
4e0f0f98be02 rationalized data structure for distinctness theorems
blanchet
parents: 49046
diff changeset
   350
            fun mk_thm discI _ [] = refl RS discI
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   351
              | mk_thm _ not_discI [distinct] = distinct RS not_discI;
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   352
            fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   353
          in
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   354
            map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   355
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   356
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
   357
        val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   358
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
   359
        val disc_exclude_thms =
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   360
          if has_alternate_disc_def then
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
            []
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
          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
   363
            let
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   364
              fun mk_goal [] = []
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   365
                | mk_goal [((_, true), (_, true))] = []
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   366
                | mk_goal [(((_, disc), _), ((_, disc'), _))] =
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   367
                  [Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (betapply (disc, v)),
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   368
                     HOLogic.mk_Trueprop (HOLogic.mk_not (betapply (disc', v)))))];
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   369
              fun prove tac goal = Skip_Proof.prove lthy [] [] goal (K tac);
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   370
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
   371
              val bundles = ms ~~ discD_thms ~~ discs ~~ no_discs;
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   372
              val half_pairss = mk_half_pairss bundles;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   373
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
   374
              val goal_halvess = map mk_goal half_pairss;
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   375
              val half_thmss =
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
                map3 (fn [] => K (K []) | [goal] => fn [((((m, discD), _), _), _)] => fn disc_thm =>
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
   377
                  [prove (mk_half_disc_exclude_tac m discD disc_thm) goal])
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
   378
                goal_halvess half_pairss (flat disc_thmss');
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   379
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
   380
              val goal_other_halvess = map (mk_goal o map swap) half_pairss;
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   381
              val other_half_thmss =
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
   382
                map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss goal_other_halvess;
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
   383
            in
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   384
              interleave (flat half_thmss) (flat other_half_thmss)
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   385
            end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   386
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
   387
        val disc_exhaust_thms =
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   388
          if has_alternate_disc_def orelse no_discs_at_all then
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
   389
            []
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   390
          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
   391
            let
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   392
              fun mk_prem disc = mk_imp_p [HOLogic.mk_Trueprop (betapply (disc, v))];
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   393
              val goal = fold_rev Logic.all [p, v] (mk_imp_p (map mk_prem 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
   394
            in
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   395
              [Skip_Proof.prove lthy [] [] goal (fn _ =>
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   396
                 mk_disc_exhaust_tac n exhaust_thm discI_thms)]
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   397
            end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   398
49118
b815fa776b91 renamed theorem
blanchet
parents: 49117
diff changeset
   399
        val collapse_thms =
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   400
          let
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   401
            fun mk_goal ctr disc sels =
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   402
              let
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   403
                val prem = HOLogic.mk_Trueprop (betapply (disc, v));
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   404
                val concl =
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   405
                  mk_Trueprop_eq ((null sels ? swap)
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   406
                    (Term.list_comb (ctr, map (fn sel => sel $ v) sels), v));
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   407
              in
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   408
                if prem aconv concl then NONE
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   409
                else SOME (Logic.all v (Logic.mk_implies (prem, concl)))
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   410
              end;
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   411
            val goals = map3 mk_goal ctrs discs selss;
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   412
          in
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   413
            map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   414
              Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
49137
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   415
                mk_collapse_tac ctxt m discD sel_thms)
5c8fefe0f103 fixed bugs in one-constructor case
blanchet
parents: 49136
diff changeset
   416
              |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   417
            |> map_filter I
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   418
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   419
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
   420
        val case_eq_thm =
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   421
          let
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   422
            fun mk_core f sels = Term.list_comb (f, map (fn sel => sel $ v) sels);
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   423
            fun mk_rhs _ [f] [sels] = mk_core f sels
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   424
              | mk_rhs (disc :: discs) (f :: fs) (sels :: selss) =
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   425
                Const (@{const_name If}, HOLogic.boolT --> B --> B --> B) $
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   426
                  betapply (disc, v) $ mk_core f sels $ mk_rhs discs fs selss;
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   427
            val goal = mk_Trueprop_eq (caseB_fs $ v, mk_rhs discs fs selss);
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   428
          in
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   429
            Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
49153
c15a7123605c made "mk_case_eq_tac" work in the case where the first constructor is an "alternate" constructor (in which case its discriminator paradoxically starts with a negation)
blanchet
parents: 49152
diff changeset
   430
              mk_case_eq_tac ctxt n exhaust_thm' case_thms disc_thmss' sel_thmss)
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   431
            |> singleton (Proof_Context.export names_lthy lthy)
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   432
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   433
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   434
        val (case_cong_thm, weak_case_cong_thm) =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   435
          let
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   436
            fun mk_prem xctr xs f g =
49045
7d9631754bba minor fixes (for compatibility with existing datatype package)
blanchet
parents: 49044
diff changeset
   437
              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (w, xctr),
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   438
                mk_Trueprop_eq (f, g)));
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   439
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   440
            val v_eq_w = mk_Trueprop_eq (v, w);
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   441
            val case_fs = mk_caseB_term eta_fs;
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   442
            val case_gs = mk_caseB_term eta_gs;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   443
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   444
            val goal =
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   445
              Logic.list_implies (v_eq_w :: map4 mk_prem xctrs xss fs gs,
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   446
                 mk_Trueprop_eq (case_fs $ v, case_gs $ w));
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   447
            val goal_weak = Logic.mk_implies (v_eq_w, mk_Trueprop_eq (case_fs $ v, case_fs $ w));
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   448
          in
49049
blanchet
parents: 49048
diff changeset
   449
            (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac exhaust_thm' case_thms),
49033
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   450
             Skip_Proof.prove lthy [] [] goal_weak (K (etac arg_cong 1)))
23ef2d429931 generate "weak_case_cong" property
blanchet
parents: 49032
diff changeset
   451
            |> pairself (singleton (Proof_Context.export names_lthy lthy))
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   452
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   453
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   454
        val (split_thm, split_asm_thm) =
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   455
          let
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   456
            fun mk_conjunct xctr xs f_xs =
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   457
              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
   458
            fun mk_disjunct xctr xs f_xs =
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   459
              list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (v, xctr),
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   460
                HOLogic.mk_not (q $ f_xs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   461
49130
3c26e17b2849 implemented "mk_case_tac" -- and got rid of "cheat_tac"
blanchet
parents: 49129
diff changeset
   462
            val lhs = q $ (mk_caseB_term eta_fs $ v);
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   463
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   464
            val goal =
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   465
              mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   466
            val goal_asm =
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   467
              mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   468
                (map3 mk_disjunct xctrs xss xfs)));
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   469
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   470
            val split_thm =
49049
blanchet
parents: 49048
diff changeset
   471
              Skip_Proof.prove lthy [] [] goal
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   472
                (fn _ => mk_split_tac exhaust_thm' case_thms inject_thmss distinct_thmsss)
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   473
              |> singleton (Proof_Context.export names_lthy lthy)
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   474
            val split_asm_thm =
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   475
              Skip_Proof.prove lthy [] [] goal_asm (fn {context = ctxt, ...} =>
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   476
                mk_split_asm_tac ctxt split_thm)
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   477
              |> singleton (Proof_Context.export names_lthy lthy)
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   478
          in
49044
c4a34ae5504d generate "split_asm" property
blanchet
parents: 49043
diff changeset
   479
            (split_thm, split_asm_thm)
49043
bd3e33ee762d generate "split" property
blanchet
parents: 49034
diff changeset
   480
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   481
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   482
        val notes =
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   483
          [(case_congN, [case_cong_thm]),
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   484
           (case_eqN, [case_eq_thm]),
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   485
           (casesN, case_thms),
49118
b815fa776b91 renamed theorem
blanchet
parents: 49117
diff changeset
   486
           (collapseN, collapse_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
   487
           (discsN, disc_thms),
49122
83515378d4d7 renamed "disc_exclus" theorem to "disc_exclude"
blanchet
parents: 49121
diff changeset
   488
           (disc_excludeN, disc_exclude_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
   489
           (disc_exhaustN, disc_exhaust_thms),
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   490
           (distinctN, distinct_thms),
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   491
           (exhaustN, [exhaust_thm]),
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   492
           (injectN, flat inject_thmss),
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   493
           (nchotomyN, [nchotomy_thm]),
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   494
           (selsN, flat sel_thmss),
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   495
           (splitN, [split_thm]),
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   496
           (split_asmN, [split_asm_thm]),
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   497
           (weak_case_cong_thmsN, [weak_case_cong_thm])]
49116
3d520eec2746 allow pseudo-definition of is_Cons in terms of is_Nil (and similarly for other two-constructor datatypes)
blanchet
parents: 49114
diff changeset
   498
          |> filter_out (null o snd)
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   499
          |> map (fn (thmN, thms) =>
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   500
            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   501
      in
49052
3510b943dd5b optimized "mk_split_tac" + use "notes"
blanchet
parents: 49050
diff changeset
   502
        lthy |> Local_Theory.notes notes |> snd
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   503
      end;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   504
  in
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   505
    (goalss, after_qed, lthy')
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   506
  end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   507
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   508
fun wrap_data tacss = (fn (goalss, after_qed, lthy) =>
49111
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   509
  map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   510
  |> (fn thms => after_qed thms lthy)) oo
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   511
  prepare_wrap_data (K I) (* FIXME? (singleton o Type_Infer_Context.infer_types) *)
49111
9d511132394e export "wrap" function
blanchet
parents: 49075
diff changeset
   512
49114
c0d77c85e0b8 allow "*" to indicate no discriminator
blanchet
parents: 49113
diff changeset
   513
val parse_bindings = Parse.$$$ "[" |-- Parse.list Parse.binding --| Parse.$$$ "]";
49057
1a7b9e4c3153 made parser a bit more flexible
blanchet
parents: 49056
diff changeset
   514
val parse_bindingss = Parse.$$$ "[" |-- Parse.list parse_bindings --| Parse.$$$ "]";
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   515
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
   516
val wrap_data_cmd = (fn (goalss, after_qed, lthy) =>
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   517
  Proof.theorem NONE after_qed (map (map (rpair [])) goalss) lthy) oo
49121
9e0acaa470ab more work on FP sugar
blanchet
parents: 49120
diff changeset
   518
  prepare_wrap_data Syntax.read_term;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   519
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   520
val _ =
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
   521
  Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
49023
5afe918dd476 changed order of arguments to "bnf_sugar"
blanchet
parents: 49022
diff changeset
   522
    (((Parse.$$$ "[" |-- Parse.list Parse.term --| Parse.$$$ "]") -- Parse.term --
49057
1a7b9e4c3153 made parser a bit more flexible
blanchet
parents: 49056
diff changeset
   523
      Scan.optional (parse_bindings -- Scan.optional parse_bindingss []) ([], []))
49074
d8af889dcbe3 renamed three BNF/(co)datatype-related commands
blanchet
parents: 49057
diff changeset
   524
     >> wrap_data_cmd);
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   525
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   526
end;