src/HOL/Codatatype/Tools/bnf_sugar.ML
author blanchet
Thu, 30 Aug 2012 16:50:03 +0200
changeset 49032 c2a7bedd57d8
parent 49031 632ee0da3c5b
child 49033 23ef2d429931
permissions -rw-r--r--
generate "case_cong" property
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Codatatype/Tools/bnf_sugar.ML
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
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     5
Sugar on top of a BNF.
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     6
*)
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     7
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     8
signature BNF_SUGAR =
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
     9
sig
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    10
end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    11
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    12
structure BNF_Sugar : BNF_SUGAR =
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    13
struct
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    14
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    15
open BNF_Util
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    16
open BNF_FP_Util
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    17
open BNF_Sugar_Tactics
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    18
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    19
val case_congN = "case_cong"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    20
val case_discsN = "case_discs"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    21
val casesN = "cases"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    22
val ctr_selsN = "ctr_sels"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    23
val disc_disjointN = "disc_disjoint"
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    24
val disc_exhaustN = "disc_exhaust"
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    25
val discsN = "discs"
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    26
val distinctN = "distinct"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    27
val selsN = "sels"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    28
val splitN = "split"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    29
val split_asmN = "split_asm"
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    30
val weak_case_cong_thmsN = "weak_case_cong"
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    31
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    32
fun mk_half_pairs [] = []
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    33
  | mk_half_pairs (x :: xs) = fold_rev (cons o pair x) xs (mk_half_pairs xs);
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    34
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    35
fun index_of_half_row _ 0 = 0
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    36
  | index_of_half_row n j = index_of_half_row n (j - 1) + n - j;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    37
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    38
fun index_of_half_cell n j k = index_of_half_row n j + k - (j + 1);
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
    39
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    40
val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    41
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    42
fun eta_expand_caseof_arg f xs = fold_rev Term.lambda xs (Term.list_comb (f, xs));
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    43
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    44
fun prepare_sugar prep_term (((raw_ctrs, raw_caseof), disc_names), sel_namess) no_defs_lthy =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    45
  let
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
    46
    (* TODO: sanity checks on arguments *)
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    47
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    48
    (* TODO: normalize types of constructors w.r.t. each other *)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    49
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    50
    val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    51
    val caseof0 = prep_term no_defs_lthy raw_caseof;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
    52
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    53
    val n = length ctrs0;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    54
    val ks = 1 upto n;
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
    55
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    56
    val (T_name, As0) = dest_Type (body_type (fastype_of (hd ctrs0)));
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    57
    val b = Binding.qualified_name T_name;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    58
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    59
    val (As, B) =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    60
      no_defs_lthy
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    61
      |> mk_TFrees (length As0)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    62
      ||> the_single o fst o mk_TFrees 1;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    63
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    64
    fun mk_undef T Ts = Const (@{const_name undefined}, Ts ---> T);
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    65
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    66
    fun mk_ctr Ts ctr =
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    67
      let val Ts0 = snd (dest_Type (body_type (fastype_of ctr))) in
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    68
        Term.subst_atomic_types (Ts0 ~~ Ts) ctr
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    69
      end;
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    70
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    71
    fun mk_caseof Ts T =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    72
      let val (binders, body) = strip_type (fastype_of caseof0) in
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    73
        Term.subst_atomic_types ((body, T) :: (snd (dest_Type (List.last binders)) ~~ Ts)) caseof0
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
    74
      end;
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
    75
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    76
    val T = Type (T_name, As);
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    77
    val ctrs = map (mk_ctr As) ctrs0;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    78
    val ctr_Tss = map (binder_types o fastype_of) ctrs;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    79
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    80
    val ms = map length ctr_Tss;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    81
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
    82
    val caseofB = mk_caseof As B;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    83
    val caseofB_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    84
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    85
    val (((((((xss, yss), fs), gs), (v, v')), w), p), names_lthy) = no_defs_lthy |>
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    86
      mk_Freess "x" ctr_Tss
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    87
      ||>> mk_Freess "y" ctr_Tss
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    88
      ||>> mk_Frees "f" caseofB_Ts
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    89
      ||>> mk_Frees "g" caseofB_Ts
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    90
      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "v") T
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    91
      ||>> yield_singleton (mk_Frees "w") T
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    92
      ||>> yield_singleton (mk_Frees "P") HOLogic.boolT;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    93
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    94
    val xctrs = map2 (curry Term.list_comb) ctrs xss;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
    95
    val yctrs = map2 (curry Term.list_comb) ctrs yss;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    96
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    97
    val eta_fs = map2 eta_expand_caseof_arg fs xss;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
    98
    val eta_gs = map2 eta_expand_caseof_arg gs xss;
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
    99
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   100
    val exist_xs_v_eq_ctrs =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   101
      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
   102
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   103
    fun mk_sel_caseof_args k xs x T =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   104
      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
   105
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   106
    fun disc_spec b exist_xs_v_eq_ctr =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   107
      mk_Trueprop_eq (Free (Binding.name_of b, T --> HOLogic.boolT) $ v, exist_xs_v_eq_ctr);
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   108
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   109
    fun sel_spec b x xs k =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   110
      let val T' = fastype_of x in
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   111
        mk_Trueprop_eq (Free (Binding.name_of b, T --> T') $ v,
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   112
          Term.list_comb (mk_caseof As T', mk_sel_caseof_args k xs x T') $ v)
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   113
      end;
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   114
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   115
    val (((raw_discs, (_, raw_disc_defs)), (raw_selss, (_, raw_sel_defss))), (lthy', lthy)) =
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   116
      no_defs_lthy
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   117
      |> apfst (apsnd split_list o split_list) o fold_map2 (fn b => fn exist_xs_v_eq_ctr =>
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   118
        Specification.definition (SOME (b, NONE, NoSyn),
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   119
          ((Thm.def_binding b, []), disc_spec b exist_xs_v_eq_ctr))) disc_names exist_xs_v_eq_ctrs
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   120
      ||>> apfst (apsnd split_list o split_list) o fold_map3 (fn bs => fn xs => fn k =>
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   121
        apfst (apsnd split_list o split_list) o fold_map2 (fn b => fn x =>
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   122
          Specification.definition (SOME (b, NONE, NoSyn),
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   123
            ((Thm.def_binding b, []), sel_spec b x xs k))) bs xs) sel_namess xss ks
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   124
      ||> `Local_Theory.restore;
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   125
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   126
    (*transforms defined frees into consts (and more)*)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   127
    val phi = Proof_Context.export_morphism lthy lthy';
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   128
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   129
    val disc_defs = map (Morphism.thm phi) raw_disc_defs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   130
    val sel_defss = map (map (Morphism.thm phi)) raw_sel_defss;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   131
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   132
    val discs0 = map (Morphism.term phi) raw_discs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   133
    val selss0 = map (map (Morphism.term phi)) raw_selss;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   134
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   135
    fun mk_disc_or_sel Ts t =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   136
      Term.subst_atomic_types (snd (dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   137
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   138
    val discs = map (mk_disc_or_sel As) discs0;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   139
    val selss = map (map (mk_disc_or_sel As)) selss0;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   140
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   141
    fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   142
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   143
    val goal_exhaust =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   144
      let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (v, xctr)]) in
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   145
        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
   146
      end;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   147
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   148
    val goal_injects =
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   149
      let
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   150
        fun mk_goal _ _ [] [] = NONE
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   151
          | mk_goal xctr yctr xs ys =
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   152
            SOME (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   153
              Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)));
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   154
      in
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   155
        map_filter I (map4 mk_goal xctrs yctrs xss yss)
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   156
      end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   157
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   158
    val goal_half_distincts =
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   159
      map (HOLogic.mk_Trueprop o HOLogic.mk_not o HOLogic.mk_eq) (mk_half_pairs xctrs);
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   160
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   161
    val goal_cases =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   162
      let
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   163
        val lhs0 = Term.list_comb (caseofB, eta_fs);
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   164
        fun mk_goal xctr xs f = mk_Trueprop_eq (lhs0 $ xctr, Term.list_comb (f, xs));
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   165
      in
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   166
        map3 mk_goal xctrs xss fs
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   167
      end;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   168
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   169
    val goals = [[goal_exhaust], goal_injects, goal_half_distincts, goal_cases];
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   170
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   171
    fun after_qed thmss lthy =
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   172
      let
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   173
        val [[exhaust_thm], inject_thms, half_distinct_thms, case_thms] = thmss;
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   174
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   175
        val exhaust_thm' =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   176
          let val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As) in
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   177
            Drule.instantiate' [] [SOME (certify lthy v)]
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   178
              (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes exhaust_thm))
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   179
          end;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   180
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   181
        val other_half_distinct_thms = map (fn thm => thm RS not_sym) half_distinct_thms;
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   182
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   183
        val nchotomy_thm =
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   184
          let
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   185
            val goal =
49022
005ce926a932 define selectors and discriminators
blanchet
parents: 49020
diff changeset
   186
              HOLogic.mk_Trueprop (HOLogic.mk_all (fst v', snd v',
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   187
                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
   188
          in
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   189
            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
   190
          end;
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   191
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   192
        val sel_thmss =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   193
          let
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   194
            fun mk_thm k xs goal_case case_thm x sel_def =
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   195
              let
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   196
                val T = fastype_of x;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   197
                val cTs =
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   198
                  map ((fn T' => certifyT lthy (if T' = B then T else T')) o TFree)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   199
                    (rev (Term.add_tfrees goal_case []));
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   200
                val cxs = map (certify lthy) (mk_sel_caseof_args k xs x T);
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   201
              in
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   202
                Local_Defs.fold lthy [sel_def]
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   203
                  (Drule.instantiate' (map SOME cTs) (map SOME cxs) case_thm)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   204
              end;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   205
            fun mk_thms k xs goal_case case_thm sel_defs =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   206
              map2 (mk_thm k xs goal_case case_thm) xs sel_defs;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   207
          in
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   208
            map5 mk_thms ks xss goal_cases case_thms sel_defss
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   209
          end;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   210
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   211
        val discD_thms = map (fn def => def RS iffD1) disc_defs;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   212
        val discI_thms =
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   213
          map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   214
        val not_disc_thms =
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   215
          map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   216
                  (Local_Defs.unfold lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   217
            ms disc_defs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   218
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   219
        val disc_thms =
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   220
          let
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   221
            fun get_distinct_thm k k' =
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   222
              if k > k' then nth half_distinct_thms (index_of_half_cell n (k' - 1) (k - 1))
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   223
              else nth other_half_distinct_thms (index_of_half_cell n (k' - 1) (k' - 1))
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   224
            fun mk_thm ((k, discI), not_disc) k' =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   225
              if k = k' then refl RS discI else get_distinct_thm k k' RS not_disc;
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   226
          in
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   227
            map_product mk_thm (ks ~~ discI_thms ~~ not_disc_thms) ks
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   228
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   229
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   230
        val disc_disjoint_thms =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   231
          let
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   232
            fun get_disc_thm k k' = nth disc_thms ((k' - 1) * n + (k - 1));
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   233
            fun mk_goal ((_, disc), (_, disc')) =
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   234
              Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (disc $ v),
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   235
                HOLogic.mk_Trueprop (HOLogic.mk_not (disc' $ v))));
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   236
            fun prove tac goal = Skip_Proof.prove lthy [] [] goal (K tac);
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   237
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   238
            val bundles = ks ~~ ms ~~ discD_thms ~~ discs;
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   239
            val half_pairs = mk_half_pairs bundles;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   240
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   241
            val goal_halves = map mk_goal half_pairs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   242
            val half_thms =
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   243
              map2 (fn ((((k, m), discD), _), (((k', _), _), _)) =>
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   244
                prove (mk_half_disc_disjoint_tac m discD (get_disc_thm k k')))
49028
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   245
              half_pairs goal_halves;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   246
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   247
            val goal_other_halves = map (mk_goal o swap) half_pairs;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   248
            val other_half_thms =
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   249
              map2 (prove o mk_other_half_disc_disjoint_tac) half_thms goal_other_halves;
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   250
          in
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   251
            half_thms @ other_half_thms
487427a02bee generate "disc_distinct" theorems
blanchet
parents: 49027
diff changeset
   252
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   253
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   254
        val disc_exhaust_thm =
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   255
          let
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   256
            fun mk_prem disc = mk_imp_p [HOLogic.mk_Trueprop (disc $ v)];
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   257
            val goal = fold Logic.all [p, v] (mk_imp_p (map mk_prem discs));
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   258
          in
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   259
            Skip_Proof.prove lthy [] [] goal (fn _ => mk_disc_exhaust_tac n exhaust_thm discI_thms)
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   260
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   261
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   262
        val ctr_sel_thms =
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   263
          let
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   264
            fun mk_goal ctr disc sels =
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   265
              Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (disc $ v),
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   266
                mk_Trueprop_eq ((null sels ? swap)
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   267
                  (Term.list_comb (ctr, map (fn sel => sel $ v) sels), v))));
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   268
            val goals = map3 mk_goal ctrs discs selss;
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   269
          in
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   270
            map4 (fn goal => fn m => fn discD => fn sel_thms =>
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   271
              Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   272
                mk_ctr_sel_tac ctxt m discD sel_thms))
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   273
              goals ms discD_thms sel_thmss
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   274
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   275
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   276
        val case_disc_thm =
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   277
          let
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   278
            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
   279
            fun mk_rhs _ [f] [sels] = mk_core f sels
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   280
              | mk_rhs (disc :: discs) (f :: fs) (sels :: selss) =
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   281
                Const (@{const_name If}, HOLogic.boolT --> B --> B --> B) $
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   282
                  (disc $ v) $ mk_core f sels $ mk_rhs discs fs selss;
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   283
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   284
            val lhs = Term.list_comb (caseofB, eta_fs) $ v;
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   285
            val rhs = mk_rhs discs fs selss;
49032
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   286
            val goal = mk_Trueprop_eq (lhs, rhs);
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   287
          in
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   288
            Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   289
              mk_case_disc_tac ctxt exhaust_thm' case_thms disc_thms sel_thmss)
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   290
            |> singleton (Proof_Context.export names_lthy lthy)
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
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
        val case_cong_thm =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   294
          let
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   295
            fun mk_prem xctr xs f g =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   296
              fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   297
                mk_Trueprop_eq (f, g)));
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   298
            fun mk_caseof_term fs v = Term.list_comb (caseofB, fs) $ v;
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   299
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   300
            val goal =
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   301
              Logic.list_implies (mk_Trueprop_eq (v, w) :: map4 mk_prem xctrs xss fs gs,
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   302
                 mk_Trueprop_eq (mk_caseof_term eta_fs v, mk_caseof_term eta_gs w));
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   303
          in
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   304
            Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   305
              mk_case_cong_tac ctxt exhaust_thm' case_thms)
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   306
            |> singleton (Proof_Context.export names_lthy lthy)
c2a7bedd57d8 generate "case_cong" property
blanchet
parents: 49031
diff changeset
   307
          end;
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   308
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   309
        val weak_case_cong_thms = TrueI;
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   310
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   311
        val split_thms = [];
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   312
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   313
        val split_asm_thms = [];
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   314
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   315
        (* case syntax *)
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   316
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   317
        fun note thmN thms =
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   318
          snd o Local_Theory.note
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   319
            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms);
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   320
      in
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   321
        lthy
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   322
        |> note case_congN [case_cong_thm]
49031
632ee0da3c5b generate "case_disc" property
blanchet
parents: 49030
diff changeset
   323
        |> note case_discsN [case_disc_thm]
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   324
        |> note casesN case_thms
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   325
        |> note ctr_selsN ctr_sel_thms
49027
fc3b9b49c92d added discriminator theorems
blanchet
parents: 49025
diff changeset
   326
        |> note discsN disc_thms
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   327
        |> note disc_disjointN disc_disjoint_thms
49029
f0ecfa9575a9 generate "disc_exhaust" property
blanchet
parents: 49028
diff changeset
   328
        |> note disc_exhaustN [disc_exhaust_thm]
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   329
        |> note distinctN (half_distinct_thms @ other_half_distinct_thms)
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   330
        |> note exhaustN [exhaust_thm]
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   331
        |> note injectN inject_thms
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   332
        |> note nchotomyN [nchotomy_thm]
49030
d0f4f113e43d generate "ctr_sels" theorems
blanchet
parents: 49029
diff changeset
   333
        |> note selsN (flat sel_thmss)
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   334
        |> note splitN split_thms
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   335
        |> note split_asmN split_asm_thms
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   336
        |> note weak_case_cong_thmsN [weak_case_cong_thms]
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   337
      end;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   338
  in
49025
7e89b0520e83 more work on sugar
blanchet
parents: 49023
diff changeset
   339
    (goals, after_qed, lthy')
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   340
  end;
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   341
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   342
val parse_binding_list = Parse.$$$ "[" |--  Parse.list Parse.binding --| Parse.$$$ "]";
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   343
49019
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   344
val bnf_sugar_cmd = (fn (goalss, after_qed, lthy) =>
fc4decdba5ce more work on BNF sugar
blanchet
parents: 49017
diff changeset
   345
  Proof.theorem NONE after_qed (map (map (rpair [])) goalss) lthy) oo
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 49019
diff changeset
   346
  prepare_sugar Syntax.read_term;
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   347
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   348
val _ =
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   349
  Outer_Syntax.local_theory_to_proof @{command_spec "bnf_sugar"} "adds sugar on top of a BNF"
49023
5afe918dd476 changed order of arguments to "bnf_sugar"
blanchet
parents: 49022
diff changeset
   350
    (((Parse.$$$ "[" |-- Parse.list Parse.term --| Parse.$$$ "]") -- Parse.term --
5afe918dd476 changed order of arguments to "bnf_sugar"
blanchet
parents: 49022
diff changeset
   351
      parse_binding_list -- (Parse.$$$ "[" |-- Parse.list parse_binding_list --| Parse.$$$ "]"))
5afe918dd476 changed order of arguments to "bnf_sugar"
blanchet
parents: 49022
diff changeset
   352
      >> bnf_sugar_cmd);
49017
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   353
66fc7fc2d49b started work on datatype sugar
blanchet
parents:
diff changeset
   354
end;