src/HOL/Tools/inductive_package.ML
author mengj
Fri, 18 Nov 2005 07:08:18 +0100
changeset 18198 95330fc0ea8d
parent 17985 d5d576b72371
child 18222 a8ccacce3b52
permissions -rw-r--r--
-- combined common CNF functions used by HOL and FOL axioms, the difference between conversion of HOL and FOL theorems only comes in when theorems are converted to ResClause.clause or ResHolClause.clause format.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Tools/inductive_package.ML
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     2
    ID:         $Id$
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
     4
    Author:     Stefan Berghofer, TU Muenchen
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
     5
    Author:     Markus Wenzel, TU Muenchen
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     6
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
     7
(Co)Inductive Definition module for HOL.
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     8
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
     9
Features:
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    10
  * least or greatest fixedpoints
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    11
  * user-specified product and sum constructions
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    12
  * mutually recursive definitions
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    13
  * definitions involving arbitrary monotone operators
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    14
  * automatically proves introduction and elimination rules
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    15
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    16
The recursive sets must *already* be declared as constants in the
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    17
current theory!
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    18
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    19
  Introduction rules have the form
8316
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
    20
  [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    21
  where M is some monotone operator (usually the identity)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    22
  P(x) is any side condition on the free variables
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    23
  ti, t are any terms
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    24
  Sj, Sk are two of the sets being defined in mutual recursion
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    25
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    26
Sums are used only for mutual recursion.  Products are used only to
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    27
derive "streamlined" induction rules for relations.
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    28
*)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    29
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    30
signature INDUCTIVE_PACKAGE =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    31
sig
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    32
  val quiet_mode: bool ref
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
    33
  val trace: bool ref
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    34
  val unify_consts: theory -> term list -> term list -> term list * term list
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
    35
  val split_rule_vars: term list -> thm -> thm
9116
9df44b5c610b get_inductive now returns None instead of raising an exception.
berghofe
parents: 9072
diff changeset
    36
  val get_inductive: theory -> string -> ({names: string list, coind: bool} *
9df44b5c610b get_inductive now returns None instead of raising an exception.
berghofe
parents: 9072
diff changeset
    37
    {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
9df44b5c610b get_inductive now returns None instead of raising an exception.
berghofe
parents: 9072
diff changeset
    38
     intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}) option
12400
f12f95e216e0 added the_mk_cases;
wenzelm
parents: 12338
diff changeset
    39
  val the_mk_cases: theory -> string -> string -> thm
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
    40
  val print_inductives: theory -> unit
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    41
  val mono_add_global: theory attribute
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    42
  val mono_del_global: theory attribute
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    43
  val get_monos: theory -> thm list
10910
058775a575db export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents: 10804
diff changeset
    44
  val inductive_forall_name: string
058775a575db export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents: 10804
diff changeset
    45
  val inductive_forall_def: thm
058775a575db export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents: 10804
diff changeset
    46
  val rulify: thm -> thm
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    47
  val inductive_cases: ((bstring * Attrib.src list) * string list) list -> theory -> theory
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
    48
  val inductive_cases_i: ((bstring * theory attribute list) * term list) list -> theory -> theory
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    49
  val add_inductive_i: bool -> bool -> bstring -> bool -> bool -> bool -> term list ->
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
    50
    ((bstring * term) * theory attribute list) list -> thm list -> theory -> theory *
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    51
      {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
    52
       intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}
11628
e57a6e51715e inductive: no collective atts;
wenzelm
parents: 11502
diff changeset
    53
  val add_inductive: bool -> bool -> string list ->
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15574
diff changeset
    54
    ((bstring * string) * Attrib.src list) list -> (thmref * Attrib.src list) list ->
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
    55
    theory -> theory *
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    56
      {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
    57
       intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
    58
  val setup: (theory -> theory) list
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    59
end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    60
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
    61
structure InductivePackage: INDUCTIVE_PACKAGE =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    62
struct
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
    63
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
    64
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    65
(** theory context references **)
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    66
15525
396268ad58b3 HOL.order -> Orderings.order due to restructering
nipkow
parents: 15463
diff changeset
    67
val mono_name = "Orderings.mono";
17010
5abc26872268 changed reference to Lfp.lfp to FixedPint.lfp, ditto for gfp
avigad
parents: 16975
diff changeset
    68
val gfp_name = "FixedPoint.gfp";
5abc26872268 changed reference to Lfp.lfp to FixedPint.lfp, ditto for gfp
avigad
parents: 16975
diff changeset
    69
val lfp_name = "FixedPoint.lfp";
12259
3949e7aba298 Set.vimage;
wenzelm
parents: 12180
diff changeset
    70
val vimage_name = "Set.vimage";
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
    71
val Const _ $ (vimage_f $ _) $ _ = HOLogic.dest_Trueprop (Thm.concl_of vimageD);
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
    72
11991
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    73
val inductive_forall_name = "HOL.induct_forall";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    74
val inductive_forall_def = thm "induct_forall_def";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    75
val inductive_conj_name = "HOL.induct_conj";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    76
val inductive_conj_def = thm "induct_conj_def";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    77
val inductive_conj = thms "induct_conj";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    78
val inductive_atomize = thms "induct_atomize";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    79
val inductive_rulify1 = thms "induct_rulify1";
da6ee05d9f3d use HOL.induct_XXX;
wenzelm
parents: 11880
diff changeset
    80
val inductive_rulify2 = thms "induct_rulify2";
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    81
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    82
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
    83
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
    84
(** theory data **)
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    85
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    86
(* data kind 'HOL/inductive' *)
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    87
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    88
type inductive_info =
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    89
  {names: string list, coind: bool} * {defs: thm list, elims: thm list, raw_induct: thm,
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    90
    induct: thm, intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm};
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    91
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    92
structure InductiveData = TheoryDataFun
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    93
(struct
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    94
  val name = "HOL/inductive";
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    95
  type T = inductive_info Symtab.table * thm list;
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    96
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    97
  val empty = (Symtab.empty, []);
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
    98
  val copy = I;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    99
  val extend = I;
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   100
  fun merge _ ((tab1, monos1), (tab2, monos2)) =
11502
e80a712982e1 prefer immediate monos;
wenzelm
parents: 11358
diff changeset
   101
    (Symtab.merge (K true) (tab1, tab2), Drule.merge_rules (monos1, monos2));
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   102
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   103
  fun print thy (tab, monos) =
16364
dc9f7066d80a refer to name spaces values instead of names;
wenzelm
parents: 16287
diff changeset
   104
    [Pretty.strs ("(co)inductives:" ::
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   105
      map #1 (NameSpace.extern_table (Sign.const_space thy, tab))),
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   106
     Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm_sg thy) monos)]
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8634
diff changeset
   107
    |> Pretty.chunks |> Pretty.writeln;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   108
end);
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   109
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   110
val print_inductives = InductiveData.print;
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   111
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   112
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   113
(* get and put data *)
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   114
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17325
diff changeset
   115
val get_inductive = Symtab.lookup o #1 o InductiveData.get;
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   116
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   117
fun the_inductive thy name =
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   118
  (case get_inductive thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   119
    NONE => error ("Unknown (co)inductive set " ^ quote name)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   120
  | SOME info => info);
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   121
12400
f12f95e216e0 added the_mk_cases;
wenzelm
parents: 12338
diff changeset
   122
val the_mk_cases = (#mk_cases o #2) oo the_inductive;
f12f95e216e0 added the_mk_cases;
wenzelm
parents: 12338
diff changeset
   123
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   124
fun put_inductives names info thy =
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   125
  let
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17325
diff changeset
   126
    fun upd ((tab, monos), name) = (Symtab.update_new (name, info) tab, monos);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   127
    val tab_monos = Library.foldl upd (InductiveData.get thy, names)
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   128
      handle Symtab.DUP name => error ("Duplicate definition of (co)inductive set " ^ quote name);
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   129
  in InductiveData.put tab_monos thy end;
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   130
8277
493707fcd8a6 added cases_of, cases;
wenzelm
parents: 8100
diff changeset
   131
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   132
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   133
(** monotonicity rules **)
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   134
9831
9b883c416aef tuned "mono" att setup;
wenzelm
parents: 9804
diff changeset
   135
val get_monos = #2 o InductiveData.get;
9b883c416aef tuned "mono" att setup;
wenzelm
parents: 9804
diff changeset
   136
fun map_monos f = InductiveData.map (Library.apsnd f);
8277
493707fcd8a6 added cases_of, cases;
wenzelm
parents: 8100
diff changeset
   137
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   138
fun mk_mono thm =
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   139
  let
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   140
    fun eq2mono thm' = [standard (thm' RS (thm' RS eq_to_mono))] @
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   141
      (case concl_of thm of
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   142
          (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   143
        | _ => [standard (thm' RS (thm' RS eq_to_mono2))]);
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   144
    val concl = concl_of thm
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   145
  in
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   146
    if Logic.is_equals concl then
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   147
      eq2mono (thm RS meta_eq_to_obj_eq)
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   148
    else if can (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl then
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   149
      eq2mono thm
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   150
    else [thm]
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   151
  end;
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   152
8634
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   153
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   154
(* attributes *)
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   155
9831
9b883c416aef tuned "mono" att setup;
wenzelm
parents: 9804
diff changeset
   156
fun mono_add_global (thy, thm) = (map_monos (Drule.add_rules (mk_mono thm)) thy, thm);
9b883c416aef tuned "mono" att setup;
wenzelm
parents: 9804
diff changeset
   157
fun mono_del_global (thy, thm) = (map_monos (Drule.del_rules (mk_mono thm)) thy, thm);
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   158
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   159
val mono_attr =
8634
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   160
 (Attrib.add_del_args mono_add_global mono_del_global,
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   161
  Attrib.add_del_args Attrib.undef_local_attribute Attrib.undef_local_attribute);
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   162
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   163
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   164
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   165
(** misc utilities **)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   166
5662
72a2e33d3b9e Added quiet_mode flag.
berghofe
parents: 5553
diff changeset
   167
val quiet_mode = ref false;
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   168
val trace = ref false;  (*for debugging*)
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   169
fun message s = if ! quiet_mode then () else writeln s;
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   170
fun clean_message s = if ! quick_and_dirty then () else message s;
5662
72a2e33d3b9e Added quiet_mode flag.
berghofe
parents: 5553
diff changeset
   171
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   172
fun coind_prefix true = "co"
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   173
  | coind_prefix false = "";
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   174
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   175
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   176
(*the following code ensures that each recursive set always has the
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   177
  same type in all introduction rules*)
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   178
fun unify_consts thy cs intr_ts =
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   179
  (let
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16786
diff changeset
   180
    val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   181
    fun varify (t, (i, ts)) =
16876
f57b38cced32 Logic.incr_tvar;
wenzelm
parents: 16861
diff changeset
   182
      let val t' = map_term_types (Logic.incr_tvar (i + 1)) (#1 (Type.varify (t, [])))
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   183
      in (maxidx_of_term t', t'::ts) end;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   184
    val (i, cs') = foldr varify (~1, []) cs;
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   185
    val (i', intr_ts') = foldr varify (i, []) intr_ts;
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16786
diff changeset
   186
    val rec_consts = fold add_term_consts_2 cs' [];
7446b4be013b tuned fold on terms;
wenzelm
parents: 16786
diff changeset
   187
    val intr_consts = fold add_term_consts_2 intr_ts' [];
16934
9ef19e3c7fdd Sign.typ_unify;
wenzelm
parents: 16876
diff changeset
   188
    fun unify (cname, cT) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   189
      let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts)
16934
9ef19e3c7fdd Sign.typ_unify;
wenzelm
parents: 16876
diff changeset
   190
      in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
9ef19e3c7fdd Sign.typ_unify;
wenzelm
parents: 16876
diff changeset
   191
    val (env, _) = fold unify rec_consts (Vartab.empty, i');
16287
7a03b4b4df67 Type.freeze;
wenzelm
parents: 16122
diff changeset
   192
    val subst = Type.freeze o map_term_types (Envir.norm_type env)
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   193
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   194
  in (map subst cs', map subst intr_ts')
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   195
  end) handle Type.TUNIFY =>
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   196
    (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   197
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   198
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   199
(*make injections used in mutually recursive definitions*)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   200
fun mk_inj cs sumT c x =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   201
  let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   202
    fun mk_inj' T n i =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   203
      if n = 1 then x else
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   204
      let val n2 = n div 2;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   205
          val Type (_, [T1, T2]) = T
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   206
      in
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   207
        if i <= n2 then
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 15032
diff changeset
   208
          Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   209
        else
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 15032
diff changeset
   210
          Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   211
      end
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   212
  in mk_inj' sumT (length cs) (1 + find_index_eq c cs)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   213
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   214
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   215
(*make "vimage" terms for selecting out components of mutually rec.def*)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   216
fun mk_vimage cs sumT t c = if length cs < 2 then t else
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   217
  let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   218
    val cT = HOLogic.dest_setT (fastype_of c);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   219
    val vimageT = [cT --> sumT, HOLogic.mk_setT sumT] ---> HOLogic.mk_setT cT
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   220
  in
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   221
    Const (vimage_name, vimageT) $
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   222
      Abs ("y", cT, mk_inj cs sumT c (Bound 0)) $ t
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   223
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   224
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   225
(** proper splitting **)
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   226
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   227
fun prod_factors p (Const ("Pair", _) $ t $ u) =
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   228
      p :: prod_factors (1::p) t @ prod_factors (2::p) u
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   229
  | prod_factors p _ = [];
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   230
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   231
fun mg_prod_factors ts (t $ u) fs = if t mem ts then
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   232
        let val f = prod_factors [] u
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17261
diff changeset
   233
        in AList.update (op =) (t, f inter (AList.lookup (op =) fs t) |> the_default f) fs end
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   234
      else mg_prod_factors ts u (mg_prod_factors ts t fs)
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   235
  | mg_prod_factors ts (Abs (_, _, t)) fs = mg_prod_factors ts t fs
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   236
  | mg_prod_factors ts _ fs = fs;
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   237
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   238
fun prodT_factors p ps (T as Type ("*", [T1, T2])) =
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   239
      if p mem ps then prodT_factors (1::p) ps T1 @ prodT_factors (2::p) ps T2
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   240
      else [T]
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   241
  | prodT_factors _ _ T = [T];
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   242
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   243
fun ap_split p ps (Type ("*", [T1, T2])) T3 u =
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   244
      if p mem ps then HOLogic.split_const (T1, T2, T3) $
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   245
        Abs ("v", T1, ap_split (2::p) ps T2 T3 (ap_split (1::p) ps T1
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   246
          (prodT_factors (2::p) ps T2 ---> T3) (incr_boundvars 1 u) $ Bound 0))
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   247
      else u
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   248
  | ap_split _ _ _ _ u =  u;
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   249
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   250
fun mk_tuple p ps (Type ("*", [T1, T2])) (tms as t::_) =
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   251
      if p mem ps then HOLogic.mk_prod (mk_tuple (1::p) ps T1 tms,
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   252
        mk_tuple (2::p) ps T2 (Library.drop (length (prodT_factors (1::p) ps T1), tms)))
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   253
      else t
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   254
  | mk_tuple _ _ _ (t::_) = t;
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   255
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   256
fun split_rule_var' ((t as Var (v, Type ("fun", [T1, T2])), ps), rl) =
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   257
      let val T' = prodT_factors [] ps T1 ---> T2
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   258
          val newt = ap_split [] ps T1 T2 (Var (v, T'))
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   259
          val cterm = Thm.cterm_of (Thm.theory_of_thm rl)
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   260
      in
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   261
          instantiate ([], [(cterm t, cterm newt)]) rl
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   262
      end
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   263
  | split_rule_var' (_, rl) = rl;
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   264
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   265
val remove_split = rewrite_rule [split_conv RS eq_reflection];
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   266
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   267
fun split_rule_vars vs rl = standard (remove_split (foldr split_rule_var'
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   268
  rl (mg_prod_factors vs (Thm.prop_of rl) [])));
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   269
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   270
fun split_rule vs rl = standard (remove_split (foldr split_rule_var'
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   271
  rl (List.mapPartial (fn (t as Var ((a, _), _)) =>
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   272
      Option.map (pair t) (AList.lookup (op =) vs a)) (term_vars (Thm.prop_of rl)))));
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   273
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   274
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   275
(** process rules **)
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   276
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   277
local
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   278
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   279
fun err_in_rule thy name t msg =
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   280
  error (cat_lines ["Ill-formed introduction rule " ^ quote name,
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   281
    Sign.string_of_term thy t, msg]);
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   282
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   283
fun err_in_prem thy name t p msg =
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   284
  error (cat_lines ["Ill-formed premise", Sign.string_of_term thy p,
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   285
    "in introduction rule " ^ quote name, Sign.string_of_term thy t, msg]);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   286
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   287
val bad_concl = "Conclusion of introduction rule must have form \"t : S_i\"";
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   288
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   289
val all_not_allowed =
11358
416ea5c009f5 now checks for leading meta-quantifiers and complains, instead of
paulson
parents: 11036
diff changeset
   290
    "Introduction rule must not have a leading \"!!\" quantifier";
416ea5c009f5 now checks for leading meta-quantifiers and complains, instead of
paulson
parents: 11036
diff changeset
   291
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   292
fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   293
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   294
in
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   295
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   296
fun check_rule thy cs ((name, rule), att) =
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   297
  let
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   298
    val concl = Logic.strip_imp_concl rule;
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   299
    val prems = Logic.strip_imp_prems rule;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   300
    val aprems = map (atomize_term thy) prems;
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   301
    val arule = Logic.list_implies (aprems, concl);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   302
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   303
    fun check_prem (prem, aprem) =
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   304
      if can HOLogic.dest_Trueprop aprem then ()
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   305
      else err_in_prem thy name rule prem "Non-atomic premise";
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   306
  in
11358
416ea5c009f5 now checks for leading meta-quantifiers and complains, instead of
paulson
parents: 11036
diff changeset
   307
    (case concl of
416ea5c009f5 now checks for leading meta-quantifiers and complains, instead of
paulson
parents: 11036
diff changeset
   308
      Const ("Trueprop", _) $ (Const ("op :", _) $ t $ u) =>
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   309
        if u mem cs then
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   310
          if exists (Logic.occs o rpair t) cs then
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   311
            err_in_rule thy name rule "Recursion term on left of member symbol"
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   312
          else List.app check_prem (prems ~~ aprems)
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   313
        else err_in_rule thy name rule bad_concl
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   314
      | Const ("all", _) $ _ => err_in_rule thy name rule all_not_allowed
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   315
      | _ => err_in_rule thy name rule bad_concl);
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   316
    ((name, arule), att)
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   317
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   318
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   319
val rulify =
13709
ec00ba43aee8 - No longer applies norm_hhf_rule
berghofe
parents: 13657
diff changeset
   320
  standard o
11036
3c30f7b97a50 use hol_simplify;
wenzelm
parents: 11005
diff changeset
   321
  hol_simplify inductive_rulify2 o hol_simplify inductive_rulify1 o
3c30f7b97a50 use hol_simplify;
wenzelm
parents: 11005
diff changeset
   322
  hol_simplify inductive_conj;
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   323
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   324
end;
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   325
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   326
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   327
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   328
(** properties of (co)inductive sets **)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   329
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   330
(* elimination rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   331
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   332
fun mk_elims cs cTs params intr_ts intr_names =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   333
  let
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   334
    val used = foldr add_term_names [] intr_ts;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   335
    val [aname, pname] = variantlist (["a", "P"], used);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   336
    val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   337
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   338
    fun dest_intr r =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   339
      let val Const ("op :", _) $ t $ u =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   340
        HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   341
      in (u, t, Logic.strip_imp_prems r) end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   342
8380
c96953faf0a4 removed tune_names;
wenzelm
parents: 8375
diff changeset
   343
    val intrs = map dest_intr intr_ts ~~ intr_names;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   344
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   345
    fun mk_elim (c, T) =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   346
      let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   347
        val a = Free (aname, T);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   348
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   349
        fun mk_elim_prem (_, t, ts) =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   350
          list_all_free (map dest_Free ((foldr add_term_frees [] (t::ts)) \\ params),
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   351
            Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (a, t)) :: ts, P));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   352
        val c_intrs = (List.filter (equal c o #1 o #1) intrs);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   353
      in
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   354
        (Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (a, c)) ::
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   355
          map mk_elim_prem (map #1 c_intrs), P), map #2 c_intrs)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   356
      end
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   357
  in
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   358
    map mk_elim (cs ~~ cTs)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   359
  end;
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   360
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   361
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   362
(* premises and conclusions of induction rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   363
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   364
fun mk_indrule cs cTs params intr_ts =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   365
  let
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   366
    val used = foldr add_term_names [] intr_ts;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   367
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   368
    (* predicates for induction rule *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   369
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   370
    val preds = map Free (variantlist (if length cs < 2 then ["P"] else
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   371
      map (fn i => "P" ^ string_of_int i) (1 upto length cs), used) ~~
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   372
        map (fn T => T --> HOLogic.boolT) cTs);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   373
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   374
    (* transform an introduction rule into a premise for induction rule *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   375
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   376
    fun mk_ind_prem r =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   377
      let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   378
        val frees = map dest_Free ((add_term_frees (r, [])) \\ params);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   379
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   380
        val pred_of = AList.lookup (op aconv) (cs ~~ preds);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   381
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   382
        fun subst (s as ((m as Const ("op :", T)) $ t $ u)) =
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   383
              (case pred_of u of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   384
                  NONE => (m $ fst (subst t) $ fst (subst u), NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   385
                | SOME P => (HOLogic.mk_binop inductive_conj_name (s, P $ t), SOME (s, P $ t)))
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   386
          | subst s =
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   387
              (case pred_of s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   388
                  SOME P => (HOLogic.mk_binop "op Int"
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   389
                    (s, HOLogic.Collect_const (HOLogic.dest_setT
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   390
                      (fastype_of s)) $ P), NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   391
                | NONE => (case s of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   392
                     (t $ u) => (fst (subst t) $ fst (subst u), NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   393
                   | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   394
                   | _ => (s, NONE)));
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   395
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   396
        fun mk_prem (s, prems) = (case subst s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   397
              (_, SOME (t, u)) => t :: u :: prems
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   398
            | (t, _) => t :: prems);
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   399
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   400
        val Const ("op :", _) $ t $ u =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   401
          HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   402
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   403
      in list_all_free (frees,
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   404
           Logic.list_implies (map HOLogic.mk_Trueprop (foldr mk_prem
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   405
             [] (map HOLogic.dest_Trueprop (Logic.strip_imp_prems r))),
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   406
               HOLogic.mk_Trueprop (valOf (pred_of u) $ t)))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   407
      end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   408
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   409
    val ind_prems = map mk_ind_prem intr_ts;
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   410
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   411
    val factors = Library.fold (mg_prod_factors preds) ind_prems [];
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   412
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   413
    (* make conclusions for induction rules *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   414
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   415
    fun mk_ind_concl ((c, P), (ts, x)) =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   416
      let val T = HOLogic.dest_setT (fastype_of c);
17485
c39871c52977 introduced AList module
haftmann
parents: 17412
diff changeset
   417
          val ps = AList.lookup (op =) factors P |> the_default [];
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   418
          val Ts = prodT_factors [] ps T;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   419
          val (frees, x') = foldr (fn (T', (fs, s)) =>
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   420
            ((Free (s, T'))::fs, Symbol.bump_string s)) ([], x) Ts;
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   421
          val tuple = mk_tuple [] ps T frees;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   422
      in ((HOLogic.mk_binop "op -->"
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   423
        (HOLogic.mk_mem (tuple, c), P $ tuple))::ts, x')
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   424
      end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   425
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   426
    val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   427
        (fst (foldr mk_ind_concl ([], "xa") (cs ~~ preds))))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   428
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   429
  in (preds, ind_prems, mutual_ind_concl,
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   430
    map (apfst (fst o dest_Free)) factors)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   431
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   432
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   433
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   434
(* prepare cases and induct rules *)
8316
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   435
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   436
(*
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   437
  transform mutual rule:
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   438
    HH ==> (x1:A1 --> P1 x1) & ... & (xn:An --> Pn xn)
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   439
  into i-th projection:
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   440
    xi:Ai ==> HH ==> Pi xi
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   441
*)
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   442
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   443
fun project_rules [name] rule = [(name, rule)]
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   444
  | project_rules names mutual_rule =
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   445
      let
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   446
        val n = length names;
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   447
        fun proj i =
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   448
          (if i < n then (fn th => th RS conjunct1) else I)
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   449
            (Library.funpow (i - 1) (fn th => th RS conjunct2) mutual_rule)
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   450
            RS mp |> Thm.permute_prems 0 ~1 |> Drule.standard;
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   451
      in names ~~ map proj (1 upto n) end;
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   452
12172
wenzelm
parents: 12165
diff changeset
   453
fun add_cases_induct no_elim no_induct names elims induct =
8316
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   454
  let
9405
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   455
    fun cases_spec (name, elim) thy =
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   456
      thy
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   457
      |> Theory.add_path (Sign.base_name name)
10279
b223a9a3350e InductAttrib;
wenzelm
parents: 10212
diff changeset
   458
      |> (#1 o PureThy.add_thms [(("cases", elim), [InductAttrib.cases_set_global name])])
9405
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   459
      |> Theory.parent_path;
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   460
    val cases_specs = if no_elim then [] else map2 cases_spec (names, elims);
8316
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   461
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   462
    fun induct_spec (name, th) = #1 o PureThy.add_thms
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   463
      [(("", RuleCases.save induct th), [InductAttrib.induct_set_global name])];
12172
wenzelm
parents: 12165
diff changeset
   464
    val induct_specs = if no_induct then [] else map induct_spec (project_rules names induct);
9405
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   465
  in Library.apply (cases_specs @ induct_specs) end;
8316
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   466
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   467
74639e19eca0 add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents: 8312
diff changeset
   468
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   469
(** proofs for (co)inductive sets **)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   470
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   471
(* prove monotonicity -- NOT subject to quick_and_dirty! *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   472
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   473
fun prove_mono setT fp_fun monos thy =
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   474
 (message "  Proving monotonicity ...";
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   475
  standard (Goal.prove thy [] []   (*NO quick_and_dirty here!*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   476
    (HOLogic.mk_Trueprop
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   477
      (Const (mono_name, (setT --> setT) --> HOLogic.boolT) $ fp_fun))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   478
    (fn _ => EVERY [rtac monoI 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   479
      REPEAT (ares_tac (List.concat (map mk_mono monos) @ get_monos thy) 1)])));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   480
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   481
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   482
(* prove introduction rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   483
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   484
fun prove_intrs coind mono fp_def intr_ts rec_sets_defs thy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   485
  let
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   486
    val _ = clean_message "  Proving the introduction rules ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   487
13657
c1637d60a846 Now applies standard' to "unfold" theorem (due to flex-flex constraints).
berghofe
parents: 13626
diff changeset
   488
    val unfold = standard' (mono RS (fp_def RS
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10065
diff changeset
   489
      (if coind then def_gfp_unfold else def_lfp_unfold)));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   490
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   491
    fun select_disj 1 1 = []
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   492
      | select_disj _ 1 = [rtac disjI1]
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   493
      | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   494
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   495
    val intrs = (1 upto (length intr_ts) ~~ intr_ts) |> map (fn (i, intr) =>
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   496
      rulify (standard (SkipProof.prove thy [] [] intr (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   497
       [rewrite_goals_tac rec_sets_defs,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   498
        stac unfold 1,
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   499
        REPEAT (resolve_tac [vimageI2, CollectI] 1),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   500
        (*Now 1-2 subgoals: the disjunction, perhaps equality.*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   501
        EVERY1 (select_disj (length intr_ts) i),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   502
        (*Not ares_tac, since refl must be tried before any equality assumptions;
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   503
          backtracking may occur if the premises have extra variables!*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   504
        DEPTH_SOLVE_1 (resolve_tac [refl, exI, conjI] 1 APPEND assume_tac 1),
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   505
        (*Now solve the equations like Inl 0 = Inl ?b2*)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   506
        REPEAT (rtac refl 1)]))))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   507
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   508
  in (intrs, unfold) end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   509
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   510
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   511
(* prove elimination rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   512
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   513
fun prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs thy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   514
  let
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   515
    val _ = clean_message "  Proving the elimination rules ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   516
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   517
    val rules1 = [CollectE, disjE, make_elim vimageD, exE];
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   518
    val rules2 = [conjE, Inl_neq_Inr, Inr_neq_Inl] @ map make_elim [Inl_inject, Inr_inject];
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   519
  in
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   520
    mk_elims cs cTs params intr_ts intr_names |> map (fn (t, cases) =>
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   521
      SkipProof.prove thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   522
        (fn prems => EVERY
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   523
          [cut_facts_tac [hd prems] 1,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   524
           rewrite_goals_tac rec_sets_defs,
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   525
           dtac (unfold RS subst) 1,
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   526
           REPEAT (FIRSTGOAL (eresolve_tac rules1)),
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   527
           REPEAT (FIRSTGOAL (eresolve_tac rules2)),
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   528
           EVERY (map (fn prem =>
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   529
             DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_sets_defs prem, conjI] 1)) (tl prems))])
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   530
        |> standard
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   531
        |> rulify
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   532
        |> RuleCases.name cases)
8375
0544749a5e8f mk_elims, add_cases_induct: name rule cases;
wenzelm
parents: 8336
diff changeset
   533
  end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   534
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   535
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   536
(* derivation of simplified elimination rules *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   537
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   538
local
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   539
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   540
(*cprop should have the form t:Si where Si is an inductive set*)
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   541
val mk_cases_err = "mk_cases: proposition not of form \"t : S_i\"";
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   542
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   543
(*delete needless equality assumptions*)
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   544
val refl_thin = prove_goal HOL.thy "!!P. a = a ==> P ==> P" (fn _ => [assume_tac 1]);
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   545
val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE, Pair_inject];
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   546
val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   547
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   548
fun simp_case_tac solved ss i =
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   549
  EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   550
  THEN_MAYBE (if solved then no_tac else all_tac);
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   551
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   552
in
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   553
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   554
fun mk_cases_i elims ss cprop =
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   555
  let
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   556
    val prem = Thm.assume cprop;
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   557
    val tac = ALLGOALS (simp_case_tac false ss) THEN prune_params_tac;
9298
7d9b562a750b use InductMethod.simp_case_tac;
wenzelm
parents: 9235
diff changeset
   558
    fun mk_elim rl = Drule.standard (Tactic.rule_by_tactic tac (prem RS rl));
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   559
  in
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   560
    (case get_first (try mk_elim) elims of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   561
      SOME r => r
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   562
    | NONE => error (Pretty.string_of (Pretty.block
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   563
        [Pretty.str mk_cases_err, Pretty.fbrk, Display.pretty_cterm cprop])))
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   564
  end;
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   565
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6092
diff changeset
   566
fun mk_cases elims s =
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   567
  mk_cases_i elims (simpset()) (Thm.read_cterm (Thm.theory_of_thm (hd elims)) (s, propT));
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   568
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   569
fun smart_mk_cases thy ss cprop =
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   570
  let
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   571
    val c = #1 (Term.dest_Const (Term.head_of (#2 (HOLogic.dest_mem (HOLogic.dest_Trueprop
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   572
      (Logic.strip_imp_concl (Thm.term_of cprop))))))) handle TERM _ => error mk_cases_err;
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   573
    val (_, {elims, ...}) = the_inductive thy c;
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   574
  in mk_cases_i elims ss cprop end;
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   575
11682
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   576
end;
d9063229b4a1 simp_case_tac is back again from induct_method.ML;
wenzelm
parents: 11628
diff changeset
   577
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   578
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   579
(* inductive_cases(_i) *)
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   580
12609
fb073a34b537 'inductive_cases': support 'and' form;
wenzelm
parents: 12527
diff changeset
   581
fun gen_inductive_cases prep_att prep_prop args thy =
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   582
  let
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   583
    val cert_prop = Thm.cterm_of thy o prep_prop (ProofContext.init thy);
12609
fb073a34b537 'inductive_cases': support 'and' form;
wenzelm
parents: 12527
diff changeset
   584
    val mk_cases = smart_mk_cases thy (Simplifier.simpset_of thy) o cert_prop;
fb073a34b537 'inductive_cases': support 'and' form;
wenzelm
parents: 12527
diff changeset
   585
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   586
    val facts = args |> map (fn ((a, atts), props) =>
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   587
     ((a, map (prep_att thy) atts), map (Thm.no_attributes o single o mk_cases) props));
12709
e29800eba5d1 IsarThy.theorems_i;
wenzelm
parents: 12609
diff changeset
   588
  in thy |> IsarThy.theorems_i Drule.lemmaK facts |> #1 end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   589
12172
wenzelm
parents: 12165
diff changeset
   590
val inductive_cases = gen_inductive_cases Attrib.global_attribute ProofContext.read_prop;
wenzelm
parents: 12165
diff changeset
   591
val inductive_cases_i = gen_inductive_cases (K I) ProofContext.cert_prop;
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   592
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   593
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   594
(* mk_cases_meth *)
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   595
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   596
fun mk_cases_meth (ctxt, raw_props) =
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   597
  let
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   598
    val thy = ProofContext.theory_of ctxt;
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14981
diff changeset
   599
    val ss = local_simpset_of ctxt;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   600
    val cprops = map (Thm.cterm_of thy o ProofContext.read_prop ctxt) raw_props;
10743
8ea821d1e3a4 Method.erule 0;
wenzelm
parents: 10735
diff changeset
   601
  in Method.erule 0 (map (smart_mk_cases thy ss) cprops) end;
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   602
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   603
val mk_cases_args = Method.syntax (Scan.lift (Scan.repeat1 Args.name));
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   604
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   605
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   606
(* prove induction rule *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   607
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   608
fun prove_indrule cs cTs sumT rec_const params intr_ts mono
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   609
    fp_def rec_sets_defs thy =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   610
  let
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   611
    val _ = clean_message "  Proving the induction rule ...";
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   612
12922
ed70a600f0ea clarified internal theory dependencies;
wenzelm
parents: 12902
diff changeset
   613
    val sum_case_rewrites =
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   614
      (if Context.theory_name thy = "Datatype" then
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16432
diff changeset
   615
        PureThy.get_thms thy (Name "sum.cases")
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   616
      else
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   617
        (case ThyInfo.lookup_theory "Datatype" of
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   618
          NONE => []
16975
34ed8d5d4f16 Sign.read_term;
wenzelm
parents: 16934
diff changeset
   619
        | SOME thy' =>
34ed8d5d4f16 Sign.read_term;
wenzelm
parents: 16934
diff changeset
   620
            if Theory.subthy (thy', thy) then
34ed8d5d4f16 Sign.read_term;
wenzelm
parents: 16934
diff changeset
   621
              PureThy.get_thms thy' (Name "sum.cases")
34ed8d5d4f16 Sign.read_term;
wenzelm
parents: 16934
diff changeset
   622
            else []))
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   623
      |> map mk_meta_eq;
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
   624
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   625
    val (preds, ind_prems, mutual_ind_concl, factors) =
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   626
      mk_indrule cs cTs params intr_ts;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   627
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   628
    val dummy = if !trace then
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   629
                (writeln "ind_prems = ";
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   630
                 List.app (writeln o Sign.string_of_term thy) ind_prems)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   631
            else ();
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   632
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   633
    (* make predicate for instantiation of abstract induction rule *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   634
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   635
    fun mk_ind_pred _ [P] = P
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   636
      | mk_ind_pred T Ps =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   637
         let val n = (length Ps) div 2;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   638
             val Type (_, [T1, T2]) = T
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
   639
         in Const ("Datatype.sum.sum_case",
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   640
           [T1 --> HOLogic.boolT, T2 --> HOLogic.boolT, T] ---> HOLogic.boolT) $
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   641
             mk_ind_pred T1 (Library.take (n, Ps)) $ mk_ind_pred T2 (Library.drop (n, Ps))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   642
         end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   643
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   644
    val ind_pred = mk_ind_pred sumT preds;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   645
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   646
    val ind_concl = HOLogic.mk_Trueprop
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   647
      (HOLogic.all_const sumT $ Abs ("x", sumT, HOLogic.mk_binop "op -->"
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   648
        (HOLogic.mk_mem (Bound 0, rec_const), ind_pred $ Bound 0)));
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   649
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   650
    (* simplification rules for vimage and Collect *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   651
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   652
    val vimage_simps = if length cs < 2 then [] else
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   653
      map (fn c => standard (SkipProof.prove thy [] []
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   654
        (HOLogic.mk_Trueprop (HOLogic.mk_eq
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   655
          (mk_vimage cs sumT (HOLogic.Collect_const sumT $ ind_pred) c,
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   656
           HOLogic.Collect_const (HOLogic.dest_setT (fastype_of c)) $
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   657
             List.nth (preds, find_index_eq c cs))))
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   658
        (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   659
          [rtac vimage_Collect 1, rewrite_goals_tac sum_case_rewrites, rtac refl 1]))) cs;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   660
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   661
    val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   662
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   663
    val dummy = if !trace then
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   664
                (writeln "raw_fp_induct = "; print_thm raw_fp_induct)
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   665
            else ();
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   666
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   667
    val induct = standard (SkipProof.prove thy [] ind_prems ind_concl
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   668
      (fn prems => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   669
        [rewrite_goals_tac [inductive_conj_def],
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   670
         rtac (impI RS allI) 1,
13626
282fbabec862 Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents: 13197
diff changeset
   671
         DETERM (etac raw_fp_induct 1),
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   672
         rewrite_goals_tac (map mk_meta_eq (vimage_Int::Int_Collect::vimage_simps)),
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   673
         fold_goals_tac rec_sets_defs,
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   674
         (*This CollectE and disjE separates out the introduction rules*)
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   675
         REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE, exE])),
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   676
         (*Now break down the individual cases.  No disjE here in case
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   677
           some premise involves disjunction.*)
13747
bf308fcfd08e Better treatment of equality in premises of inductive definitions. Less
paulson
parents: 13709
diff changeset
   678
         REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
15416
db69af736ebb Further fix to a bug (involving equational premises) in inductive definitions
paulson
parents: 15391
diff changeset
   679
         ALLGOALS (simp_tac (HOL_basic_ss addsimps sum_case_rewrites)),
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   680
         EVERY (map (fn prem =>
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   681
           DEPTH_SOLVE_1 (ares_tac [rewrite_rule [inductive_conj_def] prem, conjI, refl] 1)) prems)]));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   682
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   683
    val lemma = standard (SkipProof.prove thy [] []
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   684
      (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   685
        [rewrite_goals_tac rec_sets_defs,
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   686
         REPEAT (EVERY
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   687
           [REPEAT (resolve_tac [conjI, impI] 1),
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   688
            TRY (dtac vimageD 1), etac allE 1, dtac mp 1, atac 1,
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
   689
            rewrite_goals_tac sum_case_rewrites,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   690
            atac 1])]))
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   691
10988
e0016a009c17 Splitting of arguments of product types in induction rules is now less
berghofe
parents: 10910
diff changeset
   692
  in standard (split_rule factors (induct RS lemma)) end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   693
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   694
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   695
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   696
(** specification of (co)inductive sets **)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   697
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   698
fun cond_declare_consts declare_consts cs paramTs cnames =
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   699
  if declare_consts then
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   700
    Theory.add_consts_i (map (fn (c, n) => (Sign.base_name n, paramTs ---> fastype_of c, NoSyn)) (cs ~~ cnames))
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   701
  else I;
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   702
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   703
fun mk_ind_def declare_consts alt_name coind cs intr_ts monos thy
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   704
      params paramTs cTs cnames =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   705
  let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   706
    val sumT = fold_bal (fn (T, U) => Type ("+", [T, U])) cTs;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   707
    val setT = HOLogic.mk_setT sumT;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   708
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   709
    val fp_name = if coind then gfp_name else lfp_name;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   710
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   711
    val used = foldr add_term_names [] intr_ts;
5149
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   712
    val [sname, xname] = variantlist (["S", "x"], used);
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   713
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   714
    (* transform an introduction rule into a conjunction  *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   715
    (*   [| t : ... S_i ... ; ... |] ==> u : S_j          *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   716
    (* is transformed into                                *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   717
    (*   x = Inj_j u & t : ... Inj_i -`` S ... & ...      *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   718
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   719
    fun transform_rule r =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   720
      let
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   721
        val frees = map dest_Free ((add_term_frees (r, [])) \\ params);
5149
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   722
        val subst = subst_free
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   723
          (cs ~~ (map (mk_vimage cs sumT (Free (sname, setT))) cs));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   724
        val Const ("op :", _) $ t $ u =
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   725
          HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   726
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   727
      in foldr (fn ((x, T), P) => HOLogic.mk_exists (x, T, P))
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   728
        (foldr1 HOLogic.mk_conj
5149
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   729
          (((HOLogic.eq_const sumT) $ Free (xname, sumT) $ (mk_inj cs sumT u t))::
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   730
            (map (subst o HOLogic.dest_Trueprop)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   731
              (Logic.strip_imp_prems r)))) frees
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   732
      end
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   733
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   734
    (* make a disjunction of all introduction rules *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   735
5149
10f0be29c0d1 Fixed bug in transform_rule.
berghofe
parents: 5120
diff changeset
   736
    val fp_fun = absfree (sname, setT, (HOLogic.Collect_const sumT) $
7710
bf8cb3fc5d64 Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents: 7349
diff changeset
   737
      absfree (xname, sumT, foldr1 HOLogic.mk_disj (map transform_rule intr_ts)));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   738
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   739
    (* add definiton of recursive sets to theory *)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   740
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   741
    val rec_name = if alt_name = "" then
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   742
      space_implode "_" (map Sign.base_name cnames) else alt_name;
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   743
    val full_rec_name = if length cs < 2 then hd cnames
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   744
      else Sign.full_name thy rec_name;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   745
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   746
    val rec_const = list_comb
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   747
      (Const (full_rec_name, paramTs ---> setT), params);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   748
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   749
    val fp_def_term = Logic.mk_equals (rec_const,
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   750
      Const (fp_name, (setT --> setT) --> setT) $ fp_fun);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   751
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   752
    val def_terms = fp_def_term :: (if length cs < 2 then [] else
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   753
      map (fn c => Logic.mk_equals (c, mk_vimage cs sumT rec_const c)) cs);
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   754
8433
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   755
    val (thy', [fp_def :: rec_sets_defs]) =
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   756
      thy
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   757
      |> cond_declare_consts declare_consts cs paramTs cnames
8433
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   758
      |> (if length cs < 2 then I
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   759
          else Theory.add_consts_i [(rec_name, paramTs ---> setT, NoSyn)])
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   760
      |> Theory.add_path rec_name
9315
f793f05024f6 adapted PureThy.add_defs_i;
wenzelm
parents: 9298
diff changeset
   761
      |> PureThy.add_defss_i false [(("defs", def_terms), [])];
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   762
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   763
    val mono = prove_mono setT fp_fun monos thy'
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   764
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   765
  in (thy', mono, fp_def, rec_sets_defs, rec_const, sumT) end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   766
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   767
fun add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   768
    intros monos thy params paramTs cTs cnames induct_cases =
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   769
  let
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   770
    val _ =
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   771
      if verbose then message ("Proofs for " ^ coind_prefix coind ^ "inductive set(s) " ^
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   772
        commas_quote (map Sign.base_name cnames)) else ();
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   773
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   774
    val ((intr_names, intr_ts), intr_atts) = apfst split_list (split_list intros);
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   775
9939
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   776
    val (thy1, mono, fp_def, rec_sets_defs, rec_const, sumT) =
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   777
      mk_ind_def declare_consts alt_name coind cs intr_ts monos thy
9072
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   778
        params paramTs cTs cnames;
a4896cf23638 Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents: 8720
diff changeset
   779
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   780
    val (intrs, unfold) = prove_intrs coind mono fp_def intr_ts rec_sets_defs thy1;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   781
    val elims = if no_elim then [] else
9939
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   782
      prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs thy1;
8312
b470bc28b59d add_cases_induct: accomodate no_elim and no_ind flags;
wenzelm
parents: 8307
diff changeset
   783
    val raw_induct = if no_ind then Drule.asm_rl else
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   784
      if coind then standard (rule_by_tactic
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   785
        (rewrite_tac [mk_meta_eq vimage_Un] THEN
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   786
          fold_tac rec_sets_defs) (mono RS (fp_def RS def_Collect_coinduct)))
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   787
      else
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   788
        prove_indrule cs cTs sumT rec_const params intr_ts mono fp_def
9939
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   789
          rec_sets_defs thy1;
12165
14e94ad99861 mutual rules declared as ``consumes 0'';
wenzelm
parents: 12128
diff changeset
   790
    val induct =
14e94ad99861 mutual rules declared as ``consumes 0'';
wenzelm
parents: 12128
diff changeset
   791
      if coind orelse no_ind orelse length cs > 1 then (raw_induct, [RuleCases.consumes 0])
14e94ad99861 mutual rules declared as ``consumes 0'';
wenzelm
parents: 12128
diff changeset
   792
      else (raw_induct RSN (2, rev_mp), [RuleCases.consumes 1]);
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   793
9939
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   794
    val (thy2, intrs') =
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   795
      thy1 |> PureThy.add_thms ((intr_names ~~ intrs) ~~ intr_atts);
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   796
    val (thy3, ([intrs'', elims'], [induct'])) =
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   797
      thy2
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   798
      |> PureThy.add_thmss
11628
e57a6e51715e inductive: no collective atts;
wenzelm
parents: 11502
diff changeset
   799
        [(("intros", intrs'), []),
11005
86f662ba3c3f more robust handling of rule cases hints;
wenzelm
parents: 10988
diff changeset
   800
          (("elims", elims), [RuleCases.consumes 1])]
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   801
      |>>> PureThy.add_thms
12165
14e94ad99861 mutual rules declared as ``consumes 0'';
wenzelm
parents: 12128
diff changeset
   802
        [((coind_prefix coind ^ "induct", rulify (#1 induct)),
14e94ad99861 mutual rules declared as ``consumes 0'';
wenzelm
parents: 12128
diff changeset
   803
         (RuleCases.case_names induct_cases :: #2 induct))]
8433
8ae16c770fc8 adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8410
diff changeset
   804
      |>> Theory.parent_path;
9939
44af7faa677e tuned handling of "intros";
wenzelm
parents: 9893
diff changeset
   805
  in (thy3,
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   806
    {defs = fp_def :: rec_sets_defs,
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   807
     mono = mono,
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   808
     unfold = unfold,
13709
ec00ba43aee8 - No longer applies norm_hhf_rule
berghofe
parents: 13657
diff changeset
   809
     intrs = intrs',
7798
42e94b618f34 return stored thms with proper naming in derivation;
wenzelm
parents: 7710
diff changeset
   810
     elims = elims',
42e94b618f34 return stored thms with proper naming in derivation;
wenzelm
parents: 7710
diff changeset
   811
     mk_cases = mk_cases elims',
10729
1b3350c4ee92 handle proper rules;
wenzelm
parents: 10569
diff changeset
   812
     raw_induct = rulify raw_induct,
7798
42e94b618f34 return stored thms with proper naming in derivation;
wenzelm
parents: 7710
diff changeset
   813
     induct = induct'})
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   814
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   815
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   816
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   817
(* external interfaces *)
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   818
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   819
fun try_term f msg thy t =
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   820
  (case Library.try f t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15525
diff changeset
   821
    SOME x => x
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   822
  | NONE => error (msg ^ Sign.string_of_term thy t));
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   823
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   824
fun add_inductive_i verbose declare_consts alt_name coind no_elim no_ind cs pre_intros monos thy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   825
  let
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   826
    val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   827
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   828
    (*parameters should agree for all mutually recursive components*)
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   829
    val (_, params) = strip_comb (hd cs);
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   830
    val paramTs = map (try_term (snd o dest_Free) "Parameter in recursive\
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   831
      \ component is not a free variable: " thy) params;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   832
10735
dfaf75f0076f simplified quick_and_dirty stuff;
wenzelm
parents: 10729
diff changeset
   833
    val cTs = map (try_term (HOLogic.dest_setT o fastype_of)
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   834
      "Recursive component not of type set: " thy) cs;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   835
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   836
    val cnames = map (try_term (fst o dest_Const o head_of)
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   837
      "Recursive set not previously declared as constant: " thy) cs;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   838
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   839
    val save_thy = thy
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   840
      |> Theory.copy |> cond_declare_consts declare_consts cs paramTs cnames;
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   841
    val intros = map (check_rule save_thy cs) pre_intros;
8401
50d5f4402305 more robust case names of induct;
wenzelm
parents: 8380
diff changeset
   842
    val induct_cases = map (#1 o #1) intros;
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   843
9405
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   844
    val (thy1, result as {elims, induct, ...}) =
11628
e57a6e51715e inductive: no collective atts;
wenzelm
parents: 11502
diff changeset
   845
      add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs intros monos
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   846
        thy params paramTs cTs cnames induct_cases;
8307
6600c6e53111 add_cases_induct: induct_method setup;
wenzelm
parents: 8293
diff changeset
   847
    val thy2 = thy1
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   848
      |> put_inductives cnames ({names = cnames, coind = coind}, result)
12172
wenzelm
parents: 12165
diff changeset
   849
      |> add_cases_induct no_elim (no_ind orelse coind orelse length cs > 1)
14235
281295a1bbaa Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents: 13747
diff changeset
   850
          cnames elims induct;
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   851
  in (thy2, result) end;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   852
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   853
fun add_inductive verbose coind c_strings intro_srcs raw_monos thy =
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   854
  let
16975
34ed8d5d4f16 Sign.read_term;
wenzelm
parents: 16934
diff changeset
   855
    val cs = map (Sign.read_term thy) c_strings;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   856
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   857
    val intr_names = map (fst o fst) intro_srcs;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   858
    fun read_rule s = Thm.read_cterm thy (s, propT)
9405
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   859
      handle ERROR => error ("The error(s) above occurred for " ^ s);
3235873fdd90 improved error msg;
wenzelm
parents: 9315
diff changeset
   860
    val intr_ts = map (Thm.term_of o read_rule o snd o fst) intro_srcs;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   861
    val intr_atts = map (map (Attrib.global_attribute thy) o snd) intro_srcs;
16432
a6e8fb94a8ca accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   862
    val (cs', intr_ts') = unify_consts thy cs intr_ts;
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   863
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   864
    val (thy', monos) = thy |> IsarThy.apply_theorems raw_monos;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   865
  in
7020
75ff179df7b7 Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents: 6851
diff changeset
   866
    add_inductive_i verbose false "" coind false false cs'
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   867
      ((intr_names ~~ intr_ts') ~~ intr_atts) monos thy'
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   868
  end;
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   869
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   870
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   871
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   872
(** package setup **)
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   873
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   874
(* setup theory *)
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   875
8634
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   876
val setup =
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8433
diff changeset
   877
 [InductiveData.init,
9625
77506775481e renamed 'mk_cases_tac' to 'ind_cases';
wenzelm
parents: 9598
diff changeset
   878
  Method.add_methods [("ind_cases", mk_cases_meth oo mk_cases_args,
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   879
    "dynamic case analysis on sets")],
9893
93d2fde0306c tuned msg;
wenzelm
parents: 9831
diff changeset
   880
  Attrib.add_attributes [("mono", mono_attr, "declaration of monotonicity rule")]];
6437
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   881
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   882
9bdfe07ba8e9 'HOL/inductive' theory data;
wenzelm
parents: 6430
diff changeset
   883
(* outer syntax *)
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   884
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 17010
diff changeset
   885
local structure P = OuterParse and K = OuterKeyword in
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   886
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   887
fun mk_ind coind ((sets, intrs), monos) =
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   888
  #1 o add_inductive true coind sets (map P.triple_swap intrs) monos;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   889
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   890
fun ind_decl coind =
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   891
  Scan.repeat1 P.term --
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   892
  (P.$$$ "intros" |--
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   893
    P.!!! (Scan.repeat1 (P.opt_thm_name ":" -- P.prop))) --
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   894
  Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) []
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   895
  >> (Toplevel.theory o mk_ind coind);
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   896
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   897
val inductiveP =
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   898
  OuterSyntax.command "inductive" "define inductive sets" K.thy_decl (ind_decl false);
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   899
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   900
val coinductiveP =
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   901
  OuterSyntax.command "coinductive" "define coinductive sets" K.thy_decl (ind_decl true);
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   902
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   903
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   904
val ind_cases =
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12798
diff changeset
   905
  P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 P.prop)
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   906
  >> (Toplevel.theory o inductive_cases);
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   907
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   908
val inductive_casesP =
9804
ee0c337327cf "inductive_cases": proper command;
wenzelm
parents: 9643
diff changeset
   909
  OuterSyntax.command "inductive_cases"
9598
65ee72db0236 raplaced "intrs" by "intrs" (new-style only);
wenzelm
parents: 9562
diff changeset
   910
    "create simplified instances of elimination rules (improper)" K.thy_script ind_cases;
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   911
12180
91c9f661b183 inductive: removed con_defs;
wenzelm
parents: 12172
diff changeset
   912
val _ = OuterSyntax.add_keywords ["intros", "monos"];
7107
ce69de572bca inductive_cases(_i): Isar interface to mk_cases;
wenzelm
parents: 7020
diff changeset
   913
val _ = OuterSyntax.add_parsers [inductiveP, coinductiveP, inductive_casesP];
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   914
5094
ddcc3c114a0e New inductive definition package
berghofe
parents:
diff changeset
   915
end;
6424
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   916
ceab9e663e08 tuned comments;
wenzelm
parents: 6394
diff changeset
   917
end;
15705
b5edb9dcec9a *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
   918