src/HOL/Library/simps_case_conv.ML
author wenzelm
Sun, 29 Sep 2024 21:16:17 +0200
changeset 81008 d0cd220d8e8b
parent 80634 a90ab1ea6458
child 82317 231b6d8231c6
permissions -rw-r--r--
more markup reports: notably "notation=..." within pretty blocks;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
     1
(*  Title:      HOL/Library/simps_case_conv.ML
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
     2
    Author:     Lars Noschinski, TU Muenchen
53433
3b356b7f7cad more standard header;
wenzelm
parents: 53429
diff changeset
     3
    Author:     Gerwin Klein, NICTA
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
     4
53433
3b356b7f7cad more standard header;
wenzelm
parents: 53429
diff changeset
     5
Convert function specifications between the representation as a list
3b356b7f7cad more standard header;
wenzelm
parents: 53429
diff changeset
     6
of equations (with patterns on the lhs) and a single equation (with a
3b356b7f7cad more standard header;
wenzelm
parents: 53429
diff changeset
     7
nested case expression on the rhs).
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
     8
*)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
     9
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    10
signature SIMPS_CASE_CONV =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    11
sig
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    12
  val to_case: Proof.context -> thm list -> thm
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    13
  val gen_to_simps: Proof.context -> thm list -> thm -> thm list
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    14
  val to_simps: Proof.context -> thm -> thm list
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    15
end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    16
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    17
structure Simps_Case_Conv: SIMPS_CASE_CONV =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    18
struct
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    19
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    20
(* Collects all type constructors in a type *)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    21
fun collect_Tcons (Type (name,Ts)) = name :: maps collect_Tcons Ts
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    22
  | collect_Tcons (TFree _) = []
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    23
  | collect_Tcons (TVar _) = []
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    24
60702
5e03e1bd1be0 case_of_simps: do not split for types with a single constructor
noschinl
parents: 60355
diff changeset
    25
fun get_type_infos ctxt =
5e03e1bd1be0 case_of_simps: do not split for types with a single constructor
noschinl
parents: 60355
diff changeset
    26
    maps collect_Tcons
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    27
    #> distinct (op =)
54401
f6950854855d ported 'Simps_Case_Conv' to use new 'Ctr_Sugar' abstraction
blanchet
parents: 53433
diff changeset
    28
    #> map_filter (Ctr_Sugar.ctr_sugar_of ctxt)
60702
5e03e1bd1be0 case_of_simps: do not split for types with a single constructor
noschinl
parents: 60355
diff changeset
    29
5e03e1bd1be0 case_of_simps: do not split for types with a single constructor
noschinl
parents: 60355
diff changeset
    30
fun get_split_ths ctxt = get_type_infos ctxt #> map #split
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    31
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    32
val strip_eq = Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    33
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    34
(*
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    35
  For a list
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    36
    f p_11 ... p_1n = t1
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    37
    f p_21 ... p_2n = t2
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    38
    ...
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    39
    f p_mn ... p_mn = tm
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    40
  of theorems, prove a single theorem
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    41
    f x1 ... xn = t
53429
9d9945941eab allowed less exhaustive patterns
noschinl
parents: 53426
diff changeset
    42
  where t is a (nested) case expression. f must not be a function
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    43
  application.
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    44
*)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    45
fun to_case ctxt ths =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    46
  let
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    47
    fun ctr_count (ctr, T) = 
59650
wenzelm
parents: 59621
diff changeset
    48
      let
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 79733
diff changeset
    49
        val tyco = dest_Type_name (body_type T)
69568
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    50
        val info = Ctr_Sugar.ctr_sugar_of ctxt tyco
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    51
        val _ = if is_none info then error ("Pattern match on non-constructor constant " ^ ctr) else ()
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    52
      in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    53
        info |> the |> #ctrs |> length
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    54
      end
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    55
    val thms = Case_Converter.to_case ctxt Case_Converter.keep_constructor_context ctr_count ths
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    56
  in
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    57
    case thms of SOME thms => hd thms
de09a7261120 new implementation for case_of_simps based on Code_Lazy's pattern matching elimination algorithm
Andreas Lochbihler
parents: 63352
diff changeset
    58
      | _ => error ("Conversion to case expression failed.")
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    59
  end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    60
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    61
local
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    62
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    63
fun was_split t =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    64
  let
56252
b72e0a9d62b9 more antiquotations;
wenzelm
parents: 55997
diff changeset
    65
    val is_free_eq_imp = is_Free o fst o HOLogic.dest_eq o fst o HOLogic.dest_imp
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    66
    val get_conjs = HOLogic.dest_conj o HOLogic.dest_Trueprop
74529
wenzelm
parents: 69593
diff changeset
    67
    val dest_alls = Term.strip_qnt_body \<^const_name>\<open>All\<close>
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    68
  in forall (is_free_eq_imp o dest_alls) (get_conjs t) end
