src/HOL/Tools/coinduction.ML
author traytel
Fri, 04 Mar 2016 17:50:22 +0100
changeset 63709 d68d10fdec78
parent 61844 007d3b34080f
child 63715 ad2c003782f9
permissions -rw-r--r--
coinduction method accepts a list of coinduction rules (takes the first matching one)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54540
5d7006e9205e moved 'coinduction' proof method to 'HOL'
blanchet
parents: 54366
diff changeset
     1
(*  Title:      HOL/Tools/coinduction.ML
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     2
    Author:     Johannes Hölzl, TU Muenchen
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     3
    Author:     Dmitriy Traytel, TU Muenchen
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     4
    Copyright   2013
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     5
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     6
Coinduction method that avoids some boilerplate compared to coinduct.
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     7
*)
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     8
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
     9
signature COINDUCTION =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    10
sig
63709
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    11
  val coinduction_context_tactic: term list -> thm list option -> thm list -> int -> context_tactic
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    12
  val coinduction_tac: Proof.context -> term list -> thm list option -> thm list -> int -> tactic
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    13
end;
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    14
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    15
structure Coinduction : COINDUCTION =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    16
struct
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    17
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    18
fun filter_in_out _ [] = ([], [])
58814
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    19
  | filter_in_out P (x :: xs) =
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    20
      let
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    21
        val (ins, outs) = filter_in_out P xs;
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    22
      in
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    23
        if P x then (x :: ins, outs) else (ins, x :: outs)
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
    24
      end;
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    25
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    26
fun ALLGOALS_SKIP skip tac st =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    27
  let fun doall n = if n = skip then all_tac else tac n THEN doall (n - 1)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    28
  in doall (Thm.nprems_of st) st  end;
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    29
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    30
fun THEN_ALL_NEW_SKIP skip tac1 tac2 i st =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    31
  st |> (tac1 i THEN (fn st' =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    32
    Seq.INTERVAL tac2 (i + skip) (i + Thm.nprems_of st' - Thm.nprems_of st) st'));
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    33
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    34
fun DELETE_PREMS_AFTER skip tac i st =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    35
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    36
    val n = nth (Thm.prems_of st) (i - 1) |> Logic.strip_assums_hyp |> length;
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    37
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59158
diff changeset
    38
    (THEN_ALL_NEW_SKIP skip tac (REPEAT_DETERM_N n o eresolve0_tac [thin_rl])) i st
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    39
  end;
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
    40
63709
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    41
fun coinduction_context_tactic raw_vars opt_raw_thms prems =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    42
  CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    43
    let
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    44
      val lhs_of_eq = HOLogic.dest_Trueprop #> HOLogic.dest_eq #> fst;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    45
      fun find_coinduct t =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    46
        Induct.find_coinductP ctxt t @
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    47
        (try (Induct.find_coinductT ctxt o fastype_of o lhs_of_eq) t |> the_default []);
63709
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    48
      val raw_thms =
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    49
        (case opt_raw_thms of
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    50
          SOME raw_thms => raw_thms
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    51
        | NONE => goal |> Logic.strip_assums_concl |> find_coinduct);
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    52
      val thy = Proof_Context.theory_of ctxt;
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    53
      val concl = Variable.focus NONE goal ctxt |> fst |> snd |> Logic.strip_imp_concl |> @{print};
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    54
      val raw_thm = (case find_first (fn thm =>
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    55
            Pattern.matches thy (Thm.concl_of thm, concl)) raw_thms of
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    56
          SOME thm => thm
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
    57
        | NONE => error "No matching coinduction rule found")
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    58
      val skip = Integer.max 1 (Rule_Cases.get_consumes raw_thm) - 1;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    59
      val cases = Rule_Cases.get raw_thm |> fst;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    60
    in
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    61
      ((Object_Logic.rulify_tac ctxt THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    62
        Method.insert_tac ctxt prems THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    63
        Object_Logic.atomize_prems_tac ctxt THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    64
        DELETE_PREMS_AFTER skip (Subgoal.FOCUS (fn {concl, context = ctxt, params, prems, ...} =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    65
          let
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    66
            val vars = raw_vars @ map (Thm.term_of o snd) params;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    67
            val names_ctxt = ctxt
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    68
              |> fold Variable.declare_names vars
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    69
              |> fold Variable.declare_thm (raw_thm :: prems);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    70
            val thm_concl = Thm.cprop_of raw_thm |> strip_imp_concl;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    71
            val (instT, inst) = Thm.match (thm_concl, concl);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    72
            val rhoTs = map (fn (v, T) => (TVar v, Thm.typ_of T)) instT;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    73
            val rhots = map (fn (v, t) => (Var v, Thm.term_of t)) inst;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    74
            val xs = hd (Thm.prems_of raw_thm) |> HOLogic.dest_Trueprop |> strip_comb |> snd
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    75
              |> map (subst_atomic_types rhoTs);
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    76
            val raw_eqs = map (fn x => (x, AList.lookup op aconv rhots x |> the)) xs;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    77
            val ((names, ctxt), Ts) = map_split (apfst fst o dest_Var o fst) raw_eqs
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    78
              |>> (fn names => Variable.variant_fixes names names_ctxt) ;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    79
            val eqs =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    80
              @{map 3} (fn name => fn T => fn (_, rhs) =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    81
                HOLogic.mk_eq (Free (name, T), rhs))
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    82
              names Ts raw_eqs;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    83
            val phi = eqs @ map (HOLogic.dest_Trueprop o Thm.prop_of) prems
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    84
              |> try (Library.foldr1 HOLogic.mk_conj)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    85
              |> the_default @{term True}
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    86
              |> Ctr_Sugar_Util.list_exists_free vars
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    87
              |> Term.map_abs_vars (Variable.revert_fixed ctxt)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    88
              |> fold_rev Term.absfree (names ~~ Ts)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    89
              |> Thm.cterm_of ctxt;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    90
            val thm = infer_instantiate' ctxt [SOME phi] raw_thm;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    91
            val e = length eqs;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    92
            val p = length prems;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    93
          in
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    94
            HEADGOAL (EVERY' [resolve_tac ctxt [thm],
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    95
              EVERY' (map (fn var =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    96
                resolve_tac ctxt
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    97
                  [infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt var)] exI]) vars),
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    98
              if p = 0 then Ctr_Sugar_Util.CONJ_WRAP' (K (resolve_tac ctxt [refl])) eqs
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
    99
              else
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   100
                REPEAT_DETERM_N e o (resolve_tac ctxt [conjI] THEN' resolve_tac ctxt [refl]) THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   101
                Ctr_Sugar_Util.CONJ_WRAP' (resolve_tac ctxt o single) prems,
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   102
              K (ALLGOALS_SKIP skip
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   103
                 (REPEAT_DETERM_N (length vars) o (eresolve_tac ctxt [exE] THEN' rotate_tac ~1) THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   104
                 DELETE_PREMS_AFTER 0 (Subgoal.FOCUS (fn {prems, params, context = ctxt, ...} =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   105
                   (case prems of
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   106
                     [] => all_tac
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   107
                   | inv :: case_prems =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   108
                       let
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   109
                         val (init, last) = funpow_yield (p + e - 1) (HOLogic.conj_elim ctxt) inv;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   110
                         val inv_thms = init @ [last];
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   111
                         val eqs = take e inv_thms;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   112
                         fun is_local_var t =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   113
                           member (fn (t, (_, t')) => t aconv (Thm.term_of t')) params t;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   114
                          val (eqs, assms') =
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   115
                            filter_in_out (is_local_var o lhs_of_eq o Thm.prop_of) eqs;
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   116
                          val assms = assms' @ drop e inv_thms
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   117
                        in
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   118
                          HEADGOAL (Method.insert_tac ctxt (assms @ case_prems)) THEN
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   119
                          Ctr_Sugar_General_Tactics.unfold_thms_tac ctxt eqs
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   120
                        end)) ctxt)))])
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   121
          end) ctxt) THEN'
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   122
        K (prune_params_tac ctxt)) i) st
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   123
      |> Seq.maps (fn st' =>
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   124
        CONTEXT_CASES (Rule_Cases.make_common ctxt (Thm.prop_of st') cases) all_tac (ctxt, st'))
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60784
diff changeset
   125
    end);
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   126
61844
007d3b34080f tuned signature;
wenzelm
parents: 61841
diff changeset
   127
