src/ZF/Tools/inductive_package.ML
author wenzelm
Sat, 11 Sep 2021 22:02:12 +0200
changeset 74294 ee04dc00bf0a
parent 74290 b2ad24b5a42c
child 74296 abc878973216
permissions -rw-r--r--
more antiquotations;
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
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     3
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     4
Fixedpoint definition module -- for Inductive/Coinductive Definitions
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     5
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     6
The functor will be instantiated for normal sums/products (inductive defs)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     7
                         and non-standard sums/products (coinductive defs)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     8
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
     9
Sums are used only for mutual recursion;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    10
Products are used only to derive "streamlined" induction rules for relations
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    11
*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    12
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    13
type inductive_result =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    14
   {defs       : thm list,             (*definitions made in thy*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    15
    bnd_mono   : thm,                  (*monotonicity for the lfp definition*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    16
    dom_subset : thm,                  (*inclusion of recursive set in dom*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    17
    intrs      : thm list,             (*introduction rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    18
    elim       : thm,                  (*case analysis theorem*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    19
    induct     : thm,                  (*main induction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    20
    mutual_induct : thm};              (*mutual induction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    21
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    22
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    23
(*Functor's result signature*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    24
signature INDUCTIVE_PACKAGE =
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    25
sig
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    26
  (*Insert definitions for the recursive sets, which
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    27
     must *already* be declared as constants in parent theory!*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    28
  val add_inductive_i: bool -> term list * term ->
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29306
diff changeset
    29
    ((binding * term) * attribute list) list ->
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    30
    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
    31
  val add_inductive: string list * string ->
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57877
diff changeset
    32
    ((binding * string) * Token.src list) list ->
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57877
diff changeset
    33
    (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list *
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57877
diff changeset
    34
    (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list ->
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    35
    theory -> theory * inductive_result
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    36
end;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    37
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    38
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    39
(*Declares functions to add fixedpoint/constructor defs to a theory.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    40
  Recursive sets must *already* be declared as constants.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    41
functor Add_inductive_def_Fun
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    42
    (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
    43
 : INDUCTIVE_PACKAGE =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    44
struct
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
    45
12227
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
    46
val co_prefix = if coind then "co" else "";
c654c2c03f1d actually store "coinduct" rule;
wenzelm
parents: 12191
diff changeset
    47
7695
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    48
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    49
(* utils *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    50
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    51
(*make distinct individual variables a1, a2, a3, ..., an. *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    52
fun mk_frees a [] = []
12902
a23dc0b7566f Symbol.bump_string;
wenzelm
parents: 12876
diff changeset
    53
  | 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
    54
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
(* add_inductive(_i) *)
6d7f9f30e6df mk_frees, assume_read moved here;
wenzelm
parents: 7570
diff changeset
    57
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    58
(*internal version, accepting terms*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    59
fun add_inductive_i verbose (rec_tms, dom_sum)
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    60
  raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy0 =
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
    61
let
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    62
  val ctxt0 = Proof_Context.init_global thy0;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    63
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30190
diff changeset
    64
  val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs;
12191
2c383ee7ff16 case_names;
wenzelm
parents: 12183
diff changeset
    65
  val (intr_names, intr_tms) = split_list (map fst intr_specs);
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33040
diff changeset
    66
  val case_names = Rule_Cases.case_names intr_names;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    67
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    68
  (*recT and rec_params should agree for all mutually recursive components*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    69
  val rec_hds = map head_of rec_tms;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    70
57877
wenzelm
parents: 56491
diff changeset
    71
  val dummy = rec_hds |> forall (fn t => is_Const t orelse
wenzelm
parents: 56491
diff changeset
    72
      error ("Recursive set not previously declared as constant: " ^
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    73
                   Syntax.string_of_term ctxt0 t));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    74
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    75
  (*Now we know they are all Consts, so get their names, type and params*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    76
  val rec_names = map (#1 o dest_Const) rec_hds
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    77
  and (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    78
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
    79
  val rec_base_names = map Long_Name.base_name rec_names;
57877
wenzelm
parents: 56491
diff changeset
    80
  val dummy = rec_base_names |> forall (fn a => Symbol_Pos.is_identifier a orelse
wenzelm
parents: 56491
diff changeset
    81
    error ("Base name of recursive set not an identifier: " ^ a));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    82
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    83
  local (*Checking the introduction rules*)
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    84
    val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy0) intr_tms;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    85
    fun intr_ok set =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
    86
        case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    87
  in
57877
wenzelm
parents: 56491
diff changeset
    88
    val dummy = intr_sets |> forall (fn t => intr_ok t orelse
wenzelm
parents: 56491
diff changeset
    89
      error ("Conclusion of rule does not name a recursive set: " ^
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    90
                Syntax.string_of_term ctxt0 t));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    91
  end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    92
57877
wenzelm
parents: 56491
diff changeset
    93
  val dummy = rec_params |> forall (fn t => is_Free t orelse
wenzelm
parents: 56491
diff changeset
    94
      error ("Param in recursion term not a free variable: " ^
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
    95
               Syntax.string_of_term ctxt0 t));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    96
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    97
  (*** Construct the fixedpoint definition ***)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43597
diff changeset
    98
  val mk_variant = singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] intr_tms));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
    99
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   100
  val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   101
74294
ee04dc00bf0a more antiquotations;
wenzelm
parents: 74290
diff changeset
   102
  fun dest_tprop \<^Const_>\<open>Trueprop for P\<close> = P
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   103
    | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   104
                            Syntax.string_of_term ctxt0 Q);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   105
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   106
  (*Makes a disjunct from an introduction rule*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   107
  fun fp_part intr = (*quantify over rule's free vars except parameters*)
16855
7563d0eb3414 no open Logic;
wenzelm
parents: 16457
diff changeset
   108
    let val prems = map dest_tprop (Logic.strip_imp_prems intr)
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   109
        val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43597
diff changeset
   110
        val exfrees = subtract (op =) rec_params (Misc_Legacy.term_frees intr)
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   111
        val zeq = FOLogic.mk_eq (Free(z', Ind_Syntax.iT), #1 (Ind_Syntax.rule_concl intr))
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29579
diff changeset
   112
    in List.foldr FOLogic.mk_exists
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32091
diff changeset
   113
             (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) exfrees
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   114
    end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   115
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   116
  (*The Part(A,h) terms -- compose injections to make h*)
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   117
  fun mk_Part (Bound 0) = Free(X', Ind_Syntax.iT) (*no mutual rec, no Part needed*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   118
    | mk_Part h = \<^const>\<open>Part\<close> $ Free(X', Ind_Syntax.iT) $ Abs (w', Ind_Syntax.iT, h);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   119
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   120
  (*Access to balanced disjoint sums via injections*)
23419
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22675
diff changeset
   121
  val parts = map mk_Part
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32091
diff changeset
   122
    (Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0}
23419
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22675
diff changeset
   123
      (length rec_tms));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   124
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   125
  (*replace each set by the corresponding Part(A,h)*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   126
  val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   127
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   128
  val fp_abs =
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
   129
    absfree (X', Ind_Syntax.iT)
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
   130
        (Ind_Syntax.mk_Collect (z', dom_sum,
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   131
            Balanced_Tree.make FOLogic.mk_disj part_intrs));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   132
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   133
  val fp_rhs = Fp.oper $ dom_sum $ fp_abs
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   134
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22101
diff changeset
   135
  val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22101
diff changeset
   136
                             error "Illegal occurrence of recursion operator"; ()))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   137
           rec_hds;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   138
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   139
  (*** Make the new theory ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   140
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   141
  (*A key definition:
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   142
    If no mutual recursion then it equals the one recursive set.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   143
    If mutual recursion then it differs from all the recursive sets. *)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   144
  val big_rec_base_name = space_implode "_" rec_base_names;
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   145
  val big_rec_name = Proof_Context.intern_const ctxt0 big_rec_base_name;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   146
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   147
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21350
diff changeset
   148
  val _ =
279b129498b6 removed conditional combinator;
wenzelm
parents: 21350
diff changeset
   149
    if verbose then
279b129498b6 removed conditional combinator;
wenzelm
parents: 21350
diff changeset
   150
      writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name)
279b129498b6 removed conditional combinator;
wenzelm
parents: 21350
diff changeset
   151
    else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   152
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   153
  (*Big_rec... is the union of the mutually recursive sets*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   154
  val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   155
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   156
  (*The individual sets must already be declared*)
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37145
diff changeset
   157
  val axpairs = map Misc_Legacy.mk_defpair
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   158
        ((big_rec_tm, fp_rhs) ::
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   159
         (case parts of
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   160
             [_] => []                        (*no mutual recursion*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   161
           | _ => rec_tms ~~          (*define the sets as Parts*)
41449
7339f0e7c513 do not open ML structures;
wenzelm
parents: 39557
diff changeset
   162
                  map (subst_atomic [(Free (X', Ind_Syntax.iT), big_rec_tm)]) parts));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   163
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   164
  (*tracing: print the fixedpoint definition*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   165
  val dummy = if !Ind_Syntax.trace then
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   166
              writeln (cat_lines (map (Syntax.string_of_term ctxt0 o #2) axpairs))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   167
          else ()
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   168
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   169
  (*add definitions of the inductive sets*)
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   170
  val (_, thy1) =
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   171
    thy0
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24255
diff changeset
   172
    |> Sign.add_path big_rec_base_name
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39288
diff changeset
   173
    |> Global_Theory.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs);
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
   174
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   175
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   176
  (*fetch fp definitions from the theory*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   177
  val big_rec_def::part_rec_defs =
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37145
diff changeset
   178
    map (Misc_Legacy.get_def thy1)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   179
        (case rec_names of [_] => rec_names
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   180
                         | _   => big_rec_base_name::rec_names);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   181
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   182
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   183
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   184
  val dummy = writeln "  Proving monotonicity...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   185
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   186
  val bnd_mono0 =
20342
wenzelm
parents: 20071
diff changeset
   187
    Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs))
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   188
      (fn {context = ctxt, ...} => EVERY
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   189
        [resolve_tac ctxt [@{thm Collect_subset} RS @{thm bnd_monoI}] 1,
60774
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   190
         REPEAT (ares_tac ctxt (@{thms basic_monos} @ monos) 1)]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   191
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   192
  val dom_subset0 = Drule.export_without_context (big_rec_def RS Fp.subs);
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   193
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   194
  val ([bnd_mono, dom_subset], thy2) = thy1
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   195
    |> Global_Theory.add_thms
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   196
      [((Binding.name "bnd_mono", bnd_mono0), []),
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   197
       ((Binding.name "dom_subset", dom_subset0), [])];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   198
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33385
diff changeset
   199
  val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   200
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   201
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   202
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   203
  val dummy = writeln "  Proving the introduction rules...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   204
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   205
  (*Mutual recursion?  Helps to derive subset rules for the
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   206
    individual sets.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   207
  val Part_trans =
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   208
      case rec_names of
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   209
           [_] => asm_rl
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33385
diff changeset
   210
         | _   => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans});
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   211
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   212
  (*To type-check recursive occurrences of the inductive sets, possibly
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   213
    enclosed in some monotonic operator M.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   214
  val rec_typechecks =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   215
     [dom_subset] RL (asm_rl :: ([Part_trans] RL monos))
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24867
diff changeset
   216
     RL [@{thm subsetD}];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   217
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   218
  (*Type-checking is hardest aspect of proof;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   219
    disjIn selects the correct disjunct after unfolding*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   220
  fun intro_tacsf disjIn ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   221
    [DETERM (stac ctxt unfold 1),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   222
     REPEAT (resolve_tac ctxt [@{thm Part_eqI}, @{thm CollectI}] 1),
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   223
     (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   224
     resolve_tac ctxt [disjIn] 2,
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   225
     (*Not ares_tac, since refl must be tried before equality assumptions;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   226
       backtracking may occur if the premises have extra variables!*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   227
     DEPTH_SOLVE_1 (resolve_tac ctxt [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2),
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   228
     (*Now solve the equations like Tcons(a,f) = Inl(?b4)*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   229
     rewrite_goals_tac ctxt con_defs,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   230
     REPEAT (resolve_tac ctxt @{thms refl} 2),
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   231
     (*Typechecking; this can fail*)
56491
a8ccf3d6a6e4 proper context for print_tac;
wenzelm
parents: 56249
diff changeset
   232
     if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:"
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   233
     else all_tac,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   234
     REPEAT (FIRSTGOAL (dresolve_tac ctxt rec_typechecks
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   235
                        ORELSE' eresolve_tac ctxt (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   236
                                              type_elims)
59069
ec6ce25a630d more accurate context;
wenzelm
parents: 58963
diff changeset
   237
                        ORELSE' hyp_subst_tac ctxt)),
56491
a8ccf3d6a6e4 proper context for print_tac;
wenzelm
parents: 56249
diff changeset
   238
     if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:"
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   239
     else all_tac,
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   240
     DEPTH_SOLVE (swap_res_tac ctxt (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   241
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   242
  (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*)
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32091
diff changeset
   243
  val mk_disj_rls = Balanced_Tree.accesses
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
   244
    {left = fn rl => rl RS @{thm disjI1},
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
   245
     right = fn rl => rl RS @{thm disjI2},
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
   246
     init = @{thm asm_rl}};
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   247
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   248
  val intrs0 =
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   249
    (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms)))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   250
    |> ListPair.map (fn (t, tacs) =>
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   251
      Goal.prove_global thy2 [] [] t
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   252
        (fn {context = ctxt, ...} => EVERY (rewrite_goals_tac ctxt part_rec_defs :: tacs ctxt)));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   253
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   254
  val ([intrs], thy3) = thy2
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   255
    |> Global_Theory.add_thmss [((Binding.name "intros", intrs0), [])];
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   256
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   257
  val ctxt3 = Proof_Context.init_global thy3;
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   258
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   259
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   260
  (********)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   261
  val dummy = writeln "  Proving the elimination rule...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   262
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   263
  (*Breaks down logical connectives in the monotonic function*)
52087
f3075fc4f5f6 more precise treatment of theory vs. Proof.context;
wenzelm
parents: 51798
diff changeset
   264
  fun basic_elim_tac ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   265
      REPEAT (SOMEGOAL (eresolve_tac ctxt (Ind_Syntax.elim_rls @ Su.free_SEs)
52087
f3075fc4f5f6 more precise treatment of theory vs. Proof.context;
wenzelm
parents: 51798
diff changeset
   266
                ORELSE' bound_hyp_subst_tac ctxt))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   267
      THEN prune_params_tac ctxt
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   268
          (*Mutual recursion: collapse references to Part(D,h)*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   269
      THEN (PRIMITIVE (fold_rule ctxt part_rec_defs));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   270
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   271
  (*Elimination*)
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   272
  val (elim, thy4) = thy3
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   273
    |> Global_Theory.add_thm
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   274
      ((Binding.name "elim",
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   275
        rule_by_tactic ctxt3 (basic_elim_tac ctxt3) (unfold RS Ind_Syntax.equals_CollectD)), []);
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   276
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   277
  val ctxt4 = Proof_Context.init_global thy4;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   278
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   279
  (*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
   280
      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
   281
      con_defs=[] for inference systems.
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   282
    Proposition A should have the form t:Si where Si is an inductive set*)
36541
de1862c4a308 more explicit treatment of context, although not fully localized;
wenzelm
parents: 35989
diff changeset
   283
  fun make_cases ctxt A =
36546
a9873318fe30 proper context for rule_by_tactic;
wenzelm
parents: 36543
diff changeset
   284
    rule_by_tactic ctxt
52087
f3075fc4f5f6 more precise treatment of theory vs. Proof.context;
wenzelm
parents: 51798
diff changeset
   285
      (basic_elim_tac ctxt THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac ctxt)
12175
5cf58a1799a7 rearranged inductive package for Isar;
wenzelm
parents: 12132
diff changeset
   286
      (Thm.assume A RS elim)
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33385
diff changeset
   287
      |> Drule.export_without_context_open;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   288
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   289
  fun induction_rules raw_induct =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   290
   let
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   291
     val dummy = writeln "  Proving the induction rule...";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   292
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   293
     (*** Prove the main induction rule ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   294
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   295
     val pred_name = "P";            (*name for predicate variables*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   296
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   297
     (*Used to make induction rules;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   298
        ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   299
        prem is a premise of an intr rule*)
74294
ee04dc00bf0a more antiquotations;
wenzelm
parents: 74290
diff changeset
   300
     fun add_induct_prem ind_alist (prem as \<^Const_>\<open>Trueprop for \<open>\<^Const_>\<open>mem for t X\<close>\<close>\<close>, iprems) =
17314
04e21a27c0ad introduces some modern-style AList operations
haftmann
parents: 17057
diff changeset
   301
          (case AList.lookup (op aconv) ind_alist X of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15461
diff changeset
   302
               SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15461
diff changeset
   303
             | 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
   304
                 let fun mk_sb (rec_tm,pred) =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   305
                             (rec_tm, \<^const>\<open>Collect\<close> $ rec_tm $ pred)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   306
                 in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   307
       | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   308
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   309
     (*Make a premise of the induction rule.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   310
     fun induct_prem ind_alist intr =
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45625
diff changeset
   311
       let val xs = subtract (op =) rec_params (Misc_Legacy.term_frees intr)
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29579
diff changeset
   312
           val iprems = List.foldr (add_induct_prem ind_alist) []
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   313
                              (Logic.strip_imp_prems intr)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   314
           val (t,X) = Ind_Syntax.rule_concl intr
17314
04e21a27c0ad introduces some modern-style AList operations
haftmann
parents: 17057
diff changeset
   315
           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
   316
           val concl = FOLogic.mk_Trueprop (pred $ t)
46215
0da9433f959e discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents: 45625
diff changeset
   317
       in fold_rev Logic.all xs (Logic.list_implies (iprems,concl)) end
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   318
       handle Bind => error"Recursion term not found in conclusion";
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   319
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   320
     (*Minimizes backtracking by delivering the correct premise to each goal.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   321
       Intro rules with extra Vars in premises still cause some backtracking *)
60774
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   322
     fun ind_tac _ [] 0 = all_tac
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   323
       | ind_tac ctxt (prem::prems) i =
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   324
             DEPTH_SOLVE_1 (ares_tac ctxt [prem, @{thm refl}] i) THEN ind_tac ctxt prems (i-1);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   325
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   326
     val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   327
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   328
     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
   329
                         intr_tms;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   330
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   331
     val dummy =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   332
      if ! Ind_Syntax.trace then
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   333
        writeln (cat_lines
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   334
          (["ind_prems:"] @ map (Syntax.string_of_term ctxt4) ind_prems @
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   335
           ["raw_induct:", Thm.string_of_thm ctxt4 raw_induct]))
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   336
      else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   337
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   338
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   339
     (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules.
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   340
       If the premises get simplified, then the proofs could fail.*)
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44241
diff changeset
   341
     val min_ss =
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   342
           (empty_simpset ctxt4
60822
4f58f3662e7d more explicit context;
wenzelm
parents: 60774
diff changeset
   343
             |> Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   344
           setSolver (mk_solver "minimal"
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   345
                      (fn ctxt => resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58936
diff changeset
   346
                                   ORELSE' assume_tac ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   347
                                   ORELSE' eresolve_tac ctxt @{thms FalseE}));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   348
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   349
     val quant_induct =
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   350
       Goal.prove_global thy4 [] ind_prems
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   351
         (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred)))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   352
         (fn {context = ctxt, prems} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   353
           [rewrite_goals_tac ctxt part_rec_defs,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   354
            resolve_tac ctxt [@{thm impI} RS @{thm allI}] 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   355
            DETERM (eresolve_tac ctxt [raw_induct] 1),
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   356
            (*Push Part inside Collect*)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24867
diff changeset
   357
            full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   358
            (*This CollectE and disjE separates out the introduction rules*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   359
            REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm disjE}])),
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   360
            (*Now break down the individual cases.  No disjE here in case
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   361
              some premise involves disjunction.*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   362
            REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm exE}, @{thm conjE}]
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   363
                               ORELSE' (bound_hyp_subst_tac ctxt))),
60774
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   364
            ind_tac ctxt (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   365
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   366
     val dummy =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   367
      if ! Ind_Syntax.trace then
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   368
        writeln ("quant_induct:\n" ^ Thm.string_of_thm ctxt4 quant_induct)
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   369
      else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   370
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   371
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   372
     (*** Prove the simultaneous induction rule ***)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   373
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   374
     (*Make distinct predicates for each inductive set*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   375
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   376
     (*The components of the element type, several if it is a product*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   377
     val elem_type = CP.pseudo_type dom_sum;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   378
     val elem_factors = CP.factors elem_type;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   379
     val elem_frees = mk_frees "za" elem_factors;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   380
     val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   381
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   382
     (*Given a recursive set and its domain, return the "fsplit" predicate
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   383
       and a conclusion for the simultaneous induction rule.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   384
       NOTE.  This will not work for mutually recursive predicates.  Previously
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   385
       a summand 'domt' was also an argument, but this required the domain of
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   386
       mutual recursion to invariably be a disjoint sum.*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   387
     fun mk_predpair rec_tm =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   388
       let val rec_name = (#1 o dest_Const o head_of) rec_tm
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   389
           val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name,
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   390
                            elem_factors ---> FOLogic.oT)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   391
           val qconcl =
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29579
diff changeset
   392
             List.foldr FOLogic.mk_all
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   393
               (FOLogic.imp $
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   394
                (\<^const>\<open>mem\<close> $ elem_tuple $ rec_tm)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   395
                      $ (list_comb (pfree, elem_frees))) elem_frees
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   396
       in  (CP.ap_split elem_type FOLogic.oT pfree,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   397
            qconcl)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   398
       end;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   399
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   400
     val (preds,qconcls) = split_list (map mk_predpair rec_tms);
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   401
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   402
     (*Used to form simultaneous induction lemma*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   403
     fun mk_rec_imp (rec_tm,pred) =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   404
         FOLogic.imp $ (\<^const>\<open>mem\<close> $ Bound 0 $ rec_tm) $
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   405
                          (pred $ Bound 0);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   406
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   407
     (*To instantiate the main induction rule*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   408
     val induct_concl =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   409
         FOLogic.mk_Trueprop
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   410
           (Ind_Syntax.mk_all_imp
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   411
            (big_rec_tm,
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   412
             Abs("z", Ind_Syntax.iT,
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32091
diff changeset
   413
                 Balanced_Tree.make FOLogic.mk_conj
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   414
                 (ListPair.map mk_rec_imp (rec_tms, preds)))))
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   415
     and mutual_induct_concl =
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32091
diff changeset
   416
      FOLogic.mk_Trueprop (Balanced_Tree.make FOLogic.mk_conj qconcls);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   417
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   418
     val dummy = if !Ind_Syntax.trace then
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   419
                 (writeln ("induct_concl = " ^
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   420
                           Syntax.string_of_term ctxt4 induct_concl);
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   421
                  writeln ("mutual_induct_concl = " ^
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   422
                           Syntax.string_of_term ctxt4 mutual_induct_concl))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   423
             else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   424
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   425
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   426
     fun lemma_tac ctxt =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   427
      FIRST' [eresolve_tac ctxt [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}],
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   428
              resolve_tac ctxt [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}],
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   429
              dresolve_tac ctxt [@{thm spec}, @{thm mp}, Pr.fsplitD]];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   430
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   431
     val need_mutual = length rec_names > 1;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   432
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   433
     val lemma = (*makes the link between the two induction rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   434
       if need_mutual then
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   435
          (writeln "  Proving the mutual induction rule...";
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   436
           Goal.prove_global thy4 [] []
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   437
             (Logic.mk_implies (induct_concl, mutual_induct_concl))
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   438
             (fn {context = ctxt, ...} => EVERY
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   439
               [rewrite_goals_tac ctxt part_rec_defs,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   440
                REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac ctxt 1)]))
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26056
diff changeset
   441
       else (writeln "  [ No mutual induction rule needed ]"; @{thm TrueI});
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   442
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   443
     val dummy =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   444
      if ! Ind_Syntax.trace then
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   445
        writeln ("lemma: " ^ Thm.string_of_thm ctxt4 lemma)
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 30609
diff changeset
   446
      else ();
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   447
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   448
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   449
     (*Mutual induction follows by freeness of Inl/Inr.*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   450
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   451
     (*Simplification largely reduces the mutual induction rule to the
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   452
       standard rule*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   453
     val mut_ss =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   454
         min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff];
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   455
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   456
     val all_defs = con_defs @ part_rec_defs;
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   457
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   458
     (*Removes Collects caused by M-operators in the intro rules.  It is very
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   459
       hard to simplify
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   460
         list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))})
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   461
       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
   462
       Instead the following rules extract the relevant conjunct.
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   463
     *)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24867
diff changeset
   464
     val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24867
diff changeset
   465
                   RLN (2,[@{thm rev_subsetD}]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   466
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   467
     (*Minimizes backtracking by delivering the correct premise to each goal*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   468
     fun mutual_ind_tac _ [] 0 = all_tac
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   469
       | mutual_ind_tac ctxt (prem::prems) i =
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   470
           DETERM
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   471
            (SELECT_GOAL
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   472
               (
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   473
                (*Simplify the assumptions and goal by unfolding Part and
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   474
                  using freeness of the Sum constructors; proves all but one
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   475
                  conjunct by contradiction*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   476
                rewrite_goals_tac ctxt all_defs  THEN
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24867
diff changeset
   477
                simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1  THEN
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   478
                IF_UNSOLVED (*simp_tac may have finished it off!*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   479
                  ((*simplify assumptions*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   480
                   (*some risk of excessive simplification here -- might have
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   481
                     to identify the bare minimum set of rewrites*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   482
                   full_simp_tac
26287
df8e5362cff9 proper antiquotations;
wenzelm
parents: 26189
diff changeset
   483
                      (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   484
                   THEN
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   485
                   (*unpackage and use "prem" in the corresponding place*)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   486
                   REPEAT (resolve_tac ctxt @{thms impI} 1)  THEN
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   487
                   resolve_tac ctxt [rewrite_rule ctxt all_defs prem] 1  THEN
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   488
                   (*prem must not be REPEATed below: could loop!*)
60774
6c28d8ed2488 proper context;
wenzelm
parents: 60642
diff changeset
   489
                   DEPTH_SOLVE (FIRSTGOAL (ares_tac ctxt [@{thm impI}] ORELSE'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   490
                                           eresolve_tac ctxt (@{thm conjE} :: @{thm mp} :: cmonos))))
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   491
               ) i)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   492
            THEN mutual_ind_tac ctxt prems (i-1);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   493
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   494
     val mutual_induct_fsplit =
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   495
       if need_mutual then
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   496
         Goal.prove_global thy4 [] (map (induct_prem (rec_tms~~preds)) intr_tms)
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   497
           mutual_induct_concl
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   498
           (fn {context = ctxt, prems} => EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59069
diff changeset
   499
             [resolve_tac ctxt [quant_induct RS lemma] 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   500
              mutual_ind_tac ctxt (rev prems) (length prems)])
35409
5c5bb83f2bae more antiquotations;
wenzelm
parents: 35237
diff changeset
   501
       else @{thm TrueI};
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   502
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   503
     (** Uncurrying the predicate in the ordinary induction rule **)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   504
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   505
     (*instantiate the variable to a tuple, if it is non-trivial, in order to
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   506
       allow the predicate to be "opened up".
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   507
       The name "x.1" comes from the "RS spec" !*)
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   508
     val inst =
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   509
         case elem_frees of [_] => I
74290
b2ad24b5a42c more antiquotations;
wenzelm
parents: 74282
diff changeset
   510
            | _ => Drule.instantiate_normalize (TVars.empty,
b2ad24b5a42c more antiquotations;
wenzelm
parents: 74282
diff changeset
   511
                    Vars.make [(\<^var>\<open>?x1::i\<close>, Thm.global_cterm_of thy4 elem_tuple)]);
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   512
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   513
     (*strip quantifier and the implication*)
35409
5c5bb83f2bae more antiquotations;
wenzelm
parents: 35237
diff changeset
   514
     val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   515
74294
ee04dc00bf0a more antiquotations;
wenzelm
parents: 74290
diff changeset
   516
     val \<^Const_>\<open>Trueprop for \<open>pred_var $ _\<close>\<close> = Thm.concl_of induct0
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   517
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   518
     val induct0 =
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   519
       CP.split_rule_var ctxt4
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   520
        (pred_var, elem_type-->FOLogic.oT, induct0)
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52145
diff changeset
   521
       |> Drule.export_without_context
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   522
     and mutual_induct = CP.remove_split ctxt4 mutual_induct_fsplit
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   523
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   524
     val ([induct, mutual_induct], thy5) =
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   525
       thy4
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   526
       |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct0),
24861
cc669ca5f382 tuned Induct interface: prefer pred'' over set'';
wenzelm
parents: 24830
diff changeset
   527
             [case_names, Induct.induct_pred big_rec_name]),
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29306
diff changeset
   528
           ((Binding.name "mutual_induct", mutual_induct), [case_names])];
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   529
    in ((induct, mutual_induct), thy5)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   530
    end;  (*of induction_rules*)
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   531
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33385
diff changeset
   532
  val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct)
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   533
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   534
  val ((induct, mutual_induct), thy5) =
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   535
    if not coind then induction_rules raw_induct
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   536
    else
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   537
      thy4
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39288
diff changeset
   538
      |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])]
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   539
      |> apfst (hd #> pair @{thm TrueI});
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   540
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   541
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   542
  val (([cases], [defs]), thy6) =
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   543
    thy5
12183
c10cea75dd56 adapted primrec/datatype to Isar;
wenzelm
parents: 12175
diff changeset
   544
    |> IndCases.declare big_rec_name make_cases
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39288
diff changeset
   545
    |> Global_Theory.add_thms
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   546
      [((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])]
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   547
    ||>> (Global_Theory.add_thmss o map Thm.no_attributes)
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   548
      [(Binding.name "defs", big_rec_def :: part_rec_defs)];
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   549
  val (named_intrs, thy7) =
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   550
    thy6
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   551
    |> Global_Theory.add_thms ((map Binding.name intr_names ~~ intrs) ~~ map #2 intr_specs)
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24255
diff changeset
   552
    ||> Sign.parent_path;
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   553
  in
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   554
    (thy7,
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   555
      {defs = defs,
70411
c533bec6e92d more accurate proof definitions (PThm nodes);
wenzelm
parents: 69597
diff changeset
   556
       bnd_mono = bnd_mono,
71080
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   557
       dom_subset = dom_subset,
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   558
       intrs = named_intrs,
64249a83bc29 clarified contexts;
wenzelm
parents: 70411
diff changeset
   559
       elim = cases,
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   560
       induct = induct,
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   561
       mutual_induct = mutual_induct})
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7695
diff changeset
   562
  end;
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   563
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   564
(*source version*)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   565
fun add_inductive (srec_tms, sdom_sum) intr_srcs
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   566
    (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
   567
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   568
    val ctxt = Proof_Context.init_global thy;
39288
f1ae2493d93f eliminated aliases of Type.constraint;
wenzelm
parents: 38522
diff changeset
   569
    val read_terms = map (Syntax.parse_term ctxt #> Type.constraint Ind_Syntax.iT)
24726
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   570
      #> Syntax.check_terms ctxt;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   571
56031
wenzelm
parents: 54895
diff changeset
   572
    val intr_atts = map (map (Attrib.attribute_cmd ctxt) o snd) intr_srcs;
17937
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   573
    val sintrs = map fst intr_srcs ~~ intr_atts;
24726
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   574
    val rec_tms = read_terms srec_tms;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   575
    val dom_sum = singleton read_terms sdom_sum;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   576
    val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs);
17937
dfc9d3e54213 removed add_inductive_x;
wenzelm
parents: 17892
diff changeset
   577
    val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs;
24726
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   578
    val monos = Attrib.eval_thms ctxt raw_monos;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   579
    val con_defs = Attrib.eval_thms ctxt raw_con_defs;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   580
    val type_intrs = Attrib.eval_thms ctxt raw_type_intrs;
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   581
    val type_elims = Attrib.eval_thms ctxt raw_type_elims;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   582
  in
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   583
    thy
24726
fcf13a91cda2 Attrib.eval_thms;
wenzelm
parents: 24712
diff changeset
   584
    |> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims)
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   585
  end;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   586
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   587
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   588
(* outer syntax *)
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   589
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   590
fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) =
61466
9a468c3a1fa1 tuned signature;
wenzelm
parents: 61268
diff changeset
   591
  #1 o add_inductive doms (map (fn ((x, y), z) => ((x, z), y)) intrs)
9a468c3a1fa1 tuned signature;
wenzelm
parents: 61268
diff changeset
   592
    (monos, con_defs, type_intrs, type_elims);
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   593
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   594
val ind_decl =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   595
  (\<^keyword>\<open>domains\<close> |-- Parse.!!! (Parse.enum1 "+" Parse.term --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   596
      ((\<^keyword>\<open>\<subseteq>\<close> || \<^keyword>\<open><=\<close>) |-- Parse.term))) --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   597
  (\<^keyword>\<open>intros\<close> |--
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   598
    Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) --
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   599
  Scan.optional (\<^keyword>\<open>monos\<close> |-- Parse.!!! Parse.thms1) [] --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   600
  Scan.optional (\<^keyword>\<open>con_defs\<close> |-- Parse.!!! Parse.thms1) [] --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   601
  Scan.optional (\<^keyword>\<open>type_intros\<close> |-- Parse.!!! Parse.thms1) [] --
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   602
  Scan.optional (\<^keyword>\<open>type_elims\<close> |-- Parse.!!! Parse.thms1) []
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   603
  >> (Toplevel.theory o mk_ind);
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   604
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   605
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   606
  Outer_Syntax.command
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62969
diff changeset
   607
    (if coind then \<^command_keyword>\<open>coinductive\<close> else \<^command_keyword>\<open>inductive\<close>)
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   608
    ("define " ^ co_prefix ^ "inductive sets") ind_decl;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   609
6051
7d457fc538e7 revised inductive definition package
paulson
parents:
diff changeset
   610
end;
12132
1ef58b332ca9 support co/inductive definitions in new-style theories;
wenzelm
parents: 11680
diff changeset
   611