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