src/ZF/ind_syntax.ML
author wenzelm
Sun, 12 Feb 2006 21:34:25 +0100
changeset 19031 0059b5b195a2
parent 18176 ae9bd644d106
child 21539 c5cf9243ad62
permissions -rw-r--r--
export exception SAME (for join); join: use internal modify, no option type (handle SAME/DUP instead); defined: simplified copy of lookup code, avoids allocation of option constructor; added replace, which does not change equal entries;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12243
wenzelm
parents: 8819
diff changeset
     1
(*  Title:      ZF/ind_syntax.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
12243
wenzelm
parents: 8819
diff changeset
     6
Abstract Syntax functions for Inductive Definitions.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
     9
(*The structure protects these items from redeclaration (somewhat!).  The 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    10
  datatype definitions in theory files refer to these items by name!
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    11
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    12
structure Ind_Syntax =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    13
struct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
4804
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    15
(*Print tracing messages during processing of "inductive" theory sections*)
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    16
val trace = ref false;
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    17
17988
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    18
fun traceIt msg thy t = 
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    19
  if !trace then (tracing (msg ^ Sign.string_of_term thy t); t)
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    20
  else t;
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    21
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    22
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    23
(** Abstract syntax definitions for ZF **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    25
val iT = Type("i",[]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    27
val mem_const = Const("op :", [iT,iT]--->FOLogic.oT);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
fun mk_all_imp (A,P) = 
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    31
    FOLogic.all_const iT $ 
4972
7fe1d30c1374 mk_all_imp: no longer creates goals that have beta-redexes
paulson
parents: 4804
diff changeset
    32
      Abs("v", iT, FOLogic.imp $ (mem_const $ Bound 0 $ A) $ 
18176
ae9bd644d106 Term.betapply;
wenzelm
parents: 17988
diff changeset
    33
	           Term.betapply(P, Bound 0));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
val Part_const = Const("Part", [iT,iT-->iT]--->iT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    37
val apply_const = Const("op `", [iT,iT]--->iT);
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    38
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
    39
val Vrecursor_const = Const("Univ.Vrecursor", [[iT,iT]--->iT, iT]--->iT);
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    40
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    41
val Collect_const = Const("Collect", [iT, iT-->FOLogic.oT] ---> iT);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    44
(*simple error-checking in the premises of an inductive definition*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    45
fun chk_prem rec_hd (Const("op &",_) $ _ $ _) =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    46
        error"Premises may not be conjuctive"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    47
  | chk_prem rec_hd (Const("op :",_) $ t $ X) = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    48
        deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    49
  | chk_prem rec_hd t = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    50
        deny (Logic.occs(rec_hd,t)) "Recursion term in side formula";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    51
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    52
(*Return the conclusion of a rule, of the form t:X*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
fun rule_concl rl = 
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    54
    let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    55
                Logic.strip_imp_concl rl
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    56
    in  (t,X)  end;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    57
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    58
(*As above, but return error message if bad*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    59
fun rule_concl_msg sign rl = rule_concl rl
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    60
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    61
                          Sign.string_of_term sign rl);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
(*For deriving cases rules.  CollectD2 discards the domain, which is redundant;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
  read_instantiate replaces a propositional variable by a formula variable*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
val equals_CollectD = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
    read_instantiate [("W","?Q")]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
        (make_elim (equalityD1 RS subsetD RS CollectD2));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    70
(** For datatype definitions **)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    71
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    72
(*Constructor name, type, mixfix info;
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    73
  internal name from mixfix, datatype sets, full premises*)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    74
type constructor_spec = 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    75
    ((string * typ * mixfix) * string * term list * term list);
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    76
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    77
fun dest_mem (Const("op :",_) $ x $ A) = (x,A)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    78
  | dest_mem _ = error "Constructor specifications must have the form x:A";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    79
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    80
(*read a constructor specification*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    81
fun read_construct sign (id, sprems, syn) =
8819
d04923e183c7 use Sign.simple_read_term;
wenzelm
parents: 7694
diff changeset
    82
    let val prems = map (Sign.simple_read_term sign FOLogic.oT) sprems
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    83
        val args = map (#1 o dest_mem) prems
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    84
        val T = (map (#2 o dest_Free) args) ---> iT
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    85
                handle TERM _ => error 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    86
                    "Bad variable in constructor specification"
568
756b0e2a6cac replaced Lexicon.scan_id by Scanner.scan_id;
wenzelm
parents: 543
diff changeset
    87
        val name = Syntax.const_name id syn  (*handle infix constructors*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    88
    in ((id,T,syn), name, args, prems) end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    89
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    90
val read_constructs = map o map o read_construct;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    92
(*convert constructor specifications into introduction rules*)
3925
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    93
fun mk_intr_tms sg (rec_tm, constructs) =
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    94
  let
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    95
    fun mk_intr ((id,T,syn), name, args, prems) =
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    96
      Logic.list_implies
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    97
        (map FOLogic.mk_Trueprop prems,
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    98
	 FOLogic.mk_Trueprop
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    99
	    (mem_const $ list_comb (Const (Sign.full_name sg name, T), args)
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
   100
	               $ rec_tm))
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   101
  in  map mk_intr constructs  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   102
3925
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
   103
fun mk_all_intr_tms sg arg = List.concat (ListPair.map (mk_intr_tms sg) arg);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   105
fun mk_Un (t1, t2) = Const("op Un", [iT,iT]--->iT) $ t1 $ t2;
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   106
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   107
val empty       = Const("0", iT)
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
   108
and univ        = Const("Univ.univ", iT-->iT)
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
   109
and quniv       = Const("QUniv.quniv", iT-->iT);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   111
(*Make a datatype's domain: form the union of its set parameters*)
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   112
fun union_params (rec_tm, cs) =
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   113
  let val (_,args) = strip_comb rec_tm
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   114
      fun is_ind arg = (type_of arg = iT)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 13150
diff changeset
   115
  in  case List.filter is_ind (args @ cs) of
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   116
         []     => empty
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   117
       | u_args => fold_bal mk_Un u_args
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   118
  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   119
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   120
(*univ or quniv constitutes the sum domain for mutual recursion;
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   121
  it is applied to the datatype parameters and to Consts occurring in the
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   122
  definition other than Nat.nat and the datatype sets themselves.
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   123
  FIXME: could insert all constant set expressions, e.g. nat->nat.*)
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   124
fun data_domain co (rec_tms, con_ty_lists) = 
13150
0c50d13d449a better error messages for datatypes not declared Const
paulson
parents: 12243
diff changeset
   125
    let val rec_hds = map head_of rec_tms
0c50d13d449a better error messages for datatypes not declared Const
paulson
parents: 12243
diff changeset
   126
        val dummy = assert_all is_Const rec_hds
0c50d13d449a better error messages for datatypes not declared Const
paulson
parents: 12243
diff changeset
   127
          (fn t => "Datatype set not previously declared as constant: " ^
16867
cf7d61d56acf tuned fold on terms and lists;
wenzelm
parents: 15570
diff changeset
   128
                   Sign.string_of_term IFOL.thy t);
13150
0c50d13d449a better error messages for datatypes not declared Const
paulson
parents: 12243
diff changeset
   129
        val rec_names = (*nat doesn't have to be added*)
0c50d13d449a better error messages for datatypes not declared Const
paulson
parents: 12243
diff changeset
   130
	    "Nat.nat" :: map (#1 o dest_Const) rec_hds
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   131
	val u = if co then quniv else univ
16867
cf7d61d56acf tuned fold on terms and lists;
wenzelm
parents: 15570
diff changeset
   132
	val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms)
cf7d61d56acf tuned fold on terms and lists;
wenzelm
parents: 15570
diff changeset
   133
          (fn t as Const (a, _) => if a mem_string rec_names then I else insert (op =) t
cf7d61d56acf tuned fold on terms and lists;
wenzelm
parents: 15570
diff changeset
   134
            | _ => I)) con_ty_lists [];
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   135
    in  u $ union_params (hd rec_tms, cs)  end;
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   136
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
(*Could go to FOL, but it's hardly general*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   139
val def_swap_iff = prove_goal IFOL.thy "a==b ==> a=c <-> c=b"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   140
 (fn [def] => [(rewtac def), (rtac iffI 1), (REPEAT (etac sym 1))]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
val def_trans = prove_goal IFOL.thy "[| f==g;  g(a)=b |] ==> f(a)=b"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
  (fn [rew,prem] => [ rewtac rew, rtac prem 1 ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
55
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   145
(*Delete needless equality assumptions*)
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   146
val refl_thin = prove_goal IFOL.thy "!!P. [| a=a;  P |] ==> P"
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   147
     (fn _ => [assume_tac 1]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   149
(*Includes rules for succ and Pair since they are common constructions*)
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   150
val elim_rls = [asm_rl, FalseE, succ_neq_0, sym RS succ_neq_0, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   151
                Pair_neq_0, sym RS Pair_neq_0, Pair_inject,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   152
                make_elim succ_inject, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   153
                refl_thin, conjE, exE, disjE];
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   154
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   155
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   156
(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   157
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   158
  | tryres (th, []) = raise THM("tryres", 0, [th]);
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   159
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   160
fun gen_make_elim elim_rls rl = 
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   161
      standard (tryres (rl, elim_rls @ [revcut_rl]));
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   162
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   163
(*Turns iff rules into safe elimination rules*)
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   164
fun mk_free_SEs iffs = map (gen_make_elim [conjE,FalseE]) (iffs RL [iffD1]);
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   165
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   166
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   167
end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   168
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   169
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   170
(*For convenient access by the user*)
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
   171
val trace_induct = Ind_Syntax.trace;