src/ZF/Tools/inductive_package.ML
author haftmann
Tue, 06 Dec 2005 09:04:09 +0100
changeset 18358 0a733e11021a
parent 17985 d5d576b72371
child 18377 0e1d025d57b3
permissions -rw-r--r--
re-oriented some result tuples in PureThy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
     1
(*  Title:      ZF/Tools/inductive_package.ML
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     2
    ID:         $Id$
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     5
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     6
Fixedpoint definition module -- for Inductive/Coinductive Definitions
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     7
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     8
The functor will be instantiated for normal sums/products (inductive defs)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     9
                         and non-standard sums/products (coinductive defs)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    10
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    11
Sums are used only for mutual recursion;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    12
Products are used only to derive "streamlined" induction rules for relations
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    13
*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    14
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    15
type inductive_result =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    16
   {defs       : thm list,             (*definitions made in thy*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    17
    bnd_mono   : thm,                  (*monotonicity for the lfp definition*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    18
    dom_subset : thm,                  (*inclusion of recursive set in dom*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    19
    intrs      : thm list,             (*introduction rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    20
    elim       : thm,                  (*case analysis theorem*)
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6093
diff changeset
    21
    mk_cases   : string -> thm,        (*generates case theorems*)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    22
    induct     : thm,                  (*main induction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    23
    mutual_induct : thm};              (*mutual induction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    24
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    25
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    26
(*Functor's result signature*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    27
signature INDUCTIVE_PACKAGE =
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    28
sig
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    29
  (*Insert definitions for the recursive sets, which
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    30
     must *already* be declared as constants in parent theory!*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    31
  val add_inductive_i: bool -> term list * term ->
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    32
    ((bstring * term) * theory attribute list) list ->
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    33
    thm list * thm list * thm list * thm list -> theory -> theory * inductive_result
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    34
  val add_inductive: string list * string ->
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    35
    ((bstring * string) * Attrib.src list) list ->
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    36
    (thmref * Attrib.src list) list * (thmref * Attrib.src list) list *
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    37
    (thmref * Attrib.src list) list * (thmref * Attrib.src list) list ->
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    38
    theory -> theory * inductive_result
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    39
end;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    40
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    41
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    42
(*Declares functions to add fixedpoint/constructor defs to a theory.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    43
  Recursive sets must *already* be declared as constants.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    44
functor Add_inductive_def_Fun
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    45
    (structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    46
 : INDUCTIVE_PACKAGE =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    47
struct
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    48
16855
7563d0eb3414 no open Logic;
wenzelm
parents: 16457
diff changeset
    49
open Ind_Syntax;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    50
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
    51
val co_prefix = if coind then "co" else "";
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
    52
7695
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    53
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    54
(* utils *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    55
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    56
(*make distinct individual variables a1, a2, a3, ..., an. *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    57
fun mk_frees a [] = []
12902
a23dc0b7566f Symbol.bump_string;
wenzelm
parents: 12876
diff changeset
    58
  | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts;
7695
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    59
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    60
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    61
(* add_inductive(_i) *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    62
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    63
(*internal version, accepting terms*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    64
fun add_inductive_i verbose (rec_tms, dom_sum)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    65
  intr_specs (monos, con_defs, type_intrs, type_elims) thy =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    66
let
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    67
  val _ = Theory.requires thy "Inductive" "(co)inductive definitions";
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    68
  val sign = sign_of thy;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    69
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
    70
  val (intr_names, intr_tms) = split_list (map fst intr_specs);
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
    71
  val case_names = RuleCases.case_names intr_names;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    72
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    73
  (*recT and rec_params should agree for all mutually recursive components*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    74
  val rec_hds = map head_of rec_tms;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    75
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    76
  val dummy = assert_all is_Const rec_hds
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    77
          (fn t => "Recursive set not previously declared as constant: " ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    78
                   Sign.string_of_term sign t);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    79
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    80
  (*Now we know they are all Consts, so get their names, type and params*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    81
  val rec_names = map (#1 o dest_Const) rec_hds
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    82
  and (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    83
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    84
  val rec_base_names = map Sign.base_name rec_names;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    85
  val dummy = assert_all Syntax.is_identifier rec_base_names
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    86
    (fn a => "Base name of recursive set not an identifier: " ^ a);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    87
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    88
  local (*Checking the introduction rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    89
    val intr_sets = map (#2 o rule_concl_msg sign) intr_tms;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    90
    fun intr_ok set =
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    91
        case head_of set of Const(a,recT) => a mem rec_names | _ => false;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    92
  in
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    93
    val dummy =  assert_all intr_ok intr_sets
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    94
       (fn t => "Conclusion of rule does not name a recursive set: " ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    95
                Sign.string_of_term sign t);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    96
  end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    97
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    98
  val dummy = assert_all is_Free rec_params
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    99
      (fn t => "Param in recursion term not a free variable: " ^
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   100
               Sign.string_of_term sign t);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   101
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   102
  (*** Construct the fixedpoint definition ***)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   103
  val mk_variant = variant (foldr add_term_names [] intr_tms);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   104
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   105
  val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   106
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   107
  fun dest_tprop (Const("Trueprop",_) $ P) = P
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   108
    | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   109
                            Sign.string_of_term sign Q);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   110
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   111
  (*Makes a disjunct from an introduction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   112
  fun fp_part intr = (*quantify over rule's free vars except parameters*)
16855
7563d0eb3414 no open Logic;
wenzelm
parents: 16457
diff changeset
   113
    let val prems = map dest_tprop (Logic.strip_imp_prems intr)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   114
        val dummy = List.app (fn rec_hd => List.app (chk_prem rec_hd) prems) rec_hds
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   115
        val exfrees = term_frees intr \\ rec_params
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   116
        val zeq = FOLogic.mk_eq (Free(z',iT), #1 (rule_concl intr))
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   117
    in foldr FOLogic.mk_exists
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   118
             (fold_bal FOLogic.mk_conj (zeq::prems)) exfrees
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   119
    end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   120
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   121
  (*The Part(A,h) terms -- compose injections to make h*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   122
  fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   123
    | mk_Part h         = Part_const $ Free(X',iT) $ Abs(w',iT,h);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   124
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   125
  (*Access to balanced disjoint sums via injections*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   126
  val parts =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   127
      map mk_Part (accesses_bal (fn t => Su.inl $ t, fn t => Su.inr $ t, Bound 0)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   128
                                (length rec_tms));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   129
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   130
  (*replace each set by the corresponding Part(A,h)*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   131
  val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   132
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   133
  val fp_abs = absfree(X', iT,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   134
                   mk_Collect(z', dom_sum,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   135
                              fold_bal FOLogic.mk_disj part_intrs));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   136
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   137
  val fp_rhs = Fp.oper $ dom_sum $ fp_abs
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   138
16855
7563d0eb3414 no open Logic;
wenzelm
parents: 16457
diff changeset
   139
  val dummy = List.app (fn rec_hd => deny (Logic.occs (rec_hd, fp_rhs))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   140
                             "Illegal occurrence of recursion operator")
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   141
           rec_hds;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   142
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   143
  (*** Make the new theory ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   144
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   145
  (*A key definition:
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   146
    If no mutual recursion then it equals the one recursive set.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   147
    If mutual recursion then it differs from all the recursive sets. *)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   148
  val big_rec_base_name = space_implode "_" rec_base_names;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   149
  val big_rec_name = Sign.intern_const sign big_rec_base_name;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   150
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   151
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   152
  val dummy = conditional verbose (fn () =>
12243
wenzelm
parents: 12227
diff changeset
   153
    writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   154
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   155
  (*Forbid the inductive definition structure from clashing with a theory
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   156
    name.  This restriction may become obsolete as ML is de-emphasized.*)
16457
e0f22edf38a5 Context.names_of;
wenzelm
parents: 15705
diff changeset
   157
  val dummy = deny (big_rec_base_name mem (Context.names_of sign))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   158
               ("Definition " ^ big_rec_base_name ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   159
                " would clash with the theory of the same name!");
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   160
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   161
  (*Big_rec... is the union of the mutually recursive sets*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   162
  val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   163
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   164
  (*The individual sets must already be declared*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   165
  val axpairs = map Logic.mk_defpair
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   166
        ((big_rec_tm, fp_rhs) ::
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   167
         (case parts of
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   168
             [_] => []                        (*no mutual recursion*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   169
           | _ => rec_tms ~~          (*define the sets as Parts*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   170
                  map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   171
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   172
  (*tracing: print the fixedpoint definition*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   173
  val dummy = if !Ind_Syntax.trace then
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   174
              List.app (writeln o Sign.string_of_term sign o #2) axpairs
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   175
          else ()
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   176
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   177
  (*add definitions of the inductive sets*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   178
  val thy1 = thy |> Theory.add_path big_rec_base_name
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17985
diff changeset
   179
                 |> (snd o PureThy.add_defs_i false (map Thm.no_attributes axpairs))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   180
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   181
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   182
  (*fetch fp definitions from the theory*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   183
  val big_rec_def::part_rec_defs =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   184
    map (get_def thy1)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   185
        (case rec_names of [_] => rec_names
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   186
                         | _   => big_rec_base_name::rec_names);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   187
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   188
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   189
  val sign1 = sign_of thy1;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   190
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   191
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   192
  val dummy = writeln "  Proving monotonicity...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   193
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   194
  val bnd_mono =
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   195
    standard (Goal.prove sign1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   196
      (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   197
        [rtac (Collect_subset RS bnd_monoI) 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   198
         REPEAT (ares_tac (basic_monos @ monos) 1)]));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   199
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   200
  val dom_subset = standard (big_rec_def RS Fp.subs);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   201
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   202
  val unfold = standard ([big_rec_def, bnd_mono] MRS Fp.Tarski);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   203
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   204
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   205
  val dummy = writeln "  Proving the introduction rules...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   206
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   207
  (*Mutual recursion?  Helps to derive subset rules for the
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   208
    individual sets.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   209
  val Part_trans =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   210
      case rec_names of
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   211
           [_] => asm_rl
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   212
         | _   => standard (Part_subset RS subset_trans);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   213
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   214
  (*To type-check recursive occurrences of the inductive sets, possibly
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   215
    enclosed in some monotonic operator M.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   216
  val rec_typechecks =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   217
     [dom_subset] RL (asm_rl :: ([Part_trans] RL monos))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   218
     RL [subsetD];
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   219
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   220
  (*Type-checking is hardest aspect of proof;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   221
    disjIn selects the correct disjunct after unfolding*)
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   222
  fun intro_tacsf disjIn =
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   223
    [DETERM (stac unfold 1),
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   224
     REPEAT (resolve_tac [Part_eqI,CollectI] 1),
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   225
     (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   226
     rtac disjIn 2,
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   227
     (*Not ares_tac, since refl must be tried before equality assumptions;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   228
       backtracking may occur if the premises have extra variables!*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   229
     DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 2 APPEND assume_tac 2),
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   230
     (*Now solve the equations like Tcons(a,f) = Inl(?b4)*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   231
     rewrite_goals_tac con_defs,
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   232
     REPEAT (rtac refl 2),
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   233
     (*Typechecking; this can fail*)
6172
8a505e0694d0 standard spelling: type-checking
paulson
parents: 6141
diff changeset
   234
     if !Ind_Syntax.trace then print_tac "The type-checking subgoal:"
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   235
     else all_tac,
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   236
     REPEAT (FIRSTGOAL (        dresolve_tac rec_typechecks
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   237
                        ORELSE' eresolve_tac (asm_rl::PartE::SigmaE2::
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   238
                                              type_elims)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   239
                        ORELSE' hyp_subst_tac)),
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   240
     if !Ind_Syntax.trace then print_tac "The subgoal after monos, type_elims:"
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   241
     else all_tac,
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   242
     DEPTH_SOLVE (swap_res_tac (SigmaI::subsetI::type_intrs) 1)];
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   243
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   244
  (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   245
  val mk_disj_rls =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   246
      let fun f rl = rl RS disjI1
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   247
          and g rl = rl RS disjI2
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   248
      in  accesses_bal(f, g, asm_rl)  end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   249
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   250
  val intrs =
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   251
    (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms)))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   252
    |> ListPair.map (fn (t, tacs) =>
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   253
      standard (Goal.prove sign1 [] [] t
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   254
        (fn _ => EVERY (rewrite_goals_tac part_rec_defs :: tacs))))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   255
    handle MetaSimplifier.SIMPLIFIER (msg, thm) => (print_thm thm; error msg);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   256
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   257
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   258
  val dummy = writeln "  Proving the elimination rule...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   259
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   260
  (*Breaks down logical connectives in the monotonic function*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   261
  val basic_elim_tac =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   262
      REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   263
                ORELSE' bound_hyp_subst_tac))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   264
      THEN prune_params_tac
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   265
          (*Mutual recursion: collapse references to Part(D,h)*)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   266
      THEN fold_tac part_rec_defs;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   267
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   268
  (*Elimination*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   269
  val elim = rule_by_tactic basic_elim_tac
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   270
                 (unfold RS Ind_Syntax.equals_CollectD)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   271
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   272
  (*Applies freeness of the given constructors, which *must* be unfolded by
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   273
      the given defs.  Cannot simply use the local con_defs because
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   274
      con_defs=[] for inference systems.
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   275
    Proposition A should have the form t:Si where Si is an inductive set*)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   276
  fun make_cases ss A =
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   277
    rule_by_tactic
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   278
      (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac ss) THEN basic_elim_tac)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   279
      (Thm.assume A RS elim)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   280
      |> Drule.standard';
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   281
  fun mk_cases a = make_cases (*delayed evaluation of body!*)
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   282
    (simpset ()) (read_cterm (Thm.sign_of_thm elim) (a, propT));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   283
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   284
  fun induction_rules raw_induct thy =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   285
   let
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   286
     val dummy = writeln "  Proving the induction rule...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   287
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   288
     (*** Prove the main induction rule ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   289
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   290
     val pred_name = "P";            (*name for predicate variables*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   291
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   292
     (*Used to make induction rules;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   293
        ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   294
        prem is a premise of an intr rule*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   295
     fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   296
                      (Const("op :",_)$t$X), iprems) =
17314
04e21a27c0ad introduces some modern-style AList operations
haftmann
parents: 17057
diff changeset
   297
          (case AList.lookup (op aconv) ind_alist X of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15461
diff changeset
   298
               SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15461
diff changeset
   299
             | NONE => (*possibly membership in M(rec_tm), for M monotone*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   300
                 let fun mk_sb (rec_tm,pred) =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   301
                             (rec_tm, Ind_Syntax.Collect_const$rec_tm$pred)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   302
                 in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   303
       | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   304
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   305
     (*Make a premise of the induction rule.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   306
     fun induct_prem ind_alist intr =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   307
       let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   308
           val iprems = foldr (add_induct_prem ind_alist) []
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   309
                              (Logic.strip_imp_prems intr)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   310
           val (t,X) = Ind_Syntax.rule_concl intr
17314
04e21a27c0ad introduces some modern-style AList operations
haftmann
parents: 17057
diff changeset
   311
           val (SOME pred) = AList.lookup (op aconv) ind_alist X
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   312
           val concl = FOLogic.mk_Trueprop (pred $ t)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   313
       in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   314
       handle Bind => error"Recursion term not found in conclusion";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   315
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   316
     (*Minimizes backtracking by delivering the correct premise to each goal.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   317
       Intro rules with extra Vars in premises still cause some backtracking *)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   318
     fun ind_tac [] 0 = all_tac
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   319
       | ind_tac(prem::prems) i =
13747
bf308fcfd08e Better treatment of equality in premises of inductive definitions. Less
paulson
parents: 13627
diff changeset
   320
             DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN ind_tac prems (i-1);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   321
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   322
     val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   323
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   324
     val ind_prems = map (induct_prem (map (rpair pred) rec_tms))
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   325
                         intr_tms;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   326
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   327
     val dummy = if !Ind_Syntax.trace then
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   328
                 (writeln "ind_prems = ";
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   329
                  List.app (writeln o Sign.string_of_term sign1) ind_prems;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   330
                  writeln "raw_induct = "; print_thm raw_induct)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   331
             else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   332
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   333
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   334
     (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules.
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   335
       If the premises get simplified, then the proofs could fail.*)
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17314
diff changeset
   336
     val min_ss = Simplifier.theory_context thy empty_ss
12725
7ede865e1fe5 renamed forall_elim_vars_safe to gen_all;
wenzelm
parents: 12720
diff changeset
   337
           setmksimps (map mk_eq o ZF_atomize o gen_all)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   338
           setSolver (mk_solver "minimal"
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   339
                      (fn prems => resolve_tac (triv_rls@prems)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   340
                                   ORELSE' assume_tac
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   341
                                   ORELSE' etac FalseE));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   342
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   343
     val quant_induct =
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   344
       standard (Goal.prove sign1 [] ind_prems
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   345
         (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred)))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   346
         (fn prems => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   347
           [rewrite_goals_tac part_rec_defs,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   348
            rtac (impI RS allI) 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   349
            DETERM (etac raw_induct 1),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   350
            (*Push Part inside Collect*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   351
            full_simp_tac (min_ss addsimps [Part_Collect]) 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   352
            (*This CollectE and disjE separates out the introduction rules*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   353
            REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE])),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   354
            (*Now break down the individual cases.  No disjE here in case
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   355
              some premise involves disjunction.*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   356
            REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE]
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   357
                               ORELSE' bound_hyp_subst_tac)),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   358
            ind_tac (rev (map (rewrite_rule part_rec_defs) prems)) (length prems)]));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   359
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   360
     val dummy = if !Ind_Syntax.trace then
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   361
                 (writeln "quant_induct = "; print_thm quant_induct)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   362
             else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   363
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   364
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   365
     (*** Prove the simultaneous induction rule ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   366
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   367
     (*Make distinct predicates for each inductive set*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   368
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   369
     (*The components of the element type, several if it is a product*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   370
     val elem_type = CP.pseudo_type dom_sum;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   371
     val elem_factors = CP.factors elem_type;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   372
     val elem_frees = mk_frees "za" elem_factors;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   373
     val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   374
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   375
     (*Given a recursive set and its domain, return the "fsplit" predicate
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   376
       and a conclusion for the simultaneous induction rule.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   377
       NOTE.  This will not work for mutually recursive predicates.  Previously
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   378
       a summand 'domt' was also an argument, but this required the domain of
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   379
       mutual recursion to invariably be a disjoint sum.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   380
     fun mk_predpair rec_tm =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   381
       let val rec_name = (#1 o dest_Const o head_of) rec_tm
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   382
           val pfree = Free(pred_name ^ "_" ^ Sign.base_name rec_name,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   383
                            elem_factors ---> FOLogic.oT)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   384
           val qconcl =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   385
             foldr FOLogic.mk_all
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   386
               (FOLogic.imp $
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   387
                (Ind_Syntax.mem_const $ elem_tuple $ rec_tm)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   388
                      $ (list_comb (pfree, elem_frees))) elem_frees
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   389
       in  (CP.ap_split elem_type FOLogic.oT pfree,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   390
            qconcl)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   391
       end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   392
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   393
     val (preds,qconcls) = split_list (map mk_predpair rec_tms);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   394
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   395
     (*Used to form simultaneous induction lemma*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   396
     fun mk_rec_imp (rec_tm,pred) =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   397
         FOLogic.imp $ (Ind_Syntax.mem_const $ Bound 0 $ rec_tm) $
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   398
                          (pred $ Bound 0);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   399
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   400
     (*To instantiate the main induction rule*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   401
     val induct_concl =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   402
         FOLogic.mk_Trueprop
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   403
           (Ind_Syntax.mk_all_imp
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   404
            (big_rec_tm,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   405
             Abs("z", Ind_Syntax.iT,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   406
                 fold_bal FOLogic.mk_conj
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   407
                 (ListPair.map mk_rec_imp (rec_tms, preds)))))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   408
     and mutual_induct_concl =
7695
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
   409
      FOLogic.mk_Trueprop(fold_bal FOLogic.mk_conj qconcls);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   410
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   411
     val dummy = if !Ind_Syntax.trace then
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   412
                 (writeln ("induct_concl = " ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   413
                           Sign.string_of_term sign1 induct_concl);
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   414
                  writeln ("mutual_induct_concl = " ^
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   415
                           Sign.string_of_term sign1 mutual_induct_concl))
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   416
             else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   417
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   418
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   419
     val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp],
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   420
                             resolve_tac [allI, impI, conjI, Part_eqI],
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   421
                             dresolve_tac [spec, mp, Pr.fsplitD]];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   422
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   423
     val need_mutual = length rec_names > 1;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   424
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   425
     val lemma = (*makes the link between the two induction rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   426
       if need_mutual then
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   427
          (writeln "  Proving the mutual induction rule...";
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   428
           standard (Goal.prove sign1 [] []
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   429
             (Logic.mk_implies (induct_concl, mutual_induct_concl))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   430
             (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   431
               [rewrite_goals_tac part_rec_defs,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   432
                REPEAT (rewrite_goals_tac [Pr.split_eq] THEN lemma_tac 1)])))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   433
       else (writeln "  [ No mutual induction rule needed ]"; TrueI);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   434
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   435
     val dummy = if !Ind_Syntax.trace then
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   436
                 (writeln "lemma = "; print_thm lemma)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   437
             else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   438
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   439
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   440
     (*Mutual induction follows by freeness of Inl/Inr.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   441
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   442
     (*Simplification largely reduces the mutual induction rule to the
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   443
       standard rule*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   444
     val mut_ss =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   445
         min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   446
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   447
     val all_defs = con_defs @ part_rec_defs;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   448
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   449
     (*Removes Collects caused by M-operators in the intro rules.  It is very
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   450
       hard to simplify
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   451
         list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))})
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   452
       where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   453
       Instead the following rules extract the relevant conjunct.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   454
     *)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   455
     val cmonos = [subset_refl RS Collect_mono] RL monos
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   456
                   RLN (2,[rev_subsetD]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   457
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   458
     (*Minimizes backtracking by delivering the correct premise to each goal*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   459
     fun mutual_ind_tac [] 0 = all_tac
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   460
       | mutual_ind_tac(prem::prems) i =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   461
           DETERM
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   462
            (SELECT_GOAL
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   463
               (
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   464
                (*Simplify the assumptions and goal by unfolding Part and
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   465
                  using freeness of the Sum constructors; proves all but one
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   466
                  conjunct by contradiction*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   467
                rewrite_goals_tac all_defs  THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   468
                simp_tac (mut_ss addsimps [Part_iff]) 1  THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   469
                IF_UNSOLVED (*simp_tac may have finished it off!*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   470
                  ((*simplify assumptions*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   471
                   (*some risk of excessive simplification here -- might have
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   472
                     to identify the bare minimum set of rewrites*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   473
                   full_simp_tac
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   474
                      (mut_ss addsimps conj_simps @ imp_simps @ quant_simps) 1
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   475
                   THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   476
                   (*unpackage and use "prem" in the corresponding place*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   477
                   REPEAT (rtac impI 1)  THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   478
                   rtac (rewrite_rule all_defs prem) 1  THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   479
                   (*prem must not be REPEATed below: could loop!*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   480
                   DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE'
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   481
                                           eresolve_tac (conjE::mp::cmonos))))
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   482
               ) i)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   483
            THEN mutual_ind_tac prems (i-1);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   484
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   485
     val mutual_induct_fsplit =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   486
       if need_mutual then
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   487
         standard (Goal.prove sign1 [] (map (induct_prem (rec_tms~~preds)) intr_tms)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   488
           mutual_induct_concl
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   489
           (fn prems => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   490
             [rtac (quant_induct RS lemma) 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   491
              mutual_ind_tac (rev prems) (length prems)]))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   492
       else TrueI;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   493
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   494
     (** Uncurrying the predicate in the ordinary induction rule **)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   495
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   496
     (*instantiate the variable to a tuple, if it is non-trivial, in order to
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   497
       allow the predicate to be "opened up".
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   498
       The name "x.1" comes from the "RS spec" !*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   499
     val inst =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   500
         case elem_frees of [_] => I
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   501
            | _ => instantiate ([], [(cterm_of sign1 (Var(("x",1), Ind_Syntax.iT)),
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   502
                                      cterm_of sign1 elem_tuple)]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   503
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   504
     (*strip quantifier and the implication*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   505
     val induct0 = inst (quant_induct RS spec RSN (2,rev_mp));
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   506
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   507
     val Const ("Trueprop", _) $ (pred_var $ _) = concl_of induct0
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   508
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   509
     val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   510
                  |> standard
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   511
     and mutual_induct = CP.remove_split mutual_induct_fsplit
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   512
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
   513
     val (thy', [induct', mutual_induct']) = thy |> PureThy.add_thms
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   514
      [((co_prefix ^ "induct", induct), [case_names, InductAttrib.induct_set_global big_rec_name]),
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
   515
       (("mutual_induct", mutual_induct), [case_names])];
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   516
    in ((thy', induct'), mutual_induct')
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   517
    end;  (*of induction_rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   518
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   519
  val raw_induct = standard ([big_rec_def, bnd_mono] MRS Fp.induct)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   520
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   521
  val ((thy2, induct), mutual_induct) =
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   522
    if not coind then induction_rules raw_induct thy1
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   523
    else (thy1 |> PureThy.add_thms [((co_prefix ^ "induct", raw_induct), [])] |> apsnd hd, TrueI)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   524
  and defs = big_rec_def :: part_rec_defs
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   525
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   526
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   527
  val (thy3, ([bnd_mono', dom_subset', elim'], [defs', intrs'])) =
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   528
    thy2
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
   529
    |> IndCases.declare big_rec_name make_cases
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   530
    |> PureThy.add_thms
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   531
      [(("bnd_mono", bnd_mono), []),
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   532
       (("dom_subset", dom_subset), []),
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
   533
       (("cases", elim), [case_names, InductAttrib.cases_set_global big_rec_name])]
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   534
    |>>> (PureThy.add_thmss o map Thm.no_attributes)
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   535
        [("defs", defs),
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   536
         ("intros", intrs)];
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   537
  val (thy4, intrs'') =
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
   538
    thy3 |> PureThy.add_thms ((intr_names ~~ intrs') ~~ map #2 intr_specs)
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   539
    |>> Theory.parent_path;
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   540
  in
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   541
    (thy4,
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   542
      {defs = defs',
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   543
       bnd_mono = bnd_mono',
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   544
       dom_subset = dom_subset',
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   545
       intrs = intrs'',
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   546
       elim = elim',
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   547
       mk_cases = mk_cases,
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   548
       induct = induct,
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   549
       mutual_induct = mutual_induct})
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   550
  end;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   551
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   552
(*source version*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   553
fun add_inductive (srec_tms, sdom_sum) intr_srcs
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   554
    (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   555
  let
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   556
    val intr_atts = map (map (Attrib.global_attribute thy) o snd) intr_srcs;
17937
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   557
    val sintrs = map fst intr_srcs ~~ intr_atts;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   558
    val read = Sign.simple_read_term thy;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   559
    val rec_tms = map (read Ind_Syntax.iT) srec_tms;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   560
    val dom_sum = read Ind_Syntax.iT sdom_sum;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   561
    val intr_tms = map (read propT o snd o fst) sintrs;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   562
    val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs;
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   563
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   564
    val (thy', (((monos, con_defs), type_intrs), type_elims)) = thy
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   565
      |> IsarThy.apply_theorems raw_monos
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   566
      |>>> IsarThy.apply_theorems raw_con_defs
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   567
      |>>> IsarThy.apply_theorems raw_type_intrs
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   568
      |>>> IsarThy.apply_theorems raw_type_elims;
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   569
  in
17937
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   570
    add_inductive_i true (rec_tms, dom_sum) intr_specs
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   571
      (monos, con_defs, type_intrs, type_elims) thy'
17937
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   572
  end
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   573
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   574
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   575
(* outer syntax *)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   576
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 16855
diff changeset
   577
local structure P = OuterParse and K = OuterKeyword in
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   578
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   579
fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   580
  #1 o add_inductive doms (map P.triple_swap intrs) (monos, con_defs, type_intrs, type_elims);
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   581
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   582
val ind_decl =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   583
  (P.$$$ "domains" |-- P.!!! (P.enum1 "+" P.term --
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   584
      ((P.$$$ "\\<subseteq>" || P.$$$ "<=") |-- P.term))) --
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   585
  (P.$$$ "intros" |--
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   586
    P.!!! (Scan.repeat1 (P.opt_thm_name ":" -- P.prop))) --
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   587
  Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) [] --
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   588
  Scan.optional (P.$$$ "con_defs" |-- P.!!! P.xthms1) [] --
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   589
  Scan.optional (P.$$$ "type_intros" |-- P.!!! P.xthms1) [] --
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12725
diff changeset
   590
  Scan.optional (P.$$$ "type_elims" |-- P.!!! P.xthms1) []
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   591
  >> (Toplevel.theory o mk_ind);
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   592
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   593
val inductiveP = OuterSyntax.command (co_prefix ^ "inductive")
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
   594
  ("define " ^ co_prefix ^ "inductive sets") K.thy_decl ind_decl;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   595
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   596
val _ = OuterSyntax.add_keywords
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   597
  ["domains", "intros", "monos", "con_defs", "type_intros", "type_elims"];
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   598
val _ = OuterSyntax.add_parsers [inductiveP];
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   599
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   600
end;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   601
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   602
end;
15705
b5edb9dcec9a *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
   603