56252
b72e0a9d62b9 more antiquotations;
wenzelm
parents: 55997
diff changeset
    69
  handle TERM _ => false
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    70
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    71
fun apply_split ctxt split thm = Seq.of_list
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    72
  let val ((_,thm'), ctxt') = Variable.import false [thm] ctxt in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    73
    (Variable.export ctxt' ctxt) (filter (was_split o Thm.prop_of) (thm' RL [split]))
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    74
  end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    75
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    76
fun forward_tac rules t = Seq.of_list ([t] RL rules)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    77
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    78
val refl_imp = refl RSN (2, mp)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    79
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    80
val get_rules_once_split =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    81
  REPEAT (forward_tac [conjunct1, conjunct2])
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    82
    THEN REPEAT (forward_tac [spec])
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    83
    THEN (forward_tac [refl_imp])
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    84
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    85
fun do_split ctxt split =
60355
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    86
  case try op RS (split, iffD1) of
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    87
    NONE => raise TERM ("malformed split rule", [Thm.prop_of split])
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    88
  | SOME split' =>
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    89
      let val split_rhs = Thm.concl_of (hd (snd (fst (Variable.import false [split'] ctxt))))
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    90
      in if was_split split_rhs
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    91
         then DETERM (apply_split ctxt split') THEN get_rules_once_split
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    92
         else raise TERM ("malformed split rule", [split_rhs])
ccafd7d193e7 simps_of_case: Better error if split rule is not an equality
noschinl
parents: 60354
diff changeset
    93
      end
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    94
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    95
val atomize_meta_eq = forward_tac [meta_eq_to_obj_eq]
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    96
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    97
in
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    98
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
    99
fun gen_to_simps ctxt splitthms thm =
60354
235d65be79c9 simps_of_case: allow Drule.dummy_thm as ignored split rule
noschinl
parents: 59936
diff changeset
   100
  let val splitthms' = filter (fn t => not (Thm.eq_thm (t, Drule.dummy_thm))) splitthms
235d65be79c9 simps_of_case: allow Drule.dummy_thm as ignored split rule
noschinl
parents: 59936
diff changeset
   101
  in
235d65be79c9 simps_of_case: allow Drule.dummy_thm as ignored split rule
noschinl
parents: 59936
diff changeset
   102
    Seq.list_of ((TRY atomize_meta_eq THEN (REPEAT (FIRST (map (do_split ctxt) splitthms')))) thm)
235d65be79c9 simps_of_case: allow Drule.dummy_thm as ignored split rule
noschinl
parents: 59936
diff changeset
   103
  end
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   104
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   105
fun to_simps ctxt thm =
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   106
  let
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   107
    val T = thm |> strip_eq |> fst |> strip_comb |> fst |> fastype_of
60702
5e03e1bd1be0 case_of_simps: do not split for types with a single constructor
noschinl
parents: 60355
diff changeset
   108
    val splitthms = get_split_ths ctxt [T]
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   109
  in gen_to_simps ctxt splitthms thm end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   110
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   111
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   112
end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   113
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   114
fun case_of_simps_cmd (bind, thms_ref) int lthy =
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   115
  let
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 54883
diff changeset
   116
    val bind' = apsnd (map (Attrib.check_src lthy)) bind
61813
wenzelm
parents: 60702
diff changeset
   117
    val thm = Attrib.eval_thms lthy thms_ref |> to_case lthy
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   118
    val (res, lthy') = Local_Theory.note (bind', [thm]) lthy 
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   119
    val _ =
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   120
      Proof_Display.print_results
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   121
        {interactive = int, pos = Position.thread_data (), proof_state = false}
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   122
        lthy' ((Thm.theoremK, ""), [res])
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   123
  in lthy' end
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   124
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   125
fun simps_of_case_cmd ((bind, thm_ref), splits_ref) int lthy =
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   126
  let
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 54883
diff changeset
   127
    val bind' = apsnd (map (Attrib.check_src lthy)) bind
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   128
    val thm = singleton (Attrib.eval_thms lthy) thm_ref
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   129
    val simps = if null splits_ref
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   130
      then to_simps lthy thm
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   131
      else gen_to_simps lthy (Attrib.eval_thms lthy splits_ref) thm
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   132
    val (res, lthy') = Local_Theory.note (bind', simps) lthy
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   133
    val _ =
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   134
      Proof_Display.print_results
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   135
        {interactive = int, pos = Position.thread_data (), proof_state = false}
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   136
        lthy' ((Thm.theoremK, ""), [res])
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   137
  in lthy' end
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   138
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   139
val _ =
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   140
  Outer_Syntax.local_theory' \<^command_keyword>\<open>case_of_simps\<close>
57628
c80fc5576271 tuned message;
wenzelm
parents: 56252
diff changeset
   141
    "turn a list of equations into a case expression"
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 61813
diff changeset
   142
    (Parse_Spec.opt_thm_name ":"  -- Parse.thms1 >> case_of_simps_cmd)
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   143
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69568
diff changeset
   144
val parse_splits = \<^keyword>\<open>(\<close> |-- Parse.reserved "splits" |-- \<^keyword>\<open>:\<close> |--
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69568
diff changeset
   145
  Parse.thms1 --| \<^keyword>\<open>)\<close>
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   146
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   147
val _ =
79733
3e30ca77ccfe improved output in simps_case_conv;
Fabian Huch <huch@in.tum.de>
parents: 74529
diff changeset
   148
  Outer_Syntax.local_theory' \<^command_keyword>\<open>simps_of_case\<close>
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   149
    "perform case split on rule"
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 61813
diff changeset
   150
    (Parse_Spec.opt_thm_name ":"  -- Parse.thm --
53426
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   151
      Scan.optional parse_splits [] >> simps_of_case_cmd)
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   152
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   153
end
92db671e0ac6 added simps_of_case and case_of_simps to convert between simps and case rules
noschinl
parents:
diff changeset
   154