src/ZF/Tools/induct_tacs.ML
author wenzelm
Sat, 10 Jun 2017 21:48:02 +0200
changeset 66061 880db47fed30
parent 63120 629a4c5e953e
child 69593 3dda49e08b9d
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
     1
(*  Title:      ZF/Tools/induct_tacs.ML
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
     4
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
     5
Induction and exhaustion tactics for Isabelle/ZF.  The theory
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
     6
information needed to support them (and to support primrec).  Also a
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
     7
function to install other sets as if they were datatypes.
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
     8
*)
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
     9
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    10
signature DATATYPE_TACTICS =
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    11
sig
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
    12
  val exhaust_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
    13
    int -> tactic
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
    14
  val induct_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
    15
    int -> tactic
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
    16
  val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56231
diff changeset
    17
  val rep_datatype: Facts.ref * Token.src list -> Facts.ref * Token.src list ->
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56231
diff changeset
    18
    (Facts.ref * Token.src list) list -> (Facts.ref * Token.src list) list -> theory -> theory
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    19
end;
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    20
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    21
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    22
(** Datatype information, e.g. associated theorems **)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    23
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    24
type datatype_info =
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
    25
  {inductive: bool,             (*true if inductive, not coinductive*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    26
   constructors : term list,    (*the constructors, as Consts*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    27
   rec_rewrites : thm list,     (*recursor equations*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    28
   case_rewrites : thm list,    (*case equations*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    29
   induct : thm,
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    30
   mutual_induct : thm,
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    31
   exhaustion : thm};
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    32
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32952
diff changeset
    33
structure DatatypesData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    34
(
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    35
  type T = datatype_info Symtab.table;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    36
  val empty = Symtab.empty;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    37
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32952
diff changeset
    38
  fun merge data : T = Symtab.merge (K true) data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    39
);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    40
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    41
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    42
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    43
(** Constructor information: needed to map constructors to datatypes **)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    44
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    45
type constructor_info =
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    46
  {big_rec_name : string,     (*name of the mutually recursive set*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    47
   constructors : term list,  (*the constructors, as Consts*)
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6112
diff changeset
    48
   free_iffs    : thm list,   (*freeness simprules*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    49
   rec_rewrites : thm list};  (*recursor equations*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    50
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    51
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32952
diff changeset
    52
structure ConstructorsData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    53
(
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    54
  type T = constructor_info Symtab.table
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    55
  val empty = Symtab.empty
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    56
  val extend = I
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32952
diff changeset
    57
  fun merge data = Symtab.merge (K true) data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22101
diff changeset
    58
);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
    59
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    60
structure DatatypeTactics : DATATYPE_TACTICS =
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    61
struct
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    62
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    63
fun datatype_info thy name =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17314
diff changeset
    64
  (case Symtab.lookup (DatatypesData.get thy) name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15462
diff changeset
    65
    SOME info => info
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15462
diff changeset
    66
  | NONE => error ("Unknown datatype " ^ quote name));
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    67
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    68
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    69
(*Given a variable, find the inductive set associated it in the assumptions*)
14153
76a6ba67bd15 new case_tac
paulson
parents: 12311
diff changeset
    70
exception Find_tname of string
76a6ba67bd15 new case_tac
paulson
parents: 12311
diff changeset
    71
37780
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    72
fun find_tname ctxt var As =
24826
78e6a3cea367 avoid unnamed infixes;
wenzelm
parents: 24725
diff changeset
    73
  let fun mk_pair (Const(@{const_name mem},_) $ Free (v,_) $ A) =
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    74
             (v, #1 (dest_Const (head_of A)))
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
    75
        | mk_pair _ = raise Match
37780
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    76
      val pairs = map_filter (try (mk_pair o FOLogic.dest_Trueprop)) As
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    77
      val x =
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    78
        (case try (dest_Free o Syntax.read_term ctxt) var of
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    79
          SOME (x, _) => x
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    80
        | _ => raise Find_tname ("Bad variable " ^ quote var))
7e91b3f98c46 eliminated OldGoals.strip_context;
wenzelm
parents: 36960
diff changeset
    81
  in case AList.lookup (op =) pairs x of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15462
diff changeset
    82
       NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15462
diff changeset
    83
     | SOME t => t
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    84
  end;
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    85
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
    86
(** generic exhaustion and induction tactic for datatypes
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
    87
    Differences from HOL:
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    88
      (1) no checking if the induction var occurs in premises, since it always
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    89
          appears in one of them, and it's hard to check for other occurrences
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    90
      (2) exhaustion works for VARIABLES in the premises, not general terms
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    91
**)
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    92
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
    93
fun exhaust_induct_tac exh ctxt var fixes i state = SUBGOAL (fn _ =>
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
    94
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41310
diff changeset
    95
    val thy = Proof_Context.theory_of ctxt
60695
757549b4bbe6 Variable.focus etc.: optional bindings provided by user;
wenzelm
parents: 59936
diff changeset
    96
    val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i NONE state
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58828
diff changeset
    97
    val tn = find_tname ctxt' var (map Thm.term_of asms)
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
    98
    val rule =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    99
        if exh then #exhaustion (datatype_info thy tn)
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   100
               else #induct  (datatype_info thy tn)
24826
78e6a3cea367 avoid unnamed infixes;
wenzelm
parents: 24725
diff changeset
   101
    val (Const(@{const_name mem},_) $ Var(ixn,_) $ _) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58828
diff changeset
   102
        (case Thm.prems_of rule of
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   103
             [] => error "induction is not available for this datatype"
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   104
           | major::_ => FOLogic.dest_Trueprop major)
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   105
  in
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
   106
    Rule_Insts.eres_inst_tac ctxt [((ixn, Position.none), var)] fixes rule i
14153
76a6ba67bd15 new case_tac
paulson
parents: 12311
diff changeset
   107
  end
76a6ba67bd15 new case_tac
paulson
parents: 12311
diff changeset
   108
  handle Find_tname msg =>
76a6ba67bd15 new case_tac
paulson
parents: 12311
diff changeset
   109
            if exh then (*try boolean case analysis instead*)
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
   110
                case_tac ctxt var fixes i
56231
b98813774a63 enforce subgoal boundaries via SUBGOAL/SUBGOAL_CASES -- clean tactical failure if out-of-range;
wenzelm
parents: 55740
diff changeset
   111
            else error msg) i state;
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   112
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   113
val exhaust_tac = exhaust_induct_tac true;
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   114
val induct_tac = exhaust_induct_tac false;
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   115
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   116
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   117
(**** declare non-datatype as datatype ****)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   118
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   119
fun rep_datatype_i elim induct case_eqns recursor_eqns thy =
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   120
  let
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   121
    (*analyze the LHS of a case equation to get a constructor*)
41310
65631ca437c9 proper identifiers for consts and types;
wenzelm
parents: 39557
diff changeset
   122
    fun const_of (Const(@{const_name IFOL.eq}, _) $ (_ $ c) $ _) = c
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   123
      | const_of eqn = error ("Ill-formed case equation: " ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26618
diff changeset
   124
                              Syntax.string_of_term_global thy eqn);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   125
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   126
    val constructors =
44058
ae85c5d64913 misc tuning -- eliminated old-fashioned rep_thm;
wenzelm
parents: 42361
diff changeset
   127
        map (head_of o const_of o FOLogic.dest_Trueprop o Thm.prop_of) case_eqns;
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   128
24826
78e6a3cea367 avoid unnamed infixes;
wenzelm
parents: 24725
diff changeset
   129
    val Const (@{const_name mem}, _) $ _ $ data =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58828
diff changeset
   130
        FOLogic.dest_Trueprop (hd (Thm.prems_of elim));
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   131
6112
5e4871c5136b datatype package improvements
paulson
parents: 6092
diff changeset
   132
    val Const(big_rec_name, _) = head_of data;
5e4871c5136b datatype package improvements
paulson
parents: 6092
diff changeset
   133
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   134
    val simps = case_eqns @ recursor_eqns;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   135
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   136
    val dt_info =
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   137
          {inductive = true,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   138
           constructors = constructors,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   139
           rec_rewrites = recursor_eqns,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   140
           case_rewrites = case_eqns,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   141
           induct = induct,
35409
5c5bb83f2bae more antiquotations;
wenzelm
parents: 33522
diff changeset
   142
           mutual_induct = @{thm TrueI},  (*No need for mutual induction*)
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   143
           exhaustion = elim};
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   144
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   145
    val con_info =
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   146
          {big_rec_name = big_rec_name,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   147
           constructors = constructors,
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   148
              (*let primrec handle definition by cases*)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   149
           free_iffs = [],  (*thus we expect the necessary freeness rewrites
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   150
                              to be in the simpset already, as is the case for
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   151
                              Nat and disjoint sum*)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   152
           rec_rewrites = (case recursor_eqns of
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   153
                               [] => case_eqns | _ => recursor_eqns)};
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   154
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   155
    (*associate with each constructor the datatype name and rewrites*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   156
    val con_pairs = map (fn c => (#1 (dest_Const c), con_info)) constructors
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   157
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6065
diff changeset
   158
  in
17223
430edc6b7826 curried_lookup/update;
wenzelm
parents: 17057
diff changeset
   159
    thy
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   160
    |> Sign.add_path (Long_Name.base_name big_rec_name)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38522
diff changeset
   161
    |> Global_Theory.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17314
diff changeset
   162
    |> DatatypesData.put (Symtab.update (big_rec_name, dt_info) (DatatypesData.get thy))
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17314
diff changeset
   163
    |> ConstructorsData.put (fold_rev Symtab.update con_pairs (ConstructorsData.get thy))
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 22846
diff changeset
   164
    |> Sign.parent_path
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   165
  end;
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents:
diff changeset
   166
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   167
fun rep_datatype raw_elim raw_induct raw_case_eqns raw_recursor_eqns thy =
24725
04b676d1a1fe Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   168
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41310
diff changeset
   169
    val ctxt = Proof_Context.init_global thy;
55740
11dd48f84441 more positions;
wenzelm
parents: 46961
diff changeset
   170
    val elim = Facts.the_single ("elimination", Position.none) (Attrib.eval_thms ctxt [raw_elim]);
11dd48f84441 more positions;
wenzelm
parents: 46961
diff changeset
   171
    val induct = Facts.the_single ("induction", Position.none) (Attrib.eval_thms ctxt [raw_induct]);
24725
04b676d1a1fe Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   172
    val case_eqns = Attrib.eval_thms ctxt raw_case_eqns;
04b676d1a1fe Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   173
    val recursor_eqns = Attrib.eval_thms ctxt raw_recursor_eqns;
04b676d1a1fe Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   174
  in rep_datatype_i elim induct case_eqns recursor_eqns thy end;
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12109
diff changeset
   175
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   176
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   177
(* theory setup *)
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   178
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   179
val _ =
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   180
  Theory.setup
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   181
    (Method.setup @{binding induct_tac}
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 62969
diff changeset
   182
      (Args.goal_spec -- Scan.lift (Args.embedded -- Parse.for_fixes) >>
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
   183
        (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt s xs)))
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   184
      "induct_tac emulation (dynamic instantiation!)" #>
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   185
    Method.setup @{binding case_tac}
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 62969
diff changeset
   186
     (Args.goal_spec -- Scan.lift (Args.embedded -- Parse.for_fixes) >>
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 59780
diff changeset
   187
        (fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (exhaust_tac ctxt s xs)))
58828
6d076fdd933d modernized setup;
wenzelm
parents: 58028
diff changeset
   188
      "datatype case_tac emulation (dynamic instantiation!)");
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   189
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   190
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   191
(* outer syntax *)
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   192
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24826
diff changeset
   193
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59788
diff changeset
   194
  Outer_Syntax.command @{command_keyword rep_datatype} "represent existing set inductively"
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 60695
diff changeset
   195
    ((@{keyword "elimination"} |-- Parse.!!! Parse.thm) --
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 60695
diff changeset
   196
     (@{keyword "induction"} |-- Parse.!!! Parse.thm) --
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 60695
diff changeset
   197
     (@{keyword "case_eqns"} |-- Parse.!!! Parse.thms1) --
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 60695
diff changeset
   198
     Scan.optional (@{keyword "recursor_eqns"} |-- Parse.!!! Parse.thms1) []
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   199
     >> (fn (((x, y), z), w) => Toplevel.theory (rep_datatype x y z w)));
12204
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   200
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   201
end;
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   202
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   203
val exhaust_tac = DatatypeTactics.exhaust_tac;
98115606ee42 Isar version of 'rep_datatype';
wenzelm
parents: 12175
diff changeset
   204
val induct_tac  = DatatypeTactics.induct_tac;