src/Tools/induct.ML
author wenzelm
Wed, 08 Jun 2011 15:56:57 +0200
changeset 44158 1fbdcebb364b
parent 43360 4638622bcaa1
child 44208 47cf4bc789aa
permissions -rw-r--r--
more robust exception pattern General.Subscript;
wenzelm@24830
     1
(*  Title:      Tools/induct.ML
wenzelm@24830
     2
    Author:     Markus Wenzel, TU Muenchen
wenzelm@24830
     3
wenzelm@26924
     4
Proof by cases, induction, and coinduction.
wenzelm@24830
     5
*)
wenzelm@24830
     6
wenzelm@37525
     7
signature INDUCT_ARGS =
wenzelm@24830
     8
sig
wenzelm@24830
     9
  val cases_default: thm
wenzelm@24830
    10
  val atomize: thm list
wenzelm@24830
    11
  val rulify: thm list
wenzelm@24830
    12
  val rulify_fallback: thm list
berghofe@34974
    13
  val equal_def: thm
berghofe@34907
    14
  val dest_def: term -> (term * term) option
berghofe@34907
    15
  val trivial_tac: int -> tactic
wenzelm@24830
    16
end;
wenzelm@24830
    17
wenzelm@24830
    18
signature INDUCT =
wenzelm@24830
    19
sig
wenzelm@24830
    20
  (*rule declarations*)
wenzelm@24830
    21
  val vars_of: term -> term list
wenzelm@24830
    22
  val dest_rules: Proof.context ->
wenzelm@24861
    23
    {type_cases: (string * thm) list, pred_cases: (string * thm) list,
wenzelm@24861
    24
      type_induct: (string * thm) list, pred_induct: (string * thm) list,
wenzelm@24861
    25
      type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list}
wenzelm@24830
    26
  val print_rules: Proof.context -> unit
wenzelm@24830
    27
  val lookup_casesT: Proof.context -> string -> thm option
wenzelm@24861
    28
  val lookup_casesP: Proof.context -> string -> thm option
wenzelm@24830
    29
  val lookup_inductT: Proof.context -> string -> thm option
wenzelm@24861
    30
  val lookup_inductP: Proof.context -> string -> thm option
wenzelm@24830
    31
  val lookup_coinductT: Proof.context -> string -> thm option
wenzelm@24861
    32
  val lookup_coinductP: Proof.context -> string -> thm option
wenzelm@24830
    33
  val find_casesT: Proof.context -> typ -> thm list
wenzelm@24861
    34
  val find_casesP: Proof.context -> term -> thm list
wenzelm@24830
    35
  val find_inductT: Proof.context -> typ -> thm list
wenzelm@24861
    36
  val find_inductP: Proof.context -> term -> thm list
wenzelm@24830
    37
  val find_coinductT: Proof.context -> typ -> thm list
wenzelm@24861
    38
  val find_coinductP: Proof.context -> term -> thm list
wenzelm@24830
    39
  val cases_type: string -> attribute
wenzelm@24861
    40
  val cases_pred: string -> attribute
wenzelm@27140
    41
  val cases_del: attribute
wenzelm@24830
    42
  val induct_type: string -> attribute
wenzelm@24861
    43
  val induct_pred: string -> attribute
wenzelm@27140
    44
  val induct_del: attribute
wenzelm@24830
    45
  val coinduct_type: string -> attribute
wenzelm@24861
    46
  val coinduct_pred: string -> attribute
wenzelm@27140
    47
  val coinduct_del: attribute
berghofe@34907
    48
  val map_simpset: (simpset -> simpset) -> Context.generic -> Context.generic
wenzelm@36602
    49
  val induct_simp_add: attribute
wenzelm@36602
    50
  val induct_simp_del: attribute
berghofe@34907
    51
  val no_simpN: string
wenzelm@24830
    52
  val casesN: string
wenzelm@24830
    53
  val inductN: string
wenzelm@24830
    54
  val coinductN: string
wenzelm@24830
    55
  val typeN: string
wenzelm@24861
    56
  val predN: string
wenzelm@24830
    57
  val setN: string
wenzelm@24830
    58
  (*proof methods*)
wenzelm@24830
    59
  val fix_tac: Proof.context -> int -> (string * typ) list -> int -> tactic
berghofe@34907
    60
  val add_defs: (binding option * (term * bool)) option list -> Proof.context ->
wenzelm@24830
    61
    (term option list * thm list) * Proof.context
wenzelm@24830
    62
  val atomize_term: theory -> term -> term
berghofe@34907
    63
  val atomize_cterm: conv
wenzelm@24830
    64
  val atomize_tac: int -> tactic
wenzelm@24830
    65
  val inner_atomize_tac: int -> tactic
wenzelm@24830
    66
  val rulified_term: thm -> theory * term
wenzelm@24830
    67
  val rulify_tac: int -> tactic
berghofe@34907
    68
  val simplified_rule: Proof.context -> thm -> thm
berghofe@34907
    69
  val simplify_tac: Proof.context -> int -> tactic
berghofe@34907
    70
  val trivial_tac: int -> tactic
berghofe@34907
    71
  val rotate_tac: int -> int -> int -> tactic
wenzelm@24830
    72
  val internalize: int -> thm -> thm
wenzelm@26940
    73
  val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq
berghofe@34974
    74
  val cases_tac: Proof.context -> bool -> term option list list -> thm option ->
wenzelm@24830
    75
    thm list -> int -> cases_tactic
wenzelm@27323
    76
  val get_inductT: Proof.context -> term option list list -> thm list list
berghofe@34907
    77
  val induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list ->
wenzelm@26924
    78
    (string * typ) list list -> term option list -> thm list option ->
wenzelm@26924
    79
    thm list -> int -> cases_tactic
wenzelm@26924
    80
  val coinduct_tac: Proof.context -> term option list -> term option list -> thm option ->
wenzelm@26924
    81
    thm list -> int -> cases_tactic
wenzelm@24830
    82
  val setup: theory -> theory
wenzelm@24830
    83
end;
wenzelm@24830
    84
wenzelm@37525
    85
functor Induct(Induct_Args: INDUCT_ARGS): INDUCT =
wenzelm@24830
    86
struct
wenzelm@24830
    87
wenzelm@37524
    88
(** variables -- ordered left-to-right, preferring right **)
wenzelm@24830
    89
wenzelm@24830
    90
fun vars_of tm =
wenzelm@24830
    91
  rev (distinct (op =) (Term.fold_aterms (fn (t as Var _) => cons t | _ => I) tm []));
wenzelm@24830
    92
wenzelm@24830
    93
local
wenzelm@24830
    94
wenzelm@37524
    95
val mk_var = Net.encode_type o #2 o Term.dest_Var;
wenzelm@24830
    96
wenzelm@24830
    97
fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle Empty =>
wenzelm@24830
    98
  raise THM ("No variables in conclusion of rule", 0, [thm]);
wenzelm@24830
    99
wenzelm@24830
   100
in
wenzelm@24830
   101
wenzelm@24830
   102
fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle Empty =>
wenzelm@24830
   103
  raise THM ("No variables in major premise of rule", 0, [thm]);
wenzelm@24830
   104
wenzelm@24830
   105
val left_var_concl = concl_var hd;
wenzelm@24830
   106
