src/HOL/Tools/Datatype/rep_datatype.ML
author wenzelm
Sat, 22 Mar 2014 18:19:57 +0100
changeset 56254 a2dd9200854d
parent 56245 84fc7dfa3cd4
child 56895 f058120aaad4
permissions -rw-r--r--
more antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Datatype/rep_datatype.ML
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
     3
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
     4
Representation of existing types as datatypes: proofs and definitions
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
     5
independent of concrete representation of datatypes (i.e. requiring
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
     6
only abstract properties: injectivity / distinctness of constructors
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
     7
and induction).
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
     8
*)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
     9
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    10
signature REP_DATATYPE =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    11
sig
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    12
  val derive_datatype_props : Datatype_Aux.config -> string list -> Datatype_Aux.descr list ->
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    13
    thm -> thm list list -> thm list list -> theory -> string list * theory
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    14
  val rep_datatype : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    15
    term list -> theory -> Proof.state
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    16
  val rep_datatype_cmd : Datatype_Aux.config -> (string list -> Proof.context -> Proof.context) ->
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    17
    string list -> theory -> Proof.state
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    18
end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    19
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    20
structure Rep_Datatype: REP_DATATYPE =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    21
struct
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
    22
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    23
(** derived definitions and proofs **)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    24
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    25
(* case distinction theorems *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    26
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
    27
fun prove_casedist_thms (config : Datatype_Aux.config)
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
    28
    new_type_names descr induct case_names_exhausts thy =
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    29
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    30
    val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    31
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    32
    val descr' = flat descr;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    33
    val recTs = Datatype_Aux.get_rec_types descr';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    34
    val newTs = take (length (hd descr)) recTs;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    35
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    36
    val maxidx = Thm.maxidx_of induct;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    37
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    38
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    39
    fun prove_casedist_thm (i, (T, t)) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    40
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    41
        val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    42
          Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    43
        val P =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    44
          Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx + 1), T), Bound 0) $
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    45
            Var (("P", 0), HOLogic.boolT));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    46
        val insts = take i dummyPs @ (P :: drop (i + 1) dummyPs);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    47
        val cert = cterm_of thy;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    48
        val insts' = map cert induct_Ps ~~ map cert insts;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    49
        val induct' =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    50
          refl RS
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    51
            (nth (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i RSN (2, rev_mp));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    52
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
    53
        Goal.prove_sorry_global thy []
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    54
          (Logic.strip_imp_prems t)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    55
          (Logic.strip_imp_concl t)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    56
          (fn {prems, ...} =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    57
            EVERY
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    58
              [rtac induct' 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    59
               REPEAT (rtac TrueI 1),
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    60
               REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    61
               REPEAT (rtac TrueI 1)])
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    62
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    63
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    64
    val casedist_thms =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    65
      map_index prove_casedist_thm (newTs ~~ Datatype_Prop.make_casedists descr);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    66
  in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    67
    thy
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    68
    |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    69
        (map single case_names_exhausts) casedist_thms
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    70
  end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    71
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    72
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    73
(* primrec combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    74
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
    75
fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    76
    injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    77
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    78
    val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    79
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    80
    val big_name = space_implode "_" new_type_names;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    81
    val thy0 = Sign.add_path big_name thy;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    82
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    83
    val descr' = flat descr;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    84
    val recTs = Datatype_Aux.get_rec_types descr';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    85
    val used = fold Term.add_tfree_namesT recTs [];
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    86
    val newTs = take (length (hd descr)) recTs;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    87
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    88
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    89
55408
479a779b89a6 use names like 'rec_mytype' and 'case_mytype' in old datatype package as well, to avoid a mixture
blanchet
parents: 55403
diff changeset
    90
    val big_rec_name' = "rec_set_" ^ big_name;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    91
    val rec_set_names' =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    92
      if length descr' = 1 then [big_rec_name']
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    93
      else map (prefix (big_rec_name' ^ "_") o string_of_int) (1 upto length descr');
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    94
    val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    95
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    96
    val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr used;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    97
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    98
    val rec_set_Ts =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
    99
      map (fn (T1, T2) => (reccomb_fn_Ts @ [T1, T2]) ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   100
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   101
    val rec_fns =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   102
      map (uncurry (Datatype_Aux.mk_Free "f")) (reccomb_fn_Ts ~~ (1 upto length reccomb_fn_Ts));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   103
    val rec_sets' =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   104
      map (fn c => list_comb (Free c, rec_fns)) (rec_set_names' ~~ rec_set_Ts);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   105
    val rec_sets =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   106
      map (fn c => list_comb (Const c, rec_fns)) (rec_set_names ~~ rec_set_Ts);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   107
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   108
    (* introduction rules for graph of primrec function *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   109
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   110
    fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   111
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   112
        fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   113
          let val free1 = Datatype_Aux.mk_Free "x" U j in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   114
            (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   115
              ((_, Datatype_Aux.DtRec m), (Us, _)) =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   116
                let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   117
                  val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   118
                  val i = length Us;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   119
                in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   120
                  (j + 1, k + 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   121
                    HOLogic.mk_Trueprop (HOLogic.list_all
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   122
                      (map (pair "x") Us, nth rec_sets' m $
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   123
                        Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   124
                    free1 :: t1s, free2 :: t2s)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   125
                end
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   126
            | _ => (j + 1, k, prems, free1 :: t1s, t2s))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   127
          end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   128
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   129
        val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   130
        val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   131
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   132
      in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   133
        (rec_intr_ts @
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   134
          [Logic.list_implies (prems, HOLogic.mk_Trueprop
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   135
            (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   136
              list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   137
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   138
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   139
    val (rec_intr_ts, _) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   140
      fold (fn ((d, T), set_name) =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   141
        fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   142
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   143
    val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   144
      thy0
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   145
      |> Sign.map_naming Name_Space.conceal
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   146
      |> Inductive.add_inductive_global
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   147
          {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 49020
diff changeset
   148
            coind = false, no_elim = false, no_ind = true, skip_mono = true}
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   149
          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   150
          (map dest_Free rec_fns)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   151
          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   152
      ||> Sign.restore_naming thy0;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   153
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   154
    (* prove uniqueness and termination of primrec combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   155
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   156
    val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   157
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   158
    fun mk_unique_tac ctxt ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   159
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   160
        val distinct_tac =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   161
          if i < length newTs then
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   162
            full_simp_tac (put_simpset HOL_ss ctxt addsimps (nth dist_rewrites i)) 1
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   163
          else full_simp_tac (put_simpset HOL_ss ctxt addsimps (flat other_dist_rewrites)) 1;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   164
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   165
        val inject =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   166
          map (fn r => r RS iffD1)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   167
            (if i < length newTs then nth constr_inject i else injects_of tname);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   168
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   169
        fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   170
          let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   171
            val k = length (filter Datatype_Aux.is_rec_type cargs);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   172
          in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   173
            (EVERY
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   174
              [DETERM tac,
55990
41c6b99c5fb7 more antiquotations;
wenzelm
parents: 55958
diff changeset
   175
                REPEAT (etac @{thm ex1E} 1), rtac @{thm ex1I} 1,
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   176
                DEPTH_SOLVE_1 (ares_tac [intr] 1),
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   177
                REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   178
                etac elim 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   179
                REPEAT_DETERM_N j distinct_tac,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   180
                TRY (dresolve_tac inject 1),
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   181
                REPEAT (etac conjE 1), hyp_subst_tac ctxt 1,
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   182
                REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   183
                TRY (hyp_subst_tac ctxt 1),
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   184
                rtac refl 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   185
                REPEAT_DETERM_N (n - j - 1) distinct_tac],
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   186
              intrs, j + 1)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   187
          end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   188
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   189
        val (tac', intrs', _) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   190
          fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   191
      in (tac', intrs') end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   192
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   193
    val rec_unique_thms =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   194
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   195
        val rec_unique_ts =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   196
          map (fn (((set_t, T1), T2), i) =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   197
            Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   198
              absfree ("y", T2) (set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   199
                (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   200
        val cert = cterm_of thy1;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   201
        val insts =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   202
          map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   203
            ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   204
        val induct' = cterm_instantiate (map cert induct_Ps ~~ map cert insts) induct;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   205
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
   206
        Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] []
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   207
          (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   208
          (fn {context = ctxt, ...} =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   209
            #1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   210
              (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   211
                  rewrite_goals_tac ctxt [mk_meta_eq @{thm choice_eq}], rec_intrs)))))
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   212
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   213
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   214
    val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   215
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   216
    (* define primrec combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   217
55408
479a779b89a6 use names like 'rec_mytype' and 'case_mytype' in old datatype package as well, to avoid a mixture
blanchet
parents: 55403
diff changeset
   218
    val big_reccomb_name = "rec_" ^ space_implode "_" new_type_names;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   219
    val reccomb_names =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   220
      map (Sign.full_bname thy1)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   221
        (if length descr' = 1 then [big_reccomb_name]
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   222
         else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto length descr'));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   223
    val reccombs =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   224
      map (fn ((name, T), T') => Const (name, reccomb_fn_Ts @ [T] ---> T'))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   225
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   226
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   227
    val (reccomb_defs, thy2) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   228
      thy1
56239
17df7145a871 tuned signature;
wenzelm
parents: 56002
diff changeset
   229
      |> Sign.add_consts (map (fn ((name, T), T') =>
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   230
            (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   231
            (reccomb_names ~~ recTs ~~ rec_result_Ts))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   232
      |> (Global_Theory.add_defs false o map Thm.no_attributes)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   233
          (map
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   234
            (fn ((((name, comb), set), T), T') =>
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46708
diff changeset
   235
              (Binding.name (Thm.def_name (Long_Name.base_name name)),
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   236
                Logic.mk_equals (comb, fold_rev lambda rec_fns (absfree ("x", T)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   237
                 (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   238
                   (set $ Free ("x", T) $ Free ("y", T')))))))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   239
            (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   240
      ||> Sign.parent_path;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   241
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   242
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   243
    (* prove characteristic equations for primrec combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   244
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   245
    val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   246
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   247
    val rec_thms =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   248
      map (fn t =>
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
   249
        Goal.prove_sorry_global thy2 [] [] t
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   250
          (fn {context = ctxt, ...} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   251
            [rewrite_goals_tac ctxt reccomb_defs,
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   252
             rtac @{thm the1_equality} 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   253
             resolve_tac rec_unique_thms 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   254
             resolve_tac rec_intrs 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   255
             REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   256
       (Datatype_Prop.make_primrecs reccomb_names descr thy2);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   257
  in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   258
    thy2
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   259
    |> Sign.add_path (space_implode "_" new_type_names)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   260
    |> Global_Theory.note_thmss ""
55641
5b466efedd2c renamed 'recs' and 'cases' theorems 'rec' and 'case' in old datatype package, for consistency with new package
blanchet
parents: 55408
diff changeset
   261
      [((Binding.name "rec", [Nitpick_Simps.add]), [(rec_thms, [])])]
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   262
    ||> Sign.parent_path
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   263
    |-> (fn thms => pair (reccomb_names, maps #2 thms))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   264
  end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   265
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   266
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   267
(* case combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   268
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
   269
fun prove_case_thms (config : Datatype_Aux.config)
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
   270
    new_type_names descr reccomb_names primrec_thms thy =
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   271
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   272
    val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   273
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   274
    val ctxt = Proof_Context.init_global thy;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   275
    val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   276
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   277
    val descr' = flat descr;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   278
    val recTs = Datatype_Aux.get_rec_types descr';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   279
    val used = fold Term.add_tfree_namesT recTs [];
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   280
    val newTs = take (length (hd descr)) recTs;
56254
a2dd9200854d more antiquotations;
wenzelm
parents: 56245
diff changeset
   281
    val T' = TFree (singleton (Name.variant_list used) "'t", @{sort type});
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   282
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   283
    fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' dt) ---> T';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   284
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   285
    val case_dummy_fns =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   286
      map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   287
        let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   288
          val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   289
          val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   290
        in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   291
55408
479a779b89a6 use names like 'rec_mytype' and 'case_mytype' in old datatype package as well, to avoid a mixture
blanchet
parents: 55403
diff changeset
   292
    val case_names0 = map (fn s => Sign.full_bname thy1 ("case_" ^ s)) new_type_names;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   293
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   294
    (* define case combinators via primrec combinators *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   295
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   296
    fun def_case ((((i, (_, _, constrs)), T as Type (Tcon, _)), name), recname) (defs, thy) =
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   297
      if is_some (Ctr_Sugar.ctr_sugar_of ctxt Tcon) then
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   298
        (defs, thy)
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   299
      else
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   300
        let
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   301
          val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   302
            let
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   303
              val Ts = map (Datatype_Aux.typ_of_dtyp descr') cargs;
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   304
              val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   305
              val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   306
              val frees = take (length cargs) frees';
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   307
              val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   308
            in
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   309
              (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   310
            end) (constrs ~~ (1 upto length constrs)));
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   311
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   312
          val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   313
          val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   314
          val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   315
          val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   316
          val def =
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   317
            (Binding.name (Thm.def_name (Long_Name.base_name name)),
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   318
              Logic.mk_equals (Const (name, caseT),
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   319
                fold_rev lambda fns1
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   320
                  (list_comb (reccomb,
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   321
                    flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns)))));
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   322
          val ([def_thm], thy') =
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   323
            thy
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   324
            |> Sign.declare_const_global decl |> snd
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   325
            |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   326
        in (defs @ [def_thm], thy') end;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   327
55402
f33235c7a93e tuned code
blanchet
parents: 55304
diff changeset
   328
    val (case_defs, thy2) =
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   329
      fold def_case (hd descr ~~ newTs ~~ case_names0 ~~ take (length newTs) reccomb_names)
55402
f33235c7a93e tuned code
blanchet
parents: 55304
diff changeset
   330
        ([], thy1);
f33235c7a93e tuned code
blanchet
parents: 55304
diff changeset
   331
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   332
    fun prove_case t =
55402
f33235c7a93e tuned code
blanchet
parents: 55304
diff changeset
   333
      Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
f33235c7a93e tuned code
blanchet
parents: 55304
diff changeset
   334
        EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]);
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   335
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   336
    fun prove_cases (Type (Tcon, _)) ts =
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   337
      (case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   338
        SOME {case_thms, ...} => case_thms
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   339
      | NONE => map prove_case ts);
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   340
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   341
    val case_thms =
55403
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   342
      map2 prove_cases newTs (Datatype_Prop.make_cases case_names0 descr thy2);
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   343
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   344
    fun case_name_of (th :: _) =
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   345
      fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   346
677569668824 avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
blanchet
parents: 55402
diff changeset
   347
    val case_names = map case_name_of case_thms;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   348
  in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   349
    thy2
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   350
    |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   351
    |> Sign.parent_path
55641
5b466efedd2c renamed 'recs' and 'cases' theorems 'rec' and 'case' in old datatype package, for consistency with new package
blanchet
parents: 55408
diff changeset
   352
    |> Datatype_Aux.store_thmss "case" new_type_names case_thms
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   353
    |-> (fn thmss => pair (thmss, case_names))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   354
  end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   355
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   356
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   357
(* case splitting *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   358
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
   359
fun prove_split_thms (config : Datatype_Aux.config)
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   360
    new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   361
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   362
    val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   363
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   364
    val descr' = flat descr;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   365
    val recTs = Datatype_Aux.get_rec_types descr';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   366
    val newTs = take (length (hd descr)) recTs;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   367
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   368
    fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   369
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   370
        val cert = cterm_of thy;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   371
        val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   372
        val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   373
        fun tac ctxt =
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   374
          EVERY [rtac exhaustion' 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   375
            ALLGOALS (asm_simp_tac
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   376
              (put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))];
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   377
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   378
        (Goal.prove_sorry_global thy [] [] t1 (tac o #context),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   379
         Goal.prove_sorry_global thy [] [] t2 (tac o #context))
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   380
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   381
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   382
    val split_thm_pairs =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   383
      map prove_split_thms
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   384
        (Datatype_Prop.make_splits case_names descr thy ~~ constr_inject ~~
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   385
          dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   386
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   387
    val (split_thms, split_asm_thms) = split_list split_thm_pairs
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   388
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   389
  in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   390
    thy
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   391
    |> Datatype_Aux.store_thms "split" new_type_names split_thms
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   392
    ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   393
    |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   394
  end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   395
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   396
fun prove_weak_case_congs new_type_names case_names descr thy =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   397
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   398
    fun prove_weak_case_cong t =
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
   399
     Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   400
       (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   401
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   402
    val weak_case_congs =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   403
      map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs case_names descr thy);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   404
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   405
  in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   406
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   407
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   408
(* additional theorems for TFL *)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   409
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45907
diff changeset
   410
fun prove_nchotomys (config : Datatype_Aux.config) new_type_names descr casedist_thms thy =
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   411
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   412
    val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   413
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   414
    fun prove_nchotomy (t, exhaustion) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   415
      let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   416
        (* For goal i, select the correct disjunct to attack, then prove it *)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   417
        fun tac ctxt i 0 =
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   418
              EVERY [TRY (rtac disjI1 i), hyp_subst_tac ctxt i, REPEAT (rtac exI i), rtac refl i]
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   419
          | tac ctxt i n = rtac disjI2 i THEN tac ctxt i (n - 1);
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   420
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
   421
        Goal.prove_sorry_global thy [] [] t
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   422
          (fn {context = ctxt, ...} =>
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   423
            EVERY [rtac allI 1,
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   424
             Datatype_Aux.exh_tac (K exhaustion) 1,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
   425
             ALLGOALS (fn i => tac ctxt i (i - 1))])
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   426
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   427
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   428
    val nchotomys =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   429
      map prove_nchotomy (Datatype_Prop.make_nchotomys descr ~~ casedist_thms);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   430
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   431
  in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   432
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   433
fun prove_case_congs new_type_names case_names descr nchotomys case_thms thy =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   434
  let
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   435
    fun prove_case_cong ((t, nchotomy), case_rewrites) =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   436
      let
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 56239
diff changeset
   437
        val Const (@{const_name Pure.imp}, _) $ tm $ _ = t;
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   438
        val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma) = tm;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   439
        val cert = cterm_of thy;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   440
        val nchotomy' = nchotomy RS spec;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   441
        val [v] = Term.add_vars (concl_of nchotomy') [];
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   442
        val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy';
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   443
      in
51551
88d1d19fb74f tuned signature and module arrangement;
wenzelm
parents: 49170
diff changeset
   444
        Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   445
          (fn {context = ctxt, prems, ...} =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   446
            let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   447
              val simplify = asm_simp_tac (put_simpset HOL_ss ctxt addsimps (prems @ case_rewrites))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   448
            in
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   449
              EVERY [
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51673
diff changeset
   450
                simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1,
46708
b138dee7bed3 prefer cut_tac, where it is clear that the special variants cut_rules_tac or cut_facts_tac are not required;
wenzelm
parents: 45909
diff changeset
   451
                cut_tac nchotomy'' 1,
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   452
                REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   453
                REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   454
            end)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   455
      end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   456
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   457
    val case_congs =
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   458
      map prove_case_cong
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   459
        (Datatype_Prop.make_case_congs case_names descr thy ~~ nchotomys ~~ case_thms);
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   460
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   461
  in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   462
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   463
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   464
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   465
(** derive datatype props **)
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   466
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   467
local
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   468
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   469
fun make_dt_info descr induct inducts rec_names rec_rewrites
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   470
    (index, (((((((((((_, (tname, _, _))), inject), distinct),
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   471
      exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   472
        (split, split_asm))) =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   473
  (tname,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   474
   {index = index,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   475
    descr = descr,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   476
    inject = inject,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   477
    distinct = distinct,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   478
    induct = induct,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   479
    inducts = inducts,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   480
    exhaust = exhaust,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   481
    nchotomy = nchotomy,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   482
    rec_names = rec_names,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   483
    rec_rewrites = rec_rewrites,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   484
    case_name = case_name,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   485
    case_rewrites = case_rewrites,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   486
    case_cong = case_cong,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   487
    weak_case_cong = weak_case_cong,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   488
    split = split,
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   489
    split_asm = split_asm});
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   490
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   491
in
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   492
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 51798
diff changeset
   493
fun derive_datatype_props config dt_names descr induct inject distinct thy2 =
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   494
  let
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   495
    val flat_descr = flat descr;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   496
    val new_type_names = map Long_Name.base_name dt_names;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   497
    val _ =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   498
      Datatype_Aux.message config
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   499
        ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   500
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   501
    val (exhaust, thy3) = thy2
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   502
      |> prove_casedist_thms config new_type_names descr induct
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   503
        (Datatype_Data.mk_case_names_exhausts flat_descr dt_names);
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   504
    val (nchotomys, thy4) = thy3
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   505
      |> prove_nchotomys config new_type_names descr exhaust;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   506
    val ((rec_names, rec_rewrites), thy5) = thy4
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   507
      |> prove_primrec_thms config new_type_names descr
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   508
        (#inject o the o Symtab.lookup (Datatype_Data.get_all thy4)) inject
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   509
        (distinct, Datatype_Data.all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr)) induct;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   510
    val ((case_rewrites, case_names), thy6) = thy5
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   511
      |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   512
    val (case_congs, thy7) = thy6
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   513
      |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   514
    val (weak_case_congs, thy8) = thy7
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   515
      |> prove_weak_case_congs new_type_names case_names descr;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   516
    val (splits, thy9) = thy8
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   517
      |> prove_split_thms config new_type_names case_names descr
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   518
        inject distinct exhaust case_rewrites;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   519
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   520
    val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   521
    val dt_infos =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   522
      map_index
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   523
        (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   524
        (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   525
          case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   526
    val dt_names = map fst dt_infos;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   527
    val prfx = Binding.qualify true (space_implode "_" new_type_names);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   528
    val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   529
    val named_rules = flat (map_index (fn (i, tname) =>
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   530
      [((Binding.empty, [Induct.induct_type tname]), [([nth inducts i], [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   531
       ((Binding.empty, [Induct.cases_type tname]), [([nth exhaust i], [])])]) dt_names);
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   532
    val unnamed_rules = map (fn induct =>
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   533
      ((Binding.empty, [Rule_Cases.inner_rule, Induct.induct_type ""]), [([induct], [])]))
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   534
        (drop (length dt_names) inducts);
51673
4dfa00e264d8 separate data used for case translation from the datatype package
traytel
parents: 51672
diff changeset
   535
4dfa00e264d8 separate data used for case translation from the datatype package
traytel
parents: 51672
diff changeset
   536
    val ctxt = Proof_Context.init_global thy9;
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55641
diff changeset
   537
    val case_combs =
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55990
diff changeset
   538
      map (Proof_Context.read_const {proper = true, strict = true} ctxt) case_names;
51673
4dfa00e264d8 separate data used for case translation from the datatype package
traytel
parents: 51672
diff changeset
   539
    val constrss = map (fn (dtname, {descr, index, ...}) =>
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55990
diff changeset
   540
      map (Proof_Context.read_const {proper = true, strict = true} ctxt o fst)
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55641
diff changeset
   541
        (#3 (the (AList.lookup op = descr index)))) dt_infos;
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   542
  in
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   543
    thy9
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   544
    |> Global_Theory.note_thmss ""
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   545
      ([((prfx (Binding.name "simps"), []), [(simps, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   546
        ((prfx (Binding.name "inducts"), []), [(inducts, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   547
        ((prfx (Binding.name "splits"), []), [(maps (fn (x, y) => [x, y]) splits, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   548
        ((Binding.empty, [Simplifier.simp_add]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   549
          [(flat case_rewrites @ flat distinct @ rec_rewrites, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   550
        ((Binding.empty, [Code.add_default_eqn_attribute]), [(rec_rewrites, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   551
        ((Binding.empty, [iff_add]), [(flat inject, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   552
        ((Binding.empty, [Classical.safe_elim NONE]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   553
          [(map (fn th => th RS notE) (flat distinct), [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   554
        ((Binding.empty, [Simplifier.cong_add]), [(weak_case_congs, [])]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   555
        ((Binding.empty, [Induct.induct_simp_add]), [(flat (distinct @ inject), [])])] @
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   556
          named_rules @ unnamed_rules)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   557
    |> snd
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   558
    |> Datatype_Data.register dt_infos
51673
4dfa00e264d8 separate data used for case translation from the datatype package
traytel
parents: 51672
diff changeset
   559
    |> Context.theory_map (fold2 Case_Translation.register case_combs constrss)
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   560
    |> Datatype_Data.interpretation_data (config, dt_names)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   561
    |> pair dt_names
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   562
  end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   563
45907
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   564
end;
4b41967bd77e clarified modules that contribute to datatype package;
wenzelm
parents: 45901
diff changeset
   565
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   566
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   567
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   568
(** declare existing type as datatype **)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   569
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   570
local
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   571
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   572
fun prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct thy1 =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   573
  let
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   574
    val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   575
    val new_type_names = map Long_Name.base_name dt_names;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   576
    val prfx = Binding.qualify true (space_implode "_" new_type_names);
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   577
    val (((inject, distinct), [(_, [induct])]), thy2) =
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   578
      thy1
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   579
      |> Datatype_Aux.store_thmss "inject" new_type_names raw_inject
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   580
      ||>> Datatype_Aux.store_thmss "distinct" new_type_names raw_distinct
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   581
      ||>> Global_Theory.note_thmss ""
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   582
        [((prfx (Binding.name "induct"), [Datatype_Data.mk_case_names_induct descr]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45891
diff changeset
   583
          [([raw_induct], [])])];
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   584
  in
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   585
    thy2
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   586
    |> derive_datatype_props config dt_names [descr] induct inject distinct
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   587
 end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   588
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   589
fun gen_rep_datatype prep_term config after_qed raw_ts thy =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   590
  let
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   591
    val ctxt = Proof_Context.init_global thy;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   592
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   593
    fun constr_of_term (Const (c, T)) = (c, T)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   594
      | constr_of_term t = error ("Not a constant: " ^ Syntax.string_of_term ctxt t);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   595
    fun no_constr (c, T) =
55304
55ac31bc08a4 more formal markup;
wenzelm
parents: 54742
diff changeset
   596
      error ("Bad constructor: " ^ Proof_Context.markup_const ctxt c ^ "::" ^
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   597
        Syntax.string_of_typ ctxt T);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   598
    fun type_of_constr (cT as (_, T)) =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   599
      let
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   600
        val frees = Term.add_tfreesT T [];
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   601
        val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type T))
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   602
          handle TYPE _ => no_constr cT
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   603
        val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   604
        val _ = if length frees <> length vs then no_constr cT else ();
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   605
      in (tyco, (vs, cT)) end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   606
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   607
    val raw_cs =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   608
      AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   609
    val _ =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   610
      (case map_filter (fn (tyco, _) =>
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   611
          if Symtab.defined (Datatype_Data.get_all thy) tyco then SOME tyco else NONE) raw_cs of
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   612
        [] => ()
49020
f379cf5d71bd more work on BNF sugar -- up to derivation of nchotomy
blanchet
parents: 46961
diff changeset
   613
      | tycos => error ("Type(s) " ^ commas_quote tycos ^ " already represented inductively"));
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   614
    val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   615
    val ms =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   616
      (case distinct (op =) (map length raw_vss) of
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   617
         [n] => 0 upto n - 1
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   618
      | _ => error "Different types in given constructors");
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   619
    fun inter_sort m =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   620
      map (fn xs => nth xs m) raw_vss
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   621
      |> foldr1 (Sorts.inter_sort (Sign.classes_of thy));
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   622
    val sorts = map inter_sort ms;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   623
    val vs = Name.invent_names Name.context Name.aT sorts;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   624
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   625
    fun norm_constr (raw_vs, (c, T)) =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   626
      (c, map_atyps
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   627
        (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   628
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   629
    val cs = map (apsnd (map norm_constr)) raw_cs;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   630
    val dtyps_of_typ = map (Datatype_Aux.dtyp_of_typ (map (rpair vs o fst) cs)) o binder_types;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   631
    val dt_names = map fst cs;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   632
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   633
    fun mk_spec (i, (tyco, constr)) =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   634
      (i, (tyco, map Datatype_Aux.DtTFree vs, (map o apsnd) dtyps_of_typ constr));
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   635
    val descr = map_index mk_spec cs;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   636
    val injs = Datatype_Prop.make_injs [descr];
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   637
    val half_distincts = Datatype_Prop.make_distincts [descr];
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   638
    val ind = Datatype_Prop.make_ind [descr];
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   639
    val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   640
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   641
    fun after_qed' raw_thms =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   642
      let
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   643
        val [[[raw_induct]], raw_inject, half_distinct] =
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   644
          unflat rules (map Drule.zero_var_indexes_list raw_thms);
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   645
            (*FIXME somehow dubious*)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   646
      in
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   647
        Proof_Context.background_theory_result  (* FIXME !? *)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   648
          (prove_rep_datatype config dt_names descr raw_inject half_distinct raw_induct)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   649
        #-> after_qed
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   650
      end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   651
  in
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   652
    ctxt
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   653
    |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   654
  end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   655
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   656
in
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   657
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   658
val rep_datatype = gen_rep_datatype Sign.cert_term;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   659
val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   660
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   661
end;
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   662
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   663
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   664
(* outer syntax *)
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   665
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   666
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46909
diff changeset
   667
  Outer_Syntax.command @{command_spec "rep_datatype"} "represent existing types inductively"
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   668
    (Scan.repeat1 Parse.term >> (fn ts =>
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   669
      Toplevel.print o
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   670
      Toplevel.theory_to_proof (rep_datatype_cmd Datatype_Aux.default_config (K I) ts)));
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   671
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents:
diff changeset
   672
end;