src/HOL/Nominal/nominal_induct.ML
author wenzelm
Sat, 17 Jun 2006 19:37:43 +0200
changeset 19903 158ea5884966
parent 19115 bc8da9b4a81c
child 20072 c4710df2c953
permissions -rw-r--r--
RuleCases.mutual_rule: ctxt; Term.internal; ProofContext.exports: simultaneous facts;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     1
(*  ID:         $Id$
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
     2
    Author:     Christian Urban and Makarius
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     3
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
     4
The nominal induct proof method.
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     5
*)
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     6
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     7
structure NominalInduct:
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
     8
sig
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
     9
  val nominal_induct_tac: Proof.context -> (string option * term) option list list ->
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    10
    (string * typ) list -> (string * typ) list list -> thm list ->
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    11
    thm list -> int -> RuleCases.cases_tactic
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    12
  val nominal_induct_method: Method.src -> Proof.context -> Method.method
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    13
end =
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    14
struct
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    15
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    16
(* proper tuples -- nested left *)
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    17
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    18
fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts;
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    19
fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts;
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    20
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    21
fun tuple_fun Ts (xi, T) =
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    22
  Library.funpow (length Ts) HOLogic.mk_split
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    23
    (Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T));
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    24
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    25
val split_all_tuples =
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    26
  Simplifier.full_simplify (HOL_basic_ss addsimps
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    27
    [split_conv, split_paired_all, unit_all_eq1, thm "fresh_unit_elim", thm "fresh_prod_elim"]);
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    28
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    29
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    30
(* prepare rule *)
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    31
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    32
(*conclusions: ?P avoiding_struct ... insts*)
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
    33
fun inst_mutual_rule ctxt insts avoiding rules =
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    34
  let
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    35
    val (concls, rule) =
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
    36
      (case RuleCases.mutual_rule ctxt rules of
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    37
        NONE => error "Failed to join given rules into one mutual rule"
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    38
      | SOME res => res);
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    39
    val (cases, consumes) = RuleCases.get rule;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    40
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    41
    val l = length rules;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    42
    val _ =
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    43
      if length insts = l then ()
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    44
      else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules");
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    45
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    46
    fun subst inst rule =
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    47
      let
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    48
        val vars = InductAttrib.vars_of (Thm.concl_of rule);
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    49
        val m = length vars and n = length inst;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    50
        val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule";
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    51
        val P :: x :: ys = vars;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    52
        val zs = Library.drop (m - n - 2, ys);
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    53
      in
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    54
        (P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) ::
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    55
        (x, tuple (map Free avoiding)) ::
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    56
        List.mapPartial (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst)
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    57
      end;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    58
     val substs =
19115
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
    59
       map2 subst insts rules |> List.concat |> distinct (op =)
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
    60
       |> map (pairself (Thm.cterm_of (ProofContext.theory_of ctxt)));
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    61
  in (((cases, concls), consumes), Drule.cterm_instantiate substs rule) end;
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    62
18299
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    63
fun rename_params_rule internal xs rule =
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    64
  let
18299
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    65
    val tune =
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
    66
      if internal then Term.internal
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
    67
      else fn x => the_default x (try Term.dest_internal x);