val right_var_concl = concl_var List.last;
wenzelm@24830
   107
wenzelm@24830
   108
end;
wenzelm@24830
   109
wenzelm@24830
   110
wenzelm@24830
   111
berghofe@34907
   112
(** constraint simplification **)
berghofe@34907
   113
berghofe@34907
   114
(* rearrange parameters and premises to allow application of one-point-rules *)
berghofe@34907
   115
berghofe@34907
   116
fun swap_params_conv ctxt i j cv =
berghofe@34907
   117
  let
berghofe@34907
   118
    fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt
berghofe@34907
   119
      | conv1 k ctxt =
berghofe@34907
   120
          Conv.rewr_conv @{thm swap_params} then_conv
wenzelm@37526
   121
          Conv.forall_conv (conv1 (k - 1) o snd) ctxt
berghofe@34907
   122
    fun conv2 0 ctxt = conv1 j ctxt
wenzelm@37526
   123
      | conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt
berghofe@34907
   124
  in conv2 i ctxt end;
berghofe@34907
   125
berghofe@34907
   126
fun swap_prems_conv 0 = Conv.all_conv
berghofe@34907
   127
  | swap_prems_conv i =
wenzelm@37526
   128
      Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv
berghofe@34907
   129
      Conv.rewr_conv Drule.swap_prems_eq
berghofe@34907
   130
wenzelm@43232
   131
fun drop_judgment ctxt = Object_Logic.drop_judgment (Proof_Context.theory_of ctxt);
berghofe@34907
   132
berghofe@34907
   133
fun find_eq ctxt t =
berghofe@34907
   134
  let
berghofe@34907
   135
    val l = length (Logic.strip_params t);
berghofe@34907
   136
    val Hs = Logic.strip_assums_hyp t;
berghofe@34907
   137
    fun find (i, t) =
wenzelm@37526
   138
      (case Induct_Args.dest_def (drop_judgment ctxt t) of
berghofe@34907
   139
        SOME (Bound j, _) => SOME (i, j)
berghofe@34907
   140
      | SOME (_, Bound j) => SOME (i, j)
wenzelm@37526
   141
      | _ => NONE);
berghofe@34907
   142
  in
wenzelm@37526
   143
    (case get_first find (map_index I Hs) of
berghofe@34907
   144
      NONE => NONE
berghofe@34907
   145
    | SOME (0, 0) => NONE
wenzelm@37526
   146
    | SOME (i, j) => SOME (i, l - j - 1, j))
berghofe@34907
   147
  end;
berghofe@34907
   148
wenzelm@37526
   149
fun mk_swap_rrule ctxt ct =
wenzelm@37526
   150
  (case find_eq ctxt (term_of ct) of
berghofe@34907
   151
    NONE => NONE
wenzelm@37526
   152
  | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct));
berghofe@34907
   153
wenzelm@37526
   154
val rearrange_eqs_simproc =
wenzelm@38963
   155
  Simplifier.simproc_global
wenzelm@37526
   156
    (Thm.theory_of_thm Drule.swap_prems_eq) "rearrange_eqs" ["all t"]
wenzelm@37526
   157
    (fn thy => fn ss => fn t =>
wenzelm@37526
   158
      mk_swap_rrule (Simplifier.the_context ss) (cterm_of thy t));
wenzelm@37526
   159
berghofe@34907
   160
berghofe@34907
   161
(* rotate k premises to the left by j, skipping over first j premises *)
berghofe@34907
   162
berghofe@34907
   163
fun rotate_conv 0 j 0 = Conv.all_conv
wenzelm@37526
   164
  | rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1)
wenzelm@37526
   165
  | rotate_conv i j k = Conv.implies_concl_conv (rotate_conv (i - 1) j k);
berghofe@34907
   166
berghofe@34907
   167
fun rotate_tac j 0 = K all_tac
wenzelm@37526
   168
  | rotate_tac j k = SUBGOAL (fn (goal, i) =>
wenzelm@37526
   169
      CONVERSION (rotate_conv
wenzelm@37526
   170
        j (length (Logic.strip_assums_hyp goal) - j - k) k) i);
wenzelm@37526
   171
berghofe@34907
   172
berghofe@34907
   173
(* rulify operators around definition *)
berghofe@34907
   174
berghofe@34907
   175
fun rulify_defs_conv ctxt ct =
wenzelm@37525
   176
  if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) andalso
wenzelm@37525
   177
    not (is_some (Induct_Args.dest_def (drop_judgment ctxt (term_of ct))))
berghofe@34907
   178
  then
berghofe@34907
   179
    (Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv
berghofe@34907
   180
     Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt))
berghofe@34907
   181
       (Conv.try_conv (rulify_defs_conv ctxt)) else_conv
wenzelm@37525
   182
     Conv.first_conv (map Conv.rewr_conv Induct_Args.rulify) then_conv
berghofe@34907
   183
       Conv.try_conv (rulify_defs_conv ctxt)) ct
berghofe@34907
   184
  else Conv.no_conv ct;
berghofe@34907
   185
berghofe@34907
   186
berghofe@34907
   187
wenzelm@24830
   188
(** induct data **)
wenzelm@24830
   189
wenzelm@24830
   190
(* rules *)
wenzelm@24830
   191
wenzelm@30562
   192
type rules = (string * thm) Item_Net.T;
wenzelm@24830
   193
wenzelm@33373
   194
fun init_rules index : rules =
wenzelm@33373
   195
  Item_Net.init
wenzelm@33373
   196
    (fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2))
wenzelm@33373
   197
    (single o index);
wenzelm@24830
   198
wenzelm@27140
   199
fun filter_rules (rs: rules) th =
wenzelm@30562
   200
  filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs);
wenzelm@27140
   201
wenzelm@30562
   202
fun lookup_rule (rs: rules) = AList.lookup (op =) (Item_Net.content rs);
wenzelm@24830
   203
wenzelm@24830
   204
fun pretty_rules ctxt kind rs =
wenzelm@30562
   205
  let val thms = map snd (Item_Net.content rs)
wenzelm@32111
   206
  in Pretty.big_list kind (map (Display.pretty_thm ctxt) thms) end;
wenzelm@24830
   207
wenzelm@24830
   208
wenzelm@24830
   209
(* context data *)
wenzelm@24830
   210
wenzelm@37525
   211
structure Data = Generic_Data
wenzelm@24830
   212
(
berghofe@34907
   213
  type T = (rules * rules) * (rules * rules) * (rules * rules) * simpset;
wenzelm@24830
   214
  val empty =
wenzelm@24830
   215
    ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)),
wenzelm@24830
   216
     (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)),
berghofe@34907
   217
     (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)),
berghofe@34907
   218
     empty_ss addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq]);
wenzelm@24830
   219
  val extend = I;
berghofe@34907
   220
  fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1),
berghofe@34907
   221
      ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) =
wenzelm@30562
   222
    ((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)),
berghofe@34907
   223
     (Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)),
berghofe@34907
   224
     (Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)),
berghofe@34907
   225
     merge_ss (simpset1, simpset2));
wenzelm@24830
   226
);
wenzelm@24830
   227
wenzelm@37525
   228
val get_local = Data.get o Context.Proof;
wenzelm@24830
   229
wenzelm@24830
   230
fun dest_rules ctxt =
berghofe@34907
   231
  let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in
