src/HOL/Tools/Datatype/datatype_abs_proofs.ML
author wenzelm
Fri, 02 Dec 2011 16:37:35 +0100
changeset 45743 857b7fcb0365
parent 45738 0430f9123e43
child 45821 c2f6c50e3d42
permissions -rw-r--r--
misc tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
     1
(*  Title:      HOL/Tools/Datatype/datatype_abs_proofs.ML
11539
0f17da240450 tuned headers;
wenzelm
parents: 11435
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     3
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
     4
Datatype package: proofs and defintions independent of concrete
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
     5
representation of datatypes  (i.e. requiring only abstract
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
     6
properties: injectivity / distinctness of constructors and induction).
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     8
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     9
signature DATATYPE_ABS_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    10
sig
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    11
  include DATATYPE_COMMON
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    12
  val prove_casedist_thms : config -> string list ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    13
    descr list -> (string * sort) list -> thm ->
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18377
diff changeset
    14
    attribute list -> theory -> thm list * theory
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    15
  val prove_primrec_thms : config -> string list ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    16
    descr list -> (string * sort) list ->
32915
a7a97960054b more appropriate abstraction over distinctness rules
haftmann
parents: 32906
diff changeset
    17
      (string -> thm list) -> thm list list -> thm list list * thm list list ->
a7a97960054b more appropriate abstraction over distinctness rules
haftmann
parents: 32906
diff changeset
    18
        thm -> theory -> (string list * thm list) * theory
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    19
  val prove_case_thms : config -> string list ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    20
    descr list -> (string * sort) list ->
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    21
      string list -> thm list -> theory -> (thm list list * string list) * theory
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    22
  val prove_split_thms : config -> string list ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    23
    descr list -> (string * sort) list ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    24
      thm list list -> thm list list -> thm list -> thm list list -> theory ->
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    25
        (thm * thm) list * theory
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    26
  val prove_nchotomys : config -> string list -> descr list ->
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    27
    (string * sort) list -> thm list -> theory -> thm list * theory
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    28
  val prove_weak_case_congs : string list -> descr list ->
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    29
    (string * sort) list -> theory -> thm list * theory
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 12910
diff changeset
    30
  val prove_case_congs : string list ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
    31
    descr list -> (string * sort) list ->
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    32
      thm list -> thm list list -> theory -> thm list * theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    33
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
    35
structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
(************************ case distinction theorems ***************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    40
fun prove_casedist_thms (config : Datatype_Aux.config)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    41
    new_type_names descr sorts induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    43
    val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    44
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32915
diff changeset
    45
    val descr' = flat descr;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    46
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
    47
    val newTs = take (length (hd descr)) recTs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    48
44058
ae85c5d64913 misc tuning -- eliminated old-fashioned rep_thm;
wenzelm
parents: 43324
diff changeset
    49
    val maxidx = Thm.maxidx_of induct;
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 7904
diff changeset
    50
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    51
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33056
diff changeset
    52
    fun prove_casedist_thm (i, (T, t)) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    53
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    54
        val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
    55
          Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    56
        val P =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    57
          Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx + 1), T), Bound 0) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    58
            Var (("P", 0), HOLogic.boolT));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    59
        val insts = take i dummyPs @ (P :: drop (i + 1) dummyPs);
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
    60
        val cert = cterm_of thy;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    61
        val insts' = map cert induct_Ps ~~ map cert insts;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    62
        val induct' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    63
          refl RS
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    64
            (nth (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i RSN (2, rev_mp));
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
    65
      in
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    66
        Skip_Proof.prove_global thy []
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    67
          (Logic.strip_imp_prems t)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    68
          (Logic.strip_imp_concl t)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    69
          (fn {prems, ...} =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    70
            EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    71
              [rtac induct' 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    72
               REPEAT (rtac TrueI 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    73
               REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    74
               REPEAT (rtac TrueI 1)])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    75
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    76
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    77
    val casedist_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    78
      map_index prove_casedist_thm (newTs ~~ Datatype_Prop.make_casedists descr sorts);
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    79
  in
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    80
    thy
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    81
    |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
    82
        (map single case_names_exhausts) casedist_thms
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
    83
  end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    86
(*************************** primrec combinators ******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    87
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    88
fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr sorts
32915
a7a97960054b more appropriate abstraction over distinctness rules
haftmann
parents: 32906
diff changeset
    89
    injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    91
    val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    92
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    93
    val big_name = space_implode "_" new_type_names;
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31902
diff changeset
    94
    val thy0 = Sign.add_path big_name thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32915
diff changeset
    96
    val descr' = flat descr;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
    97
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
45738
0430f9123e43 eliminated some legacy operations;
wenzelm
parents: 45700
diff changeset
    98
    val used = fold Term.add_tfree_namesT recTs [];
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
    99
    val newTs = take (length (hd descr)) recTs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 7904
diff changeset
   101
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   102
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   103
    val big_rec_name' = big_name ^ "_rec_set";
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20071
diff changeset
   104
    val rec_set_names' =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   105
      if length descr' = 1 then [big_rec_name']
45743
857b7fcb0365 misc tuning;
wenzelm
parents: 45738
diff changeset
   106
      else map (prefix (big_rec_name' ^ "_") o string_of_int) (1 upto (length descr'));
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28839
diff changeset
   107
    val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   108
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
   109
    val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr sorts used;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   111
    val rec_set_Ts =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   112
      map (fn (T1, T2) => (reccomb_fn_Ts @ [T1, T2]) ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   113
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   114
    val rec_fns =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   115
      map (uncurry (Datatype_Aux.mk_Free "f")) (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   116
    val rec_sets' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   117
      map (fn c => list_comb (Free c, rec_fns)) (rec_set_names' ~~ rec_set_Ts);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   118
    val rec_sets =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   119
      map (fn c => list_comb (Const c, rec_fns)) (rec_set_names ~~ rec_set_Ts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   120
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   121
    (* introduction rules for graph of primrec function *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   122
32906
ac97e8735cc2 less non-standard combinators
haftmann
parents: 32905
diff changeset
   123
    fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   124
      let
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   125
        fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   126
          let val free1 = Datatype_Aux.mk_Free "x" U j in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   127
            (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   128
              ((_, Datatype_Aux.DtRec m), (Us, _)) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   129
                let
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   130
                  val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   131
                  val i = length Us;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   132
                in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   133
                  (j + 1, k + 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   134
                    HOLogic.mk_Trueprop (HOLogic.list_all
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   135
                      (map (pair "x") Us, nth rec_sets' m $
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   136
                        Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   137
                    free1 :: t1s, free2 :: t2s)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   138
                end
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   139
            | _ => (j + 1, k, prems, free1::t1s, t2s))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   140
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   141
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   142
        val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   143
        val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   145
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   146
        (rec_intr_ts @
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   147
          [Logic.list_implies (prems, HOLogic.mk_Trueprop
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   148
            (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   149
              list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   150
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   151
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   152
    val (rec_intr_ts, _) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   153
      fold (fn ((d, T), set_name) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   154
        fold (make_rec_intr T set_name) (#3 (snd d))) (descr' ~~ recTs ~~ rec_sets') ([], 0);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   155
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   156
    val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33173
diff changeset
   157
      thy0
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33173
diff changeset
   158
      |> Sign.map_naming Name_Space.conceal
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33669
diff changeset
   159
      |> Inductive.add_inductive_global
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   160
          {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   161
            coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false}
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28839
diff changeset
   162
          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
26128
fe2d24c26e0c inductive package: simplified group handling;
wenzelm
parents: 25977
diff changeset
   163
          (map dest_Free rec_fns)
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33173
diff changeset
   164
          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33173
diff changeset
   165
      ||> Sign.restore_naming thy0
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33173
diff changeset
   166
      ||> Theory.checkpoint;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   167
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
    (* prove uniqueness and termination of primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   169
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   170
    val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
32906
ac97e8735cc2 less non-standard combinators
haftmann
parents: 32905
diff changeset
   172
    fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
      let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20071
diff changeset
   174
        val distinct_tac =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   175
          if i < length newTs then
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   176
            full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   177
          else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   178
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   179
        val inject =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   180
          map (fn r => r RS iffD1)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   181
            (if i < length newTs then nth constr_inject i else injects_of tname);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   182
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   183
        fun mk_unique_constr_tac n (cname, cargs) (tac, intr :: intrs, j) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   184
          let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   185
            val k = length (filter Datatype_Aux.is_rec_type cargs);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   186
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   187
            (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   188
              [DETERM tac,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   189
                REPEAT (etac ex1E 1), rtac ex1I 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   190
                DEPTH_SOLVE_1 (ares_tac [intr] 1),
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 12910
diff changeset
   191
                REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
                etac elim 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   193
                REPEAT_DETERM_N j distinct_tac,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20071
diff changeset
   194
                TRY (dresolve_tac inject 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   195
                REPEAT (etac conjE 1), hyp_subst_tac 1,
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 12910
diff changeset
   196
                REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   197
                TRY (hyp_subst_tac 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   198
                rtac refl 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
                REPEAT_DETERM_N (n - j - 1) distinct_tac],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
              intrs, j + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   202
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   203
        val (tac', intrs', _) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   204
          fold (mk_unique_constr_tac (length constrs)) constrs (tac, intrs, 0);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   205
      in (tac', intrs') end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   207
    val rec_unique_thms =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
      let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   209
        val rec_unique_ts =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   210
          map (fn (((set_t, T1), T2), i) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   211
            Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   212
              absfree ("y", T2) (set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   213
                (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   214
        val cert = cterm_of thy1;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   215
        val insts =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   216
          map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   217
            ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   218
        val induct' = cterm_instantiate (map cert induct_Ps ~~ map cert insts) induct;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   219
        val (tac, _) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   220
          fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   221
            (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1 THEN
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   222
                rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs));
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   223
      in
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   224
        Datatype_Aux.split_conj_thm (Skip_Proof.prove_global thy1 [] []
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   225
          (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) (K tac))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   226
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   227
35410
1ea89d2a1bd4 more antiquotations;
wenzelm
parents: 35364
diff changeset
   228
    val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   229
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   230
    (* define primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   231
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   232
    val big_reccomb_name = space_implode "_" new_type_names ^ "_rec";
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   233
    val reccomb_names =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   234
      map (Sign.full_bname thy1)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   235
        (if length descr' = 1 then [big_reccomb_name]
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   236
         else map (prefix (big_reccomb_name ^ "_") o string_of_int) (1 upto (length descr')));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   237
    val reccombs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   238
      map (fn ((name, T), T') => list_comb (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   239
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   240
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 18314
diff changeset
   241
    val (reccomb_defs, thy2) =
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 18314
diff changeset
   242
      thy1
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24699
diff changeset
   243
      |> Sign.add_consts_i (map (fn ((name, T), T') =>
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   244
            (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   245
            (reccomb_names ~~ recTs ~~ rec_result_Ts))
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38864
diff changeset
   246
      |> (Global_Theory.add_defs false o map Thm.no_attributes)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   247
          (map
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   248
            (fn ((((name, comb), set), T), T') =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   249
              (Binding.name (Long_Name.base_name name ^ "_def"),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   250
                Logic.mk_equals (comb, absfree ("x", T)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   251
                 (Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T')
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   252
                   (set $ Free ("x", T) $ Free ("y", T'))))))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   253
            (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31902
diff changeset
   254
      ||> Sign.parent_path
28361
232fcbba2e4e explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
wenzelm
parents: 28110
diff changeset
   255
      ||> Theory.checkpoint;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   256
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   257
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   258
    (* prove characteristic equations for primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   259
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   260
    val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   262
    val rec_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   263
      map (fn t =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   264
        Skip_Proof.prove_global thy2 [] [] t
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   265
          (fn _ => EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   266
            [rewrite_goals_tac reccomb_defs,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   267
             rtac @{thm the1_equality} 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   268
             resolve_tac rec_unique_thms 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   269
             resolve_tac rec_intrs 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   270
             REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   271
      (Datatype_Prop.make_primrecs new_type_names descr sorts thy2);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   272
  in
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   273
    thy2
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24699
diff changeset
   274
    |> Sign.add_path (space_implode "_" new_type_names)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38864
diff changeset
   275
    |> Global_Theory.add_thmss [((Binding.name "recs", rec_thms), [Nitpick_Simps.add])]
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24699
diff changeset
   276
    ||> Sign.parent_path
28361
232fcbba2e4e explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
wenzelm
parents: 28110
diff changeset
   277
    ||> Theory.checkpoint
32906
ac97e8735cc2 less non-standard combinators
haftmann
parents: 32905
diff changeset
   278
    |-> (fn thms => pair (reccomb_names, flat thms))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   279
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   280
8477
17231d71171a - Fixed bug in prove_casedist_thms (proof failed because of
berghofe
parents: 8436
diff changeset
   281
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   282
(***************************** case combinators *******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   283
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   284
fun prove_case_thms (config : Datatype_Aux.config)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   285
    new_type_names descr sorts reccomb_names primrec_thms thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   286
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   287
    val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   288
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31902
diff changeset
   289
    val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   290
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32915
diff changeset
   291
    val descr' = flat descr;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   292
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
45738
0430f9123e43 eliminated some legacy operations;
wenzelm
parents: 45700
diff changeset
   293
    val used = fold Term.add_tfree_namesT recTs [];
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   294
    val newTs = take (length (hd descr)) recTs;
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42375
diff changeset
   295
    val T' = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   296
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   297
    fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' sorts dt) ---> T';
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   298
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   299
    val case_dummy_fns =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   300
      map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   301
        let
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   302
          val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   303
          val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   304
        in Const (@{const_name undefined}, Ts @ Ts' ---> T') end) constrs) descr';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   305
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28839
diff changeset
   306
    val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   307
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   308
    (* define case combinators via primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   309
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   310
    val (case_defs, thy2) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   311
      fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   312
          let
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   313
            val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   314
              let
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   315
                val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   316
                val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   317
                val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   318
                val frees = take (length cargs) frees';
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   319
                val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   320
              in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   321
                (free, fold_rev (absfree o dest_Free) frees' (list_comb (free, frees)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   322
              end) (constrs ~~ (1 upto length constrs)));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   323
  
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   324
            val caseT = map (snd o dest_Free) fns1 @ [T] ---> T';
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   325
            val fns = flat (take i case_dummy_fns) @ fns2 @ flat (drop (i + 1) case_dummy_fns);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   326
            val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   327
            val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   328
            val def =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   329
              (Binding.name (Long_Name.base_name name ^ "_def"),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   330
                Logic.mk_equals (list_comb (Const (name, caseT), fns1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   331
                  list_comb (reccomb, (flat (take i case_dummy_fns)) @
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   332
                    fns2 @ (flat (drop (i + 1) case_dummy_fns)))));
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   333
            val ([def_thm], thy') =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   334
              thy
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   335
              |> Sign.declare_const_global decl |> snd
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   336
              |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   337
  
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   338
          in (defs @ [def_thm], thy') end)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   339
        (hd descr ~~ newTs ~~ case_names ~~ take (length newTs) reccomb_names) ([], thy1)
28361
232fcbba2e4e explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
wenzelm
parents: 28110
diff changeset
   340
      ||> Theory.checkpoint;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   341
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   342
    val case_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   343
      (map o map) (fn t =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   344
          Skip_Proof.prove_global thy2 [] [] t
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   345
            (fn _ =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   346
              EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   347
        (Datatype_Prop.make_cases new_type_names descr sorts thy2);
8477
17231d71171a - Fixed bug in prove_casedist_thms (proof failed because of
berghofe
parents: 8436
diff changeset
   348
  in
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   349
    thy2
33459
wenzelm
parents: 33338
diff changeset
   350
    |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31902
diff changeset
   351
    |> Sign.parent_path
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   352
    |> Datatype_Aux.store_thmss "cases" new_type_names case_thms
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   353
    |-> (fn thmss => pair (thmss, case_names))
8477
17231d71171a - Fixed bug in prove_casedist_thms (proof failed because of
berghofe
parents: 8436
diff changeset
   354
  end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   355
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   356
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   357
(******************************* case splitting *******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   358
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   359
fun prove_split_thms (config : Datatype_Aux.config)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   360
    new_type_names descr sorts constr_inject dist_rewrites casedist_thms case_thms thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   361
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   362
    val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   363
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   364
    val descr' = flat descr;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   365
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   366
    val newTs = take (length (hd descr)) recTs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   367
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   368
    fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   369
      let
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   370
        val cert = cterm_of thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   371
        val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   372
        val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   373
        val tac =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   374
          EVERY [rtac exhaustion' 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   375
            ALLGOALS (asm_simp_tac (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   376
      in
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   377
        (Skip_Proof.prove_global thy [] [] t1 (K tac),
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   378
         Skip_Proof.prove_global thy [] [] t2 (K tac))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   379
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   380
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   381
    val split_thm_pairs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   382
      map prove_split_thms
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   383
        ((Datatype_Prop.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   384
          dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   385
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   386
    val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   388
  in
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   389
    thy
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   390
    |> Datatype_Aux.store_thms "split" new_type_names split_thms
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   391
    ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   392
    |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   393
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   394
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8477
diff changeset
   395
fun prove_weak_case_congs new_type_names descr sorts thy =
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8477
diff changeset
   396
  let
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8477
diff changeset
   397
    fun prove_weak_case_cong t =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   398
     Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   399
       (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1]);
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8477
diff changeset
   400
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   401
    val weak_case_congs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   402
      map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs new_type_names descr sorts thy);
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8477
diff changeset
   403
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   404
  in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end;
8477
17231d71171a - Fixed bug in prove_casedist_thms (proof failed because of
berghofe
parents: 8436
diff changeset
   405
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   406
(************************* additional theorems for TFL ************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   407
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   408
fun prove_nchotomys (config : Datatype_Aux.config)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   409
    new_type_names descr sorts casedist_thms thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   410
  let
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   411
    val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   412
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   413
    fun prove_nchotomy (t, exhaustion) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   414
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   415
        (* For goal i, select the correct disjunct to attack, then prove it *)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   416
        fun tac i 0 = EVERY [TRY (rtac disjI1 i), hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   417
          | tac i n = rtac disjI2 i THEN tac i (n - 1);
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   418
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   419
        Skip_Proof.prove_global thy [] [] t
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   420
          (fn _ =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   421
            EVERY [rtac allI 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   422
             Datatype_Aux.exh_tac (K exhaustion) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   423
             ALLGOALS (fn i => tac i (i - 1))])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   424
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   426
    val nchotomys =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   427
      map prove_nchotomy (Datatype_Prop.make_nchotomys descr sorts ~~ casedist_thms);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   428
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   429
  in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   430
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   431
fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   432
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   433
    fun prove_case_cong ((t, nchotomy), case_rewrites) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   434
      let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   435
        val Const ("==>", _) $ tm $ _ = t;
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   436
        val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma) = tm;
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21621
diff changeset
   437
        val cert = cterm_of thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   438
        val nchotomy' = nchotomy RS spec;
29264
4ea3358fac3f use regular Term.add_vars, Term.add_frees etc.;
wenzelm
parents: 28965
diff changeset
   439
        val [v] = Term.add_vars (concl_of nchotomy') [];
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   440
        val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   441
      in
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32952
diff changeset
   442
        Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   443
          (fn {prems, ...} =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   444
            let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites)) in
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   445
              EVERY [
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   446
                simp_tac (HOL_ss addsimps [hd prems]) 1,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   447
                cut_facts_tac [nchotomy''] 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   448
                REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   449
                REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19233
diff changeset
   450
            end)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   451
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   452
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   453
    val case_congs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   454
      map prove_case_cong (Datatype_Prop.make_case_congs
9dcbf6a1829c misc tuning;
wenzelm
parents: 44241
diff changeset
   455
        new_type_names descr sorts thy ~~ nchotomys ~~ case_thms);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   456
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   457
  in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   458
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   459
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 39557
diff changeset
   460
open Datatype_Aux;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   461
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   462
end;