Nominal/nominal_induct.ML
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sat, 19 Mar 2016 21:06:48 +0000
branchNominal2-Isabelle2016
changeset 3243 c4f31f1564b7
parent 3239 67370521c09c
permissions -rw-r--r--
updated to Isabelle 2016
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
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
     8
  val nominal_induct_tac: bool -> (binding option * (term * bool)) option list list ->
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
     9
    (string * typ) list -> (string * typ) list list -> thm list -> thm list -> int -> context_tactic
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
    10
  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
    11
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
    12
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
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
    14
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
(* 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
    16
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
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
    18
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
    19
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
fun tuple_fun Ts (xi, T) =
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    21
  Library.funpow (length Ts) HOLogic.mk_case_prod
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
    22
    (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
    23
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3046
diff changeset
    24
fun split_all_tuples ctxt =
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3046
diff changeset
    25
  Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps
2632
e8732350a29f added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents: 2631
diff changeset
    26
    @{thms split_conv split_paired_all unit_all_eq1} @
e8732350a29f added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents: 2631
diff changeset
    27
    @{thms fresh_Unit_elim fresh_Pair_elim} @
e8732350a29f added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents: 2631
diff changeset
    28
    @{thms fresh_star_Unit_elim fresh_star_Pair_elim fresh_star_Un_elim} @
e8732350a29f added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents: 2631
diff changeset
    29
    @{thms fresh_star_insert_elim fresh_star_empty_elim})
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
    30
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
(* 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
    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
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
    35
  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
    36
    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
    37
    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
    38
    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
    39
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 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
    41
    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
    42
      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
    43
      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
    44
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
    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
    46
      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
    47
        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
    48
        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
    49
        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
    50
        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
    51
        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
    52
      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
    53
        (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
    54
        (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
    55
        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
    56
      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
    57
     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
    58
       map2 subst insts concls |> flat |> distinct (op =)
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    59
       |> map (fn (t, u) => (#1 (dest_Var t), Thm.cterm_of ctxt u));
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
    60
  in 
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    61
    (((cases, nconcls), consumes), infer_instantiate ctxt substs joined_rule) 
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
    62
  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
    63
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
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
    65
  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
    66
    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
    67
      if internal then Name.internal
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
    68
      else perhaps (try Name.dest_internal);
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
    69
    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
    70
    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
    71
      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
    72
        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
    73
        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
    74
        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
    75
          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
    76
          else map (tune o #1) (take (p - n) ps) @ xs;
3046
9b0324e1293b all examples work again after quotient package has been "de-localised"
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
    77
      in Logic.list_rename_params ys prem end;
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
    78
    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
    79
      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
    80
      in Logic.list_implies (map rename As, C) end;
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    81
  in Thm.renamed_prop (rename_prems (Thm.prop_of rule)) rule end;
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
    82
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
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
(* 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
    85
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    86
fun nominal_induct_tac simp def_insts avoiding fixings rules facts i (ctxt, st) =
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
    87
  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
    88
    val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list;
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3222
diff changeset
    89
    val atomized_defs = map (map (Conv.fconv_rule (Induct.atomize_cterm defs_ctxt))) defs;
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
    90
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 finish_rule =
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3046
diff changeset
    92
      split_all_tuples defs_ctxt
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
    93
      #> rename_params_rule true
2781
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2632
diff changeset
    94
        (map (Name.clean o Variable.revert_fixed defs_ctxt o fst) avoiding);
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
    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
    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
    97
      let val r' = if simp then Induct.simplified_rule ctxt r else r
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
    98
      in Rule_Cases.make_nested ctxt (Thm.prop_of r') (Induct.rulified_term ctxt r') end;
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    99
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   100
    fun context_tac _ _ =
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
   101
      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
   102
      |> inst_mutual_rule ctxt insts avoiding
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3222
diff changeset
   103
      |> Rule_Cases.consume ctxt (flat defs) facts
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
   104
      |> 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
   105
        (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
   106
          (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
   107
            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
   108
              val adefs = nth_list atomized_defs (j - 1);
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   109
              val frees = fold (Term.add_frees o Thm.prop_of) adefs [];
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
   110
              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
   111
              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
   112
            in
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   113
              Method.insert_tac ctxt (more_facts @ adefs) THEN'
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
   114
                (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
   115
                   Induct.rotate_tac k (length adefs) THEN'
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
   116
                   Induct.arbitrary_tac defs_ctxt k
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
   117
                     (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
   118
                 else
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
   119
                   Induct.arbitrary_tac defs_ctxt k xs)
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
   120
            end)
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3222
diff changeset
   121
          THEN' Induct.inner_atomize_tac ctxt) j))
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3222
diff changeset
   122
        THEN' Induct.atomize_tac ctxt) i st |> Seq.maps (fn st' =>
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
   123
            Induct.guess_instance ctxt
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3222
diff changeset
   124
              (finish_rule (Induct.internalize ctxt more_consumes rule)) i st'
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
   125
            |> Seq.maps (fn rule' =>
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   126
              CONTEXT_CASES (rule_cases ctxt rule' cases)
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   127
                (resolve_tac ctxt [rename_params_rule false [] rule'] i THEN
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   128
                  PRIMITIVE
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   129
                    (singleton (Proof_Context.export defs_ctxt
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   130
                      (Proof_Context.transfer (Proof_Context.theory_of defs_ctxt) ctxt)))) (ctxt, st'))));
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   131
  in
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   132
    (context_tac CONTEXT_THEN_ALL_NEW
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   133
      ((if simp then Induct.simplify_tac ctxt THEN' (TRY o Induct.trivial_tac ctxt)
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   134
        else K all_tac) THEN_ALL_NEW Induct.rulify_tac ctxt)) i (ctxt, st)
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
   135
  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
   136
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
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
(* 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
   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
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
   141
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
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
   143
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
   144
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
   145
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 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
   147
  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
   148
  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
   149
    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
   150
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
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
   152
  ((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
   153
      -- (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
   154
    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
   155
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
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
   157
  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
   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
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
   160
  ((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
   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
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
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
   164
  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
   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 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
   167
  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
   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 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
   170
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
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
   172
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
val nominal_induct_method =
3222
8c53bcd5c0ae updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   174
  Scan.lift (Args.mode Induct.no_simpN) -- 
8c53bcd5c0ae updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   175
    (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
8c53bcd5c0ae updated to lates Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   176
      avoiding -- fixing -- rule_spec) >>
3243
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   177
  (fn (no_simp, (((x, y), z), w)) => fn _ => fn facts =>
c4f31f1564b7 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   178
    nominal_induct_tac (not no_simp) x y z w facts 1);
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
   179
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
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
   181
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
end;