wenzelm@30562
   232
    {type_cases = Item_Net.content casesT,
wenzelm@30562
   233
     pred_cases = Item_Net.content casesP,
wenzelm@30562
   234
     type_induct = Item_Net.content inductT,
wenzelm@30562
   235
     pred_induct = Item_Net.content inductP,
wenzelm@30562
   236
     type_coinduct = Item_Net.content coinductT,
wenzelm@30562
   237
     pred_coinduct = Item_Net.content coinductP}
wenzelm@24830
   238
  end;
wenzelm@24830
   239
wenzelm@24830
   240
fun print_rules ctxt =
berghofe@34907
   241
  let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in
wenzelm@24830
   242
   [pretty_rules ctxt "coinduct type:" coinductT,
wenzelm@24861
   243
    pretty_rules ctxt "coinduct pred:" coinductP,
wenzelm@24830
   244
    pretty_rules ctxt "induct type:" inductT,
wenzelm@24861
   245
    pretty_rules ctxt "induct pred:" inductP,
wenzelm@24830
   246
    pretty_rules ctxt "cases type:" casesT,
wenzelm@24861
   247
    pretty_rules ctxt "cases pred:" casesP]
wenzelm@24830
   248
    |> Pretty.chunks |> Pretty.writeln
wenzelm@24830
   249
  end;
wenzelm@24830
   250
wenzelm@24867
   251
val _ =
wenzelm@36970
   252
  Outer_Syntax.improper_command "print_induct_rules" "print induction and cases rules"
wenzelm@36970
   253
    Keyword.diag (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
wenzelm@24867
   254
      Toplevel.keep (print_rules o Toplevel.context_of)));
wenzelm@24830
   255
wenzelm@24830
   256
wenzelm@24830
   257
(* access rules *)
wenzelm@24830
   258
wenzelm@24830
   259
val lookup_casesT = lookup_rule o #1 o #1 o get_local;
wenzelm@24861
   260
val lookup_casesP = lookup_rule o #2 o #1 o get_local;
wenzelm@24830
   261
val lookup_inductT = lookup_rule o #1 o #2 o get_local;
wenzelm@24861
   262
val lookup_inductP = lookup_rule o #2 o #2 o get_local;
wenzelm@24830
   263
val lookup_coinductT = lookup_rule o #1 o #3 o get_local;
wenzelm@24861
   264
val lookup_coinductP = lookup_rule o #2 o #3 o get_local;
wenzelm@24830
   265
wenzelm@24830
   266
wenzelm@24830
   267
fun find_rules which how ctxt x =
wenzelm@30562
   268
  map snd (Item_Net.retrieve (which (get_local ctxt)) (how x));
wenzelm@24830
   269
wenzelm@37524
   270
