Nominal/nominal_induct.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 30 Dec 2010 10:00:09 +0000
changeset 2631 e73bd379e839
child 2632 e8732350a29f
permissions -rw-r--r--
removed local fix for bug in induction_schema; added setup method for strong inductions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2631
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Author:     Christian Urban and Makarius
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
The nominal induct proof method.
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
*)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
structure NominalInduct:
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
sig
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  val nominal_induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list ->
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
    (string * typ) list -> (string * typ) list list -> thm list -> thm list -> int -> Rule_Cases.cases_tactic
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  val nominal_induct_method: (Proof.context -> Proof.method) context_parser
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
end =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
struct
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
(* proper tuples -- nested left *)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
fun tupleT Ts = HOLogic.unitT |> fold (fn T => fn U => HOLogic.mk_prodT (U, T)) Ts;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
fun tuple ts = HOLogic.unit |> fold (fn t => fn u => HOLogic.mk_prod (u, t)) ts;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
fun tuple_fun Ts (xi, T) =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  Library.funpow (length Ts) HOLogic.mk_split
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
    (Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T));
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
val split_all_tuples =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  Simplifier.full_simplify (HOL_basic_ss addsimps
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
    @{thms split_conv split_paired_all unit_all_eq1})
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
(* 
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
     @{thm fresh_unit_elim}, @{thm fresh_prod_elim}] @
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
     @{thms fresh_star_unit_elim} @ @{thms fresh_star_prod_elim})
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
*)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
(* prepare rule *)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
fun inst_mutual_rule ctxt insts avoiding rules =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
    val (nconcls, joined_rule) = Rule_Cases.strict_mutual_rule ctxt rules;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
    val concls = Logic.dest_conjunctions (Thm.concl_of joined_rule);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
    val (cases, consumes) = Rule_Cases.get joined_rule;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
    val l = length rules;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
    val _ =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
      if length insts = l then ()
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
      else error ("Bad number of instantiations for " ^ string_of_int l ^ " rules");
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
    fun subst inst concl =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
        val vars = Induct.vars_of concl;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
        val m = length vars and n = length inst;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
        val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule";
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
        val P :: x :: ys = vars;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
        val zs = drop (m - n - 2) ys;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
      in
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
        (P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) ::
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
        (x, tuple (map Free avoiding)) ::
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
        map_filter (fn (z, SOME t) => SOME (z, t) | _ => NONE) (zs ~~ inst)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
      end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
     val substs =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
       map2 subst insts concls |> flat |> distinct (op =)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
       |> map (pairself (Thm.cterm_of (ProofContext.theory_of ctxt)));
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  in 
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
    (((cases, nconcls), consumes), Drule.cterm_instantiate substs joined_rule) 
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
fun rename_params_rule internal xs rule =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
    val tune =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
      if internal then Name.internal
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
      else fn x => the_default x (try Name.dest_internal x);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
    val n = length xs;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
    fun rename prem =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
      let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
        val ps = Logic.strip_params prem;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
        val p = length ps;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
        val ys =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
          if p < n then []
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
          else map (tune o #1) (take (p - n) ps) @ xs;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
      in Logic.list_rename_params (ys, prem) end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
    fun rename_prems prop =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
      let val (As, C) = Logic.strip_horn prop
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
      in Logic.list_implies (map rename As, C) end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
(* nominal_induct_tac *)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
fun nominal_induct_tac ctxt simp def_insts avoiding fixings rules facts =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
    val thy = ProofContext.theory_of ctxt;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
    val cert = Thm.cterm_of thy;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
    val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
    val atomized_defs = map (map (Conv.fconv_rule Induct.atomize_cterm)) defs;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
    val finish_rule =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
      split_all_tuples
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
      #> rename_params_rule true
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
        (map (Name.clean o ProofContext.revert_skolem defs_ctxt o fst) avoiding);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
    fun rule_cases ctxt r =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
      let val r' = if simp then Induct.simplified_rule ctxt r else r
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
      in Rule_Cases.make_nested (Thm.prop_of r') (Induct.rulified_term r') end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  in
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
    (fn i => fn st =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
      rules
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
      |> inst_mutual_rule ctxt insts avoiding
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
      |> Rule_Cases.consume (flat defs) facts
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
      |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
        (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
          (CONJUNCTS (ALLGOALS
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
            let
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
              val adefs = nth_list atomized_defs (j - 1);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
              val frees = fold (Term.add_frees o prop_of) adefs [];
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
              val xs = nth_list fixings (j - 1);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
              val k = nth concls (j - 1) + more_consumes
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
            in
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
              Method.insert_tac (more_facts @ adefs) THEN'
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
                (if simp then
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
                   Induct.rotate_tac k (length adefs) THEN'
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
                   Induct.fix_tac defs_ctxt k
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
                     (List.partition (member op = frees) xs |> op @)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
                 else
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
                   Induct.fix_tac defs_ctxt k xs)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
            end)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
          THEN' Induct.inner_atomize_tac) j))
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
        THEN' Induct.atomize_tac) i st |> Seq.maps (fn st' =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
            Induct.guess_instance ctxt
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
              (finish_rule (Induct.internalize more_consumes rule)) i st'
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
            |> Seq.maps (fn rule' =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
              CASES (rule_cases ctxt rule' cases)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
                (Tactic.rtac (rename_params_rule false [] rule') i THEN
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
                  PRIMITIVE (singleton (ProofContext.export defs_ctxt ctxt))) st'))))
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
    THEN_ALL_NEW_CASES
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
      ((if simp then Induct.simplify_tac ctxt THEN' (TRY o Induct.trivial_tac)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
        else K all_tac)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
       THEN_ALL_NEW Induct.rulify_tac)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  end;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
(* concrete syntax *)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
local
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
val avoidingN = "avoiding";
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
val fixingN = "arbitrary";  (* to be consistent with induct; hopefully this changes again *)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
val ruleN = "rule";
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
val inst = Scan.lift (Args.$$$ "_") >> K NONE ||
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  Args.term >> (SOME o rpair false) ||
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --|
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    Scan.lift (Args.$$$ ")");
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
val def_inst =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  ((Scan.lift (Args.binding --| (Args.$$$ "\<equiv>" || Args.$$$ "==")) >> SOME)
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
      -- (Args.term >> rpair false)) >> SOME ||
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
    inst >> Option.map (pair NONE);
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
fun unless_more_args scan = Scan.unless (Scan.lift
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
  ((Args.$$$ avoidingN || Args.$$$ fixingN || Args.$$$ ruleN) -- Args.colon)) scan;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
val avoiding = Scan.optional (Scan.lift (Args.$$$ avoidingN -- Args.colon) |--
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  Scan.repeat (unless_more_args free)) [];
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
val fixing = Scan.optional (Scan.lift (Args.$$$ fixingN -- Args.colon) |--
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  Parse.and_list' (Scan.repeat (unless_more_args free))) [];
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms;
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
in
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
val nominal_induct_method =
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  Args.mode Induct.no_simpN -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  avoiding -- fixing -- rule_spec) >>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  (fn (no_simp, (((x, y), z), w)) => fn ctxt =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
    RAW_METHOD_CASES (fn facts =>
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
      HEADGOAL (nominal_induct_tac ctxt (not no_simp) x y z w facts)));
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
end
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
e73bd379e839 removed local fix for bug in induction_schema; added setup method for strong inductions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
end;