18299
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    68
    val n = length xs;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    69
    fun rename prem =
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    70
      let
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    71
        val ps = Logic.strip_params prem;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    72
        val p = length ps;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    73
        val ys =
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    74
          if p < n then []
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    75
          else map (tune o #1) (Library.take (p - n, ps)) @ xs;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    76
      in Logic.list_rename_params (ys, prem) end;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    77
    fun rename_prems prop =
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    78
      let val (As, C) = Logic.strip_horn (Thm.prop_of rule)
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    79
      in Logic.list_implies (map rename As, C) end;
af72dfc4b9f9 added rename_params_rule: recover orginal fresh names in subgoals/cases;
wenzelm
parents: 18297
diff changeset
    80
  in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end;
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    81
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    82
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
    83
(* nominal_induct_tac *)
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    84
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    85
fun nominal_induct_tac ctxt def_insts avoiding fixings rules facts =
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    86
  let
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    87
    val thy = ProofContext.theory_of ctxt;
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    88
    val cert = Thm.cterm_of thy;
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    89
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    90
    val ((insts, defs), defs_ctxt) = fold_map InductMethod.add_defs def_insts ctxt |>> split_list;
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    91
    val atomized_defs = map (map ObjectLogic.atomize_thm) defs;
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    92
19115
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
    93
    val finish_rule =
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    94
      split_all_tuples
19115
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
    95
      #> rename_params_rule true (map (ProofContext.revert_skolem defs_ctxt o fst) avoiding);
18610
05a5e950d5f1 RuleCases.make_nested;
wenzelm
parents: 18583
diff changeset
    96
    fun rule_cases r = RuleCases.make_nested true (Thm.prop_of r) (InductMethod.rulified_term r);
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
    97
  in
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
    98
    (fn i => fn st =>
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
    99
      rules
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
   100
      |> inst_mutual_rule ctxt insts avoiding
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   101
      |> RuleCases.consume (List.concat defs) facts
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   102
      |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   103
        (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   104
          (CONJUNCTS (ALLGOALS
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   105
            (Method.insert_tac (more_facts @ nth_list atomized_defs (j - 1))
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   106
              THEN' InductMethod.fix_tac defs_ctxt
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   107
                (nth concls (j - 1) + more_consumes)
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   108
                (nth_list fixings (j - 1))))
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   109
          THEN' InductMethod.inner_atomize_tac) j))
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   110
        THEN' InductMethod.atomize_tac) i st |> Seq.maps (fn st' =>
19115
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
   111
            InductMethod.guess_instance
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
   112
              (finish_rule (InductMethod.internalize more_consumes rule)) i st'
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   113
            |> Seq.maps (fn rule' =>
19115
bc8da9b4a81c distinct (op =);
wenzelm
parents: 19036
diff changeset
   114
              CASES (rule_cases rule' cases)
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   115
                (Tactic.rtac (rename_params_rule false [] rule') i THEN
19903
158ea5884966 RuleCases.mutual_rule: ctxt;
wenzelm
parents: 19115
diff changeset
   116
                  PRIMSEQ (Seq.singleton (ProofContext.exports defs_ctxt ctxt))) st'))))
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   117
    THEN_ALL_NEW_CASES InductMethod.rulify_tac
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   118
  end;
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   119
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   120
18288
feb79a6b274b proper treatment of tuple/tuple_fun -- nest to the left!
wenzelm
parents: 18283
diff changeset
   121
(* concrete syntax *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   122
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   123
local
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   124
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   125
val avoidingN = "avoiding";
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   126
val fixingN = "fixing";
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   127
val ruleN = "rule";
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   128
19036
73782d21e855 Adapted to Context.generic syntax.
berghofe
parents: 18610
diff changeset
   129
val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   130
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   131
val def_inst =
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   132
  ((Scan.lift (Args.name --| (Args.$$$ "\\<equiv>" || Args.$$$ "==")) >> SOME)
19036
73782d21e855 Adapted to Context.generic syntax.
berghofe
parents: 18610
diff changeset
   133
      -- Args.term) >> SOME ||
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   134
    inst >> Option.map (pair NONE);
18099
e956b04fea22 fixed bug with nominal induct
urbanc
parents: 18052
diff changeset
   135
19036
73782d21e855 Adapted to Context.generic syntax.
berghofe
parents: 18610
diff changeset
   136
val free = Scan.state -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
73782d21e855 Adapted to Context.generic syntax.
berghofe
parents: 18610
diff changeset
   137
  error ("Bad free variable: " ^ ProofContext.string_of_term (Context.proof_of ctxt) t));
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   138
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   139
fun unless_more_args scan = Scan.unless (Scan.lift
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   140
  ((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan;
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   141
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   142
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   143
val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |--
18297
116fe71fad51 fresh: frees instead of terms, rename corresponding params in rule;
wenzelm
parents: 18288
diff changeset
   144
  Scan.repeat (unless_more_args free)) [];
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   145
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   146
val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |--
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   147
  Args.and_list (Scan.repeat (unless_more_args free))) [];
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   148
19036
73782d21e855 Adapted to Context.generic syntax.
berghofe
parents: 18610
diff changeset
   149
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   150
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   151
in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   152
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   153
fun nominal_induct_method src =
18583
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   154
  Method.syntax
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   155
   (Args.and_list (Scan.repeat (unless_more_args def_inst)) --
96e1ef2f806f proper handling of simultaneous goals and mutual rules;
wenzelm
parents: 18311
diff changeset
   156
    avoiding -- fixing -- rule_spec) src
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   157
  #> (fn (ctxt, (((x, y), z), w)) =>
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   158
    Method.RAW_METHOD_CASES (fn facts =>
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   159
      HEADGOAL (nominal_induct_tac ctxt x y z w facts)));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   160
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   161
end;
18283
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   162
f8a49f09a202 reworked version with proper support for defs, fixes, fresh specification;
wenzelm
parents: 18265
diff changeset
   163
end;