fun coinduction_tac ctxt x1 x2 x3 x4 =
007d3b34080f tuned signature;
wenzelm
parents: 61841
diff changeset
   128
  Method.NO_CONTEXT_TACTIC ctxt (coinduction_context_tactic x1 x2 x3 x4);
007d3b34080f tuned signature;
wenzelm
parents: 61841
diff changeset
   129
007d3b34080f tuned signature;
wenzelm
parents: 61841
diff changeset
   130
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   131
local
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   132
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   133
val ruleN = "rule"
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   134
val arbitraryN = "arbitrary"
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   135
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   136
fun named_rule k arg get =
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   137
  Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|--
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   138
    (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name =>
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   139
      (case get (Context.proof_of context) name of SOME x => x
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   140
      | NONE => error ("No rule for " ^ k ^ " " ^ quote name))))));
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   141
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   142
fun rule get_type get_pred =
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 54742
diff changeset
   143
  named_rule Induct.typeN (Args.type_name {proper = false, strict = false}) get_type ||
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   144
  named_rule Induct.predN (Args.const {proper = false, strict = false}) get_pred ||
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   145
  named_rule Induct.setN (Args.const {proper = false, strict = false}) get_pred ||
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   146
  Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms;
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   147
63709
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
   148
val coinduct_rules =
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
   149
  rule Induct.lookup_coinductT Induct.lookup_coinductP;
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   150
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   151
fun unless_more_args scan = Scan.unless (Scan.lift
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   152
  ((Args.$$$ arbitraryN || Args.$$$ Induct.typeN ||
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   153
    Args.$$$ Induct.predN || Args.$$$ Induct.setN || Args.$$$ ruleN) -- Args.colon)) scan;
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   154
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   155
val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |--
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   156
  Scan.repeat1 (unless_more_args Args.term)) [];
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   157
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   158
in
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   159
58814
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
   160
val _ =
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
   161
  Theory.setup
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
   162
    (Method.setup @{binding coinduction}
63709
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
   163
      (arbitrary -- Scan.option coinduct_rules >>
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
   164
        (fn (arbitrary, opt_rules) => fn _ => fn facts =>
d68d10fdec78 coinduction method accepts a list of coinduction rules (takes the first matching one)
traytel
parents: 61844
diff changeset
   165
          Seq.DETERM (coinduction_context_tactic arbitrary opt_rules facts 1)))
58814
4c0ad4162cb7 modernized setup;
wenzelm
parents: 58634
diff changeset
   166
      "coinduction on types or predicates/sets");
54026
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   167
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   168
end;
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   169
82d9b2701a03 new coinduction method
traytel
parents:
diff changeset
   170
end;