val find_casesT = find_rules (#1 o #1) Net.encode_type;
wenzelm@24861
   271
val find_casesP = find_rules (#2 o #1) I;
wenzelm@37524
   272
val find_inductT = find_rules (#1 o #2) Net.encode_type;
wenzelm@24861
   273
val find_inductP = find_rules (#2 o #2) I;
wenzelm@37524
   274
val find_coinductT = find_rules (#1 o #3) Net.encode_type;
wenzelm@24861
   275
val find_coinductP = find_rules (#2 o #3) I;
wenzelm@24830
   276
wenzelm@24830
   277
wenzelm@24830
   278
wenzelm@24830
   279
(** attributes **)
wenzelm@24830
   280
wenzelm@24830
   281
local
wenzelm@24830
   282
wenzelm@24830
   283
fun mk_att f g name arg =
wenzelm@37525
   284
  let val (x, thm) = g arg in (Data.map (f (name, thm)) x, thm) end;
wenzelm@27140
   285
wenzelm@37526
   286
fun del_att which =
wenzelm@37526
   287
  Thm.declaration_attribute (fn th => Data.map (which (pairself (fn rs =>
wenzelm@37526
   288
    fold Item_Net.remove (filter_rules rs th) rs))));
wenzelm@24830
   289
berghofe@34907
   290
fun map1 f (x, y, z, s) = (f x, y, z, s);
berghofe@34907
   291
fun map2 f (x, y, z, s) = (x, f y, z, s);
berghofe@34907
   292
fun map3 f (x, y, z, s) = (x, y, f z, s);
berghofe@34907
   293
fun map4 f (x, y, z, s) = (x, y, z, f s);
wenzelm@24830
   294
wenzelm@33373
   295
fun add_casesT rule x = map1 (apfst (Item_Net.update rule)) x;
wenzelm@33373
   296
fun add_casesP rule x = map1 (apsnd (Item_Net.update rule)) x;
wenzelm@33373
   297
fun add_inductT rule x = map2 (apfst (Item_Net.update rule)) x;
wenzelm@33373
   298
fun add_inductP rule x = map2 (apsnd (Item_Net.update rule)) x;
wenzelm@33373
   299
fun add_coinductT rule x = map3 (apfst (Item_Net.update rule)) x;
wenzelm@33373
   300
fun add_coinductP rule x = map3 (apsnd (Item_Net.update rule)) x;
wenzelm@24830
   301
wenzelm@33368
   302
val consumes0 = Rule_Cases.consumes_default 0;
wenzelm@33368
   303
val consumes1 = Rule_Cases.consumes_default 1;
wenzelm@24830
   304
wenzelm@24830
   305
in
wenzelm@24830
   306
wenzelm@24830
   307
val cases_type = mk_att add_casesT consumes0;
wenzelm@24861
   308
val cases_pred = mk_att add_casesP consumes1;
wenzelm@27140
   309
val cases_del = del_att map1;
wenzelm@27140
   310
wenzelm@24830
   311
val induct_type = mk_att add_inductT consumes0;
wenzelm@24861
   312
val induct_pred = mk_att add_inductP consumes1;
wenzelm@27140
   313
val induct_del = del_att map2;
wenzelm@27140
   314
wenzelm@24830
   315
val coinduct_type = mk_att add_coinductT consumes0;
wenzelm@24861
   316
val coinduct_pred = mk_att add_coinductP consumes1;
wenzelm@27140
   317
val coinduct_del = del_att map3;
wenzelm@24830
   318
wenzelm@37525
   319
fun map_simpset f = Data.map (map4 f);
wenzelm@36602
   320
wenzelm@36602
   321
fun induct_simp f =
wenzelm@36602
   322
  Thm.declaration_attribute (fn thm => fn context =>
wenzelm@37526
   323
      map_simpset
wenzelm@37526
   324
        (Simplifier.with_context (Context.proof_of context) (fn ss => f (ss, [thm]))) context);
wenzelm@36602
   325
wenzelm@36602
   326
val induct_simp_add = induct_simp (op addsimps);
wenzelm@36602
   327
val induct_simp_del = induct_simp (op delsimps);
berghofe@34907
   328
wenzelm@24830
   329
end;
wenzelm@24830
   330
wenzelm@24830
   331
wenzelm@24830
   332
wenzelm@24830
   333
(** attribute syntax **)
wenzelm@24830
   334
berghofe@34907
   335
val no_simpN = "no_simp";
wenzelm@24830
   336
val casesN = "cases";
wenzelm@24830
   337
val inductN = "induct";
wenzelm@24830
   338
val coinductN = "coinduct";
wenzelm@24830
   339
wenzelm@24830
   340
val typeN = "type";
wenzelm@24861
   341
val predN = "pred";
wenzelm@24830
   342
val setN = "set";
wenzelm@24830
   343
wenzelm@24830
   344
local
wenzelm@24830
   345
wenzelm@24830
   346
fun spec k arg =
wenzelm@24830
   347
  Scan.lift (Args.$$$ k -- Args.colon) |-- arg ||
wenzelm@24830
   348
  Scan.lift (Args.$$$ k) >> K "";
wenzelm@24830
   349
wenzelm@30530
   350
fun attrib add_type add_pred del =
wenzelm@35400
   351
  spec typeN (Args.type_name false) >> add_type ||
wenzelm@35400
   352
  spec predN (Args.const false) >> add_pred ||
wenzelm@35400
   353
  spec setN (Args.const false) >> add_pred ||
wenzelm@30530
   354
  Scan.lift Args.del >> K del;
wenzelm@24830
   355
wenzelm@24830
   356
in
wenzelm@24830
   357
wenzelm@30530
   358
val attrib_setup =
wenzelm@30722
   359
  Attrib.setup @{binding cases} (attrib cases_type cases_pred cases_del)
wenzelm@30722
   360
    "declaration of cases rule" #>
wenzelm@30722
   361
  Attrib.setup @{binding induct} (attrib induct_type induct_pred induct_del)
wenzelm@30722
   362
    "declaration of induction rule" #>
wenzelm@30722
   363
  Attrib.setup @{binding coinduct} (attrib coinduct_type coinduct_pred coinduct_del)
berghofe@34907
   364
    "declaration of coinduction rule" #>
wenzelm@36602
   365
  Attrib.setup @{binding induct_simp} (Attrib.add_del induct_simp_add induct_simp_del)
berghofe@34907
   366
    "declaration of rules for simplifying induction or cases rules";
wenzelm@24830
   367
wenzelm@24830
   368
end;
wenzelm@24830
   369
wenzelm@24830
   370
wenzelm@24830
   371
wenzelm@24830
   372
(** method utils **)
wenzelm@24830
   373
wenzelm@24830
   374
(* alignment *)
wenzelm@24830
   375
wenzelm@24830
   376
fun align_left msg xs ys =
wenzelm@24830
   377
  let val m = length xs and n = length ys
haftmann@33956
   378
  in if m < n then error msg else (take n xs ~~ ys) end;
wenzelm@24830
   379
wenzelm@24830
   380
fun align_right msg xs ys =
wenzelm@24830
   381
  let val m = length xs and n = length ys
haftmann@33956
   382
  in if m < n then error msg else (drop (m - n) xs ~~ ys) end;
wenzelm@24830
   383
wenzelm@24830
   384
wenzelm@24830
   385
(* prep_inst *)
wenzelm@24830
   386
wenzelm@32445
   387
fun prep_inst ctxt align tune (tm, ts) =
wenzelm@24830
   388
  let
wenzelm@43232
   389
    val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
wenzelm@24830
   390
    fun prep_var (x, SOME t) =
wenzelm@24830
   391
          let
wenzelm@24830
   392
            val cx = cert x;
wenzelm@26626
   393
            val xT = #T (Thm.rep_cterm cx);
wenzelm@24830
   394
            val ct = cert (tune t);
wenzelm@32445
   395
            val tT = #T (Thm.rep_cterm ct);
wenzelm@24830
   396
          in
wenzelm@32445
   397
            if Type.could_unify (tT, xT) then SOME (cx, ct)
wenzelm@24830
   398
            else error (Pretty.string_of (Pretty.block
wenzelm@24830
   399
             [Pretty.str "Ill-typed instantiation:", Pretty.fbrk,
wenzelm@32445
   400
              Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1,
wenzelm@32445
   401
              Syntax.pretty_typ ctxt tT]))
wenzelm@24830
   402
          end
wenzelm@24830
   403
      | prep_var (_, NONE) = NONE;
wenzelm@24830
   404
    val xs = vars_of tm;
wenzelm@24830
   405
  in
wenzelm@24830
   406
    align "Rule has fewer variables than instantiations given" xs ts
wenzelm@24830
   407
    |> map_filter prep_var
wenzelm@24830
   408
  end;
wenzelm@24830
   409
wenzelm@24830
   410
wenzelm@24830
   411
(* trace_rules *)
wenzelm@24830
   412
wenzelm@24830
   413
fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule")
wenzelm@24830
   414
  | trace_rules ctxt _ rules = Method.trace ctxt rules;
wenzelm@24830
   415
wenzelm@24830
   416
berghofe@34974
   417
(* mark equality constraints in cases rule *)
berghofe@34974
   418
wenzelm@37525
   419
val equal_def' = Thm.symmetric Induct_Args.equal_def;
berghofe@34974
   420
berghofe@34974
   421
fun mark_constraints n ctxt = Conv.fconv_rule
berghofe@34974
   422
  (Conv.prems_conv (~1) (Conv.params_conv ~1 (K (Conv.prems_conv n
wenzelm@41494
   423
     (Raw_Simplifier.rewrite false [equal_def']))) ctxt));
berghofe@34974
   424
berghofe@34974
   425
val unmark_constraints = Conv.fconv_rule
wenzelm@41494
   426
  (Raw_Simplifier.rewrite true [Induct_Args.equal_def]);
berghofe@34974
   427
wenzelm@37526
   428
berghofe@34974
   429
(* simplify *)
berghofe@34974
   430
berghofe@34974
   431
fun simplify_conv' ctxt =
berghofe@34974
   432
  Simplifier.full_rewrite (Simplifier.context ctxt (#4 (get_local ctxt)));
berghofe@34974
   433
berghofe@34974
   434
fun simplify_conv ctxt ct =
wenzelm@37525
   435
  if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) then
berghofe@34974
   436
    (Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct
berghofe@34974
   437
  else Conv.all_conv ct;
berghofe@34974
   438
berghofe@34974
   439
fun gen_simplified_rule cv ctxt =
berghofe@34974
   440
  Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt));
berghofe@34974
   441
berghofe@34974
   442
val simplified_rule' = gen_simplified_rule simplify_conv';
berghofe@34974
   443
val simplified_rule = gen_simplified_rule simplify_conv;
berghofe@34974
   444
berghofe@34974
   445
fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt);
berghofe@34974
   446
wenzelm@37525
   447
val trivial_tac = Induct_Args.trivial_tac;
berghofe@34974
   448
berghofe@34974
   449
wenzelm@24830
   450
wenzelm@24830
   451
(** cases method **)
wenzelm@24830
   452
wenzelm@24830
   453
(*
wenzelm@24830
   454
  rule selection scheme:
wenzelm@24830
   455
          cases         - default case split
wenzelm@24861
   456
    `A t` cases ...     - predicate/set cases
wenzelm@24830
   457
          cases t       - type cases
wenzelm@24830
   458
    ...   cases ... r   - explicit rule
wenzelm@24830
   459
*)
wenzelm@24830
   460
wenzelm@24830
   461
local
wenzelm@24830
   462
wenzelm@24830
   463
fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t)
wenzelm@24830
   464
  | get_casesT _ _ = [];
wenzelm@24830
   465
wenzelm@24861
   466
fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact)
wenzelm@24861
   467
  | get_casesP _ _ = [];
wenzelm@24830
   468
wenzelm@24830
   469
in
wenzelm@24830
   470
berghofe@34974
   471
fun cases_tac ctxt simp insts opt_rule facts =
wenzelm@24830
   472
  let
wenzelm@43232
   473
    val thy = Proof_Context.theory_of ctxt;
wenzelm@24830
   474
wenzelm@24830
   475
    fun inst_rule r =
berghofe@34974
   476
      (if null insts then r
berghofe@34974
   477
       else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts
berghofe@34974
   478
         |> maps (prep_inst ctxt align_left I)
berghofe@34974
   479
         |> Drule.cterm_instantiate) r) |>
berghofe@34974
   480
      (if simp then mark_constraints (Rule_Cases.get_constraints r) ctxt else I) |>
berghofe@34974
   481
      pair (Rule_Cases.get r);
wenzelm@24830
   482
wenzelm@24830
   483
    val ruleq =
wenzelm@24830
   484
      (case opt_rule of
wenzelm@24830
   485
        SOME r => Seq.single (inst_rule r)
wenzelm@24830
   486
      | NONE =>
wenzelm@37525
   487
          (get_casesP ctxt facts @ get_casesT ctxt insts @ [Induct_Args.cases_default])
wenzelm@24830
   488
          |> tap (trace_rules ctxt casesN)
wenzelm@24830
   489
          |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
wenzelm@24830
   490
  in
wenzelm@24830
   491
    fn i => fn st =>
wenzelm@24830
   492
      ruleq
wenzelm@33368
   493
      |> Seq.maps (Rule_Cases.consume [] facts)
wenzelm@24830
   494
      |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
berghofe@34974
   495
        let val rule' =
berghofe@34974
   496
          (if simp then simplified_rule' ctxt #> unmark_constraints else I) rule
berghofe@34974
   497
        in
berghofe@34974
   498
          CASES (Rule_Cases.make_common (thy,
berghofe@34974
   499
              Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
berghofe@34974
   500
            ((Method.insert_tac more_facts THEN' Tactic.rtac rule' THEN_ALL_NEW
berghofe@34974
   501
                (if simp then TRY o trivial_tac else K all_tac)) i) st
berghofe@34974
   502
        end)
wenzelm@24830
   503
  end;
wenzelm@24830
   504
wenzelm@24830
   505
end;
wenzelm@24830
   506
wenzelm@24830
   507
wenzelm@24830
   508
wenzelm@24830
   509
(** induct method **)
wenzelm@24830
   510
wenzelm@24830
   511
val conjunction_congs = [@{thm Pure.all_conjunction}, @{thm imp_conjunction}];
wenzelm@24830
   512
wenzelm@24830
   513
wenzelm@24830
   514
(* atomize *)
wenzelm@24830
   515
wenzelm@24830
   516
fun atomize_term thy =
wenzelm@41494
   517
  Raw_Simplifier.rewrite_term thy Induct_Args.atomize []
wenzelm@35625
   518
  #> Object_Logic.drop_judgment thy;
wenzelm@24830
   519
wenzelm@41494
   520
val atomize_cterm = Raw_Simplifier.rewrite true Induct_Args.atomize;
wenzelm@24830
   521
wenzelm@37525
   522
val atomize_tac = Simplifier.rewrite_goal_tac Induct_Args.atomize;
wenzelm@24830
   523
wenzelm@24830
   524
val inner_atomize_tac =
wenzelm@24830
   525
  Simplifier.rewrite_goal_tac (map Thm.symmetric conjunction_congs) THEN' atomize_tac;
wenzelm@24830
   526
wenzelm@24830
   527
wenzelm@24830
   528
(* rulify *)
wenzelm@24830
   529
wenzelm@24830
   530
fun rulify_term thy =
wenzelm@41494
   531
  Raw_Simplifier.rewrite_term thy (Induct_Args.rulify @ conjunction_congs) [] #>
wenzelm@41494
   532
  Raw_Simplifier.rewrite_term thy Induct_Args.rulify_fallback [];
wenzelm@24830
   533
wenzelm@24830
   534
fun rulified_term thm =
wenzelm@24830
   535
  let
wenzelm@24830
   536
    val thy = Thm.theory_of_thm thm;
wenzelm@24830
   537
    val rulify = rulify_term thy;
wenzelm@24830
   538
    val (As, B) = Logic.strip_horn (Thm.prop_of thm);
wenzelm@24830
   539
  in (thy, Logic.list_implies (map rulify As, rulify B)) end;
wenzelm@24830
   540
wenzelm@24830
   541
val rulify_tac =
wenzelm@37525
   542
  Simplifier.rewrite_goal_tac (Induct_Args.rulify @ conjunction_congs) THEN'
wenzelm@37525
   543
  Simplifier.rewrite_goal_tac Induct_Args.rulify_fallback THEN'
wenzelm@24830
   544
  Goal.conjunction_tac THEN_ALL_NEW
wenzelm@24830
   545
  (Simplifier.rewrite_goal_tac [@{thm Pure.conjunction_imp}] THEN' Goal.norm_hhf_tac);
wenzelm@24830
   546
wenzelm@24830
   547
wenzelm@24830
   548
(* prepare rule *)
wenzelm@24830
   549
wenzelm@32445
   550
fun rule_instance ctxt inst rule =
wenzelm@32445
   551
  Drule.cterm_instantiate (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule;
wenzelm@24830
   552
wenzelm@24830
   553
fun internalize k th =
wenzelm@24830
   554
  th |> Thm.permute_prems 0 k
wenzelm@24830
   555
  |> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) atomize_cterm);
wenzelm@24830
   556
wenzelm@24830
   557
wenzelm@24830
   558
(* guess rule instantiation -- cannot handle pending goal parameters *)
wenzelm@24830
   559
wenzelm@24830
   560
local
wenzelm@24830
   561
wenzelm@32043
   562
fun dest_env thy env =
wenzelm@24830
   563
  let
wenzelm@24830
   564
    val cert = Thm.cterm_of thy;
wenzelm@24830
   565
    val certT = Thm.ctyp_of thy;
wenzelm@32043
   566
    val pairs = Vartab.dest (Envir.term_env env);
wenzelm@32043
   567
    val types = Vartab.dest (Envir.type_env env);
wenzelm@24830
   568
    val ts = map (cert o Envir.norm_term env o #2 o #2) pairs;
wenzelm@24830
   569
    val xs = map2 (curry (cert o Var)) (map #1 pairs) (map (#T o Thm.rep_cterm) ts);
wenzelm@32043
   570
  in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end;
wenzelm@24830
   571
wenzelm@24830
   572
in
wenzelm@24830
   573
wenzelm@26940
   574
fun guess_instance ctxt rule i st =
wenzelm@24830
   575
  let
wenzelm@43232
   576
    val thy = Proof_Context.theory_of ctxt;
wenzelm@26626
   577
    val maxidx = Thm.maxidx_of st;
wenzelm@24830
   578
    val goal = Thm.term_of (Thm.cprem_of st i);  (*exception Subscript*)
wenzelm@29276
   579
    val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal));
wenzelm@24830
   580
  in
wenzelm@24830
   581
    if not (null params) then
wenzelm@24830
   582
      (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^
wenzelm@43156
   583
        commas_quote (map (Syntax.string_of_term ctxt o Syntax_Trans.mark_boundT) params));
wenzelm@24830
   584
      Seq.single rule)
wenzelm@24830
   585
    else
wenzelm@24830
   586
      let
wenzelm@24830
   587
        val rule' = Thm.incr_indexes (maxidx + 1) rule;
wenzelm@24830
   588
        val concl = Logic.strip_assums_concl goal;
wenzelm@24830
   589
      in
wenzelm@32043
   590
        Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule'))
wenzelm@24830
   591
        |> Seq.map (fn env => Drule.instantiate (dest_env thy env) rule')
wenzelm@24830
   592
      end
wenzelm@44158
   593
  end handle General.Subscript => Seq.empty;
wenzelm@24830
   594
wenzelm@24830
   595
end;
wenzelm@24830
   596
wenzelm@24830
   597
wenzelm@24830
   598
(* special renaming of rule parameters *)
wenzelm@24830
   599
wenzelm@24830
   600
fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] =
wenzelm@24830
   601
      let
wenzelm@43360
   602
        val x = Name.clean (Variable.revert_fixed ctxt z);
wenzelm@24830
   603
        fun index i [] = []
wenzelm@24830
   604
          | index i (y :: ys) =
wenzelm@24830
   605
              if x = y then x ^ string_of_int i :: index (i + 1) ys
wenzelm@24830
   606
              else y :: index i ys;
wenzelm@24830
   607
        fun rename_params [] = []
wenzelm@24830
   608
          | rename_params ((y, Type (U, _)) :: ys) =
wenzelm@24830
   609
              (if U = T then x else y) :: rename_params ys
wenzelm@24830
   610
          | rename_params ((y, _) :: ys) = y :: rename_params ys;
wenzelm@24830
   611
        fun rename_asm A =
wenzelm@24830
   612
          let
wenzelm@24830
   613
            val xs = rename_params (Logic.strip_params A);
wenzelm@24830
   614
            val xs' =
wenzelm@28375
   615
              (case filter (fn x' => x' = x) xs of
wenzelm@24830
   616
                [] => xs | [_] => xs | _ => index 1 xs);
wenzelm@24830
   617
          in Logic.list_rename_params (xs', A) end;
wenzelm@24830
   618
        fun rename_prop p =
wenzelm@24830
   619
          let val (As, C) = Logic.strip_horn p
wenzelm@24830
   620
          in Logic.list_implies (map rename_asm As, C) end;
wenzelm@24830
   621
        val cp' = cterm_fun rename_prop (Thm.cprop_of thm);
wenzelm@24830
   622
        val thm' = Thm.equal_elim (Thm.reflexive cp') thm;
wenzelm@33368
   623
      in [Rule_Cases.save thm thm'] end
wenzelm@24830
   624
  | special_rename_params _ _ ths = ths;
wenzelm@24830
   625
wenzelm@24830
   626
wenzelm@24830
   627
(* fix_tac *)
wenzelm@24830
   628
wenzelm@24830
   629
local
wenzelm@24830
   630
wenzelm@24830
   631
fun goal_prefix k ((c as Const ("all", _)) $ Abs (a, T, B)) = c $ Abs (a, T, goal_prefix k B)
wenzelm@24830
   632
  | goal_prefix 0 _ = Term.dummy_pattern propT
wenzelm@24830
   633
  | goal_prefix k ((c as Const ("==>", _)) $ A $ B) = c $ A $ goal_prefix (k - 1) B
wenzelm@24830
   634
  | goal_prefix _ _ = Term.dummy_pattern propT;
wenzelm@24830
   635
wenzelm@24830
   636
fun goal_params k (Const ("all", _) $ Abs (_, _, B)) = goal_params k B + 1
wenzelm@24830
   637
  | goal_params 0 _ = 0
wenzelm@24830
   638
  | goal_params k (Const ("==>", _) $ _ $ B) = goal_params (k - 1) B
wenzelm@24830
   639
  | goal_params _ _ = 0;
wenzelm@24830
   640
wenzelm@24830
   641
fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) =>
wenzelm@24830
   642
  let
wenzelm@43232
   643
    val thy = Proof_Context.theory_of ctxt;
wenzelm@24830
   644
    val cert = Thm.cterm_of thy;
wenzelm@24830
   645
wenzelm@24830
   646
    val v = Free (x, T);
wenzelm@24830
   647
    fun spec_rule prfx (xs, body) =
wenzelm@24830
   648
      @{thm Pure.meta_spec}
wenzelm@43360
   649
      |> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1)
wenzelm@24830
   650
      |> Thm.lift_rule (cert prfx)
wenzelm@24830
   651
      |> `(Thm.prop_of #> Logic.strip_assums_concl)
wenzelm@24830
   652
      |-> (fn pred $ arg =>
wenzelm@24830
   653
        Drule.cterm_instantiate
wenzelm@24830
   654
          [(cert (Term.head_of pred), cert (Logic.rlist_abs (xs, body))),
wenzelm@24830
   655
           (cert (Term.head_of arg), cert (Logic.rlist_abs (xs, v)))]);
wenzelm@24830
   656
wenzelm@24830
   657
    fun goal_concl k xs (Const ("all", _) $ Abs (a, T, B)) = goal_concl k ((a, T) :: xs) B
wenzelm@24830
   658
      | goal_concl 0 xs B =
wenzelm@24830
   659
          if not (Term.exists_subterm (fn t => t aconv v) B) then NONE
wenzelm@24830
   660
          else SOME (xs, Term.absfree (x, T, Term.incr_boundvars 1 B))
wenzelm@24830
   661
      | goal_concl k xs (Const ("==>", _) $ _ $ B) = goal_concl (k - 1) xs B
wenzelm@24830
   662
      | goal_concl _ _ _ = NONE;
wenzelm@24830
   663
  in
wenzelm@24830
   664
    (case goal_concl n [] goal of
wenzelm@24830
   665
      SOME concl =>
wenzelm@24830
   666
        (compose_tac (false, spec_rule (goal_prefix n goal) concl, 1) THEN' rtac asm_rl) i
wenzelm@24830
   667
    | NONE => all_tac)
wenzelm@24830
   668
  end);
wenzelm@24830
   669
wenzelm@24832
   670
fun miniscope_tac p = CONVERSION o
wenzelm@41494
   671
  Conv.params_conv p (K (Raw_Simplifier.rewrite true [Thm.symmetric Drule.norm_hhf_eq]));
wenzelm@24830
   672
wenzelm@24830
   673
in
wenzelm@24830
   674
wenzelm@24830
   675
fun fix_tac _ _ [] = K all_tac
wenzelm@24830
   676
  | fix_tac ctxt n xs = SUBGOAL (fn (goal, i) =>
wenzelm@24830
   677
     (EVERY' (map (meta_spec_tac ctxt n) xs) THEN'
wenzelm@24832
   678
      (miniscope_tac (goal_params n goal) ctxt)) i);
wenzelm@24830
   679
wenzelm@24830
   680
end;
wenzelm@24830
   681
wenzelm@24830
   682
wenzelm@24830
   683
(* add_defs *)
wenzelm@24830
   684
wenzelm@24830
   685
fun add_defs def_insts =
wenzelm@24830
   686
  let
berghofe@34907
   687
    fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt)
berghofe@34907
   688
      | add (SOME (SOME x, (t, _))) ctxt =
wenzelm@28083
   689
          let val ([(lhs, (_, th))], ctxt') =
wenzelm@35624
   690
            Local_Defs.add_defs [((x, NoSyn), (Thm.empty_binding, t))] ctxt
wenzelm@24830
   691
          in ((SOME lhs, [th]), ctxt') end
berghofe@34907
   692
      | add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt)
berghofe@34907
   693
      | add (SOME (NONE, (t, _))) ctxt =
berghofe@34907
   694
          let
berghofe@34907
   695
            val ([s], _) = Name.variants ["x"] (Variable.names_of ctxt);
berghofe@34907
   696
            val ([(lhs, (_, th))], ctxt') =
wenzelm@35624
   697
              Local_Defs.add_defs [((Binding.name s, NoSyn),
berghofe@34907
   698
                (Thm.empty_binding, t))] ctxt
berghofe@34907
   699
          in ((SOME lhs, [th]), ctxt') end
wenzelm@24830
   700
      | add NONE ctxt = ((NONE, []), ctxt);
wenzelm@24830
   701
  in fold_map add def_insts #> apfst (split_list #> apsnd flat) end;
wenzelm@24830
   702
wenzelm@24830
   703
wenzelm@24830
   704
(* induct_tac *)
wenzelm@24830
   705
wenzelm@24830
   706
(*
wenzelm@24830
   707
  rule selection scheme:
wenzelm@24861
   708
    `A x` induct ...     - predicate/set induction
wenzelm@24830
   709
          induct x       - type induction
wenzelm@24830
   710
    ...   induct ... r   - explicit rule
wenzelm@24830
   711
*)
wenzelm@24830
   712
wenzelm@24830
   713
fun get_inductT ctxt insts =
wenzelm@32188
   714
  fold_rev (map_product cons) (insts |> map
wenzelm@27323
   715
      ((fn [] => NONE | ts => List.last ts) #>
wenzelm@27323
   716
        (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #>
wenzelm@27323
   717
        find_inductT ctxt)) [[]]
wenzelm@33368
   718
  |> filter_out (forall Rule_Cases.is_inner_rule);
wenzelm@24830
   719
wenzelm@24861
   720
fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact))
wenzelm@24861
   721
  | get_inductP _ _ = [];
wenzelm@24830
   722
berghofe@34907
   723
fun induct_tac ctxt simp def_insts arbitrary taking opt_rule facts =
wenzelm@24830
   724
  let
wenzelm@43232
   725
    val thy = Proof_Context.theory_of ctxt;
wenzelm@24830
   726
wenzelm@24830
   727
    val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list;
berghofe@34907
   728
    val atomized_defs = map (map (Conv.fconv_rule atomize_cterm)) defs;
wenzelm@24830
   729
wenzelm@24830
   730
    fun inst_rule (concls, r) =
wenzelm@33368
   731
      (if null insts then `Rule_Cases.get r
wenzelm@24830
   732
       else (align_left "Rule has fewer conclusions than arguments given"
wenzelm@24830
   733
          (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts
wenzelm@32445
   734
        |> maps (prep_inst ctxt align_right (atomize_term thy))
wenzelm@33368
   735
        |> Drule.cterm_instantiate) r |> pair (Rule_Cases.get r))
wenzelm@24830
   736
      |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th));
wenzelm@24830
   737
wenzelm@24830
   738
    val ruleq =
wenzelm@24830
   739
      (case opt_rule of
wenzelm@33368
   740
        SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs))
wenzelm@24830
   741
      | NONE =>
wenzelm@24861
   742
          (get_inductP ctxt facts @
wenzelm@24830
   743
            map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts))
wenzelm@33368
   744
          |> map_filter (Rule_Cases.mutual_rule ctxt)
wenzelm@24830
   745
          |> tap (trace_rules ctxt inductN o map #2)
wenzelm@24830
   746
          |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
wenzelm@24830
   747
berghofe@34907
   748
    fun rule_cases ctxt rule =
berghofe@34907
   749
      let val rule' = (if simp then simplified_rule ctxt else I)
berghofe@34907
   750
        (Rule_Cases.internalize_params rule);
berghofe@34907
   751
      in Rule_Cases.make_nested (Thm.prop_of rule') (rulified_term rule') end;
wenzelm@24830
   752
  in
wenzelm@24830
   753
    (fn i => fn st =>
wenzelm@24830
   754
      ruleq
wenzelm@33368
   755
      |> Seq.maps (Rule_Cases.consume (flat defs) facts)
wenzelm@24830
   756
      |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
wenzelm@24830
   757
        (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
wenzelm@24830
   758
          (CONJUNCTS (ALLGOALS
berghofe@34907
   759
            let
berghofe@34907
   760
              val adefs = nth_list atomized_defs (j - 1);
berghofe@34907
   761
              val frees = fold (Term.add_frees o prop_of) adefs [];
berghofe@34907
   762
              val xs = nth_list arbitrary (j - 1);
berghofe@34907
   763
              val k = nth concls (j - 1) + more_consumes
berghofe@34907
   764
            in
berghofe@34907
   765
              Method.insert_tac (more_facts @ adefs) THEN'
berghofe@34907
   766
                (if simp then
berghofe@34907
   767
                   rotate_tac k (length adefs) THEN'
berghofe@34907
   768
                   fix_tac defs_ctxt k
berghofe@34907
   769
                     (List.partition (member op = frees) xs |> op @)
berghofe@34907
   770
                 else
berghofe@34907
   771
                   fix_tac defs_ctxt k xs)
berghofe@34907
   772
             end)
wenzelm@24830
   773
          THEN' inner_atomize_tac) j))
wenzelm@24830
   774
        THEN' atomize_tac) i st |> Seq.maps (fn st' =>
wenzelm@26940
   775
            guess_instance ctxt (internalize more_consumes rule) i st'
wenzelm@32445
   776
            |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
wenzelm@24830
   777
            |> Seq.maps (fn rule' =>
berghofe@34907
   778
              CASES (rule_cases ctxt rule' cases)
wenzelm@24830
   779
                (Tactic.rtac rule' i THEN
wenzelm@43232
   780
                  PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st'))))
berghofe@34907
   781
    THEN_ALL_NEW_CASES
berghofe@34907
   782
      ((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac)
berghofe@34907
   783
        else K all_tac)
berghofe@34907
   784
       THEN_ALL_NEW rulify_tac)
wenzelm@24830
   785
  end;
wenzelm@24830
   786
wenzelm@24830
   787
wenzelm@24830
   788
wenzelm@24830
   789
(** coinduct method **)
wenzelm@24830
   790
wenzelm@24830
   791
(*
wenzelm@24830
   792
  rule selection scheme:
wenzelm@24861
   793
    goal "A x" coinduct ...   - predicate/set coinduction
wenzelm@24830
   794
               coinduct x     - type coinduction
wenzelm@24830
   795
               coinduct ... r - explicit rule
wenzelm@24830
   796
*)
wenzelm@24830
   797
wenzelm@24830
   798
local
wenzelm@24830
   799
wenzelm@24830
   800
fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t)
wenzelm@24830
   801
  | get_coinductT _ _ = [];
wenzelm@24830
   802
wenzelm@24861
   803
fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal);
wenzelm@24861
   804
wenzelm@24861
   805
fun main_prop_of th =
wenzelm@33368
   806
  if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th;
wenzelm@24830
   807
wenzelm@24830
   808
in
wenzelm@24830
   809
wenzelm@26924
   810
fun coinduct_tac ctxt inst taking opt_rule facts =
wenzelm@24830
   811
  let
wenzelm@43232
   812
    val thy = Proof_Context.theory_of ctxt;
wenzelm@24830
   813
wenzelm@24830
   814
    fun inst_rule r =
wenzelm@33368
   815
      if null inst then `Rule_Cases.get r
wenzelm@32445
   816
      else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r
wenzelm@33368
   817
        |> pair (Rule_Cases.get r);
wenzelm@24830
   818
wenzelm@24830
   819
    fun ruleq goal =
wenzelm@24830
   820
      (case opt_rule of
wenzelm@24830
   821
        SOME r => Seq.single (inst_rule r)
wenzelm@24830
   822
      | NONE =>
wenzelm@24861
   823
          (get_coinductP ctxt goal @ get_coinductT ctxt inst)
wenzelm@24830
   824
          |> tap (trace_rules ctxt coinductN)
wenzelm@24830
   825
          |> Seq.of_list |> Seq.maps (Seq.try inst_rule));
wenzelm@24830
   826
  in
wenzelm@24830
   827
    SUBGOAL_CASES (fn (goal, i) => fn st =>
wenzelm@24830
   828
      ruleq goal
wenzelm@33368
   829
      |> Seq.maps (Rule_Cases.consume [] facts)
wenzelm@24830
   830
      |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
wenzelm@26940
   831
        guess_instance ctxt rule i st
wenzelm@32445
   832
        |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
wenzelm@24830
   833
        |> Seq.maps (fn rule' =>
berghofe@34907
   834
          CASES (Rule_Cases.make_common (thy,
berghofe@34907
   835
              Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
wenzelm@24830
   836
            (Method.insert_tac more_facts i THEN Tactic.rtac rule' i) st)))
wenzelm@24830
   837
  end;
wenzelm@24830
   838
wenzelm@24830
   839
end;
wenzelm@24830
   840
wenzelm@24830
   841
wenzelm@24830
   842
wenzelm@24830
   843
(** concrete syntax **)
wenzelm@24830
   844
wenzelm@24830
   845
val arbitraryN = "arbitrary";
wenzelm@24830
   846
val takingN = "taking";
wenzelm@24830
   847
val ruleN = "rule";
wenzelm@24830
   848
wenzelm@24830
   849
local
wenzelm@24830
   850
wenzelm@24830
   851
fun single_rule [rule] = rule
wenzelm@24830
   852
  | single_rule _ = error "Single rule expected";
wenzelm@24830
   853
wenzelm@24830
   854
fun named_rule k arg get =
wenzelm@24830
   855
  Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|--
wenzelm@24830
   856
    (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name =>
wenzelm@24830
   857
      (case get (Context.proof_of context) name of SOME x => x
wenzelm@24830
   858
      | NONE => error ("No rule for " ^ k ^ " " ^ quote name))))));
wenzelm@24830
   859
wenzelm@24861
   860
fun rule get_type get_pred =
wenzelm@35400
   861
  named_rule typeN (Args.type_name false) get_type ||
wenzelm@35400
   862
  named_rule predN (Args.const false) get_pred ||
wenzelm@35400
   863
  named_rule setN (Args.const false) get_pred ||
wenzelm@24830
   864
  Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms;
wenzelm@24830
   865
wenzelm@24861
   866
val cases_rule = rule lookup_casesT lookup_casesP >> single_rule;
wenzelm@24861
   867
val induct_rule = rule lookup_inductT lookup_inductP;
wenzelm@24861
   868
val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule;
wenzelm@24830
   869
wenzelm@24830
   870
val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME;
wenzelm@24830
   871
berghofe@34907
   872
val inst' = Scan.lift (Args.$$$ "_") >> K NONE ||
berghofe@34907
   873
  Args.term >> (SOME o rpair false) ||
berghofe@34907
   874
  Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --|
berghofe@34907
   875
    Scan.lift (Args.$$$ ")");
berghofe@34907
   876
wenzelm@24830
   877
val def_inst =
wenzelm@28083
   878
  ((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME)
berghofe@34907
   879
      -- (Args.term >> rpair false)) >> SOME ||
berghofe@34907
   880
    inst' >> Option.map (pair NONE);
wenzelm@24830
   881
wenzelm@27370
   882
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
wenzelm@27370
   883
  error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
wenzelm@24830
   884
wenzelm@24830
   885
fun unless_more_args scan = Scan.unless (Scan.lift
wenzelm@24830
   886
  ((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN ||
wenzelm@24861
   887
    Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan;
wenzelm@24830
   888
wenzelm@24830
   889
val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |--
wenzelm@36970
   890
  Parse.and_list1' (Scan.repeat (unless_more_args free))) [];
wenzelm@24830
   891
wenzelm@24830
   892
val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |--
wenzelm@24830
   893
  Scan.repeat1 (unless_more_args inst)) [];
wenzelm@24830
   894
wenzelm@24830
   895
in
wenzelm@24830
   896
wenzelm@30722
   897
val cases_setup =
wenzelm@30722
   898
  Method.setup @{binding cases}
berghofe@34974
   899
    (Args.mode no_simpN --
wenzelm@36970
   900
      (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >>
berghofe@34974
   901
      (fn (no_simp, (insts, opt_rule)) => fn ctxt =>
berghofe@34974
   902
        METHOD_CASES (fn facts => Seq.DETERM (HEADGOAL
berghofe@34974
   903
          (cases_tac ctxt (not no_simp) insts opt_rule facts)))))
wenzelm@30722
   904
    "case analysis on types or predicates/sets";
wenzelm@24830
   905
wenzelm@30722
   906
val induct_setup =
wenzelm@30722
   907
  Method.setup @{binding induct}
wenzelm@36970
   908
    (Args.mode no_simpN -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
berghofe@34907
   909
      (arbitrary -- taking -- Scan.option induct_rule)) >>
berghofe@34907
   910
      (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt =>
wenzelm@30722
   911
        RAW_METHOD_CASES (fn facts =>
wenzelm@36970
   912
          Seq.DETERM
wenzelm@36970
   913
            (HEADGOAL (induct_tac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))))
wenzelm@30722
   914
    "induction on types or predicates/sets";
wenzelm@24830
   915
wenzelm@30722
   916
val coinduct_setup =
wenzelm@30722
   917
  Method.setup @{binding coinduct}
wenzelm@30722
   918
    (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
wenzelm@30722
   919
      (fn ((insts, taking), opt_rule) => fn ctxt =>
wenzelm@30722
   920
        RAW_METHOD_CASES (fn facts =>
wenzelm@30722
   921
          Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))))
wenzelm@30722
   922
    "coinduction on types or predicates/sets";
wenzelm@24830
   923
wenzelm@24830
   924
end;
wenzelm@24830
   925
wenzelm@24830
   926
wenzelm@24830
   927
wenzelm@24830
   928
(** theory setup **)
wenzelm@24830
   929
wenzelm@30722
   930
val setup = attrib_setup #> cases_setup  #> induct_setup #> coinduct_setup;
wenzelm@24830
   931
wenzelm@24830
   932
end;