src/Pure/Isar/rule_cases.ML
author hoelzl
Tue, 17 May 2011 15:00:39 +0200
changeset 42867 760094e49a2c
parent 42496 65ec88b369fd
child 44045 2814ff2a6e3e
permissions -rw-r--r--
Collect intro-rules for sigma-algebras * * * sets_Collect shouldn't be intro rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/rule_cases.ML
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     3
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
     4
Annotations and local contexts of rules.
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     5
*)
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     6
18188
cb53a778455e added THEN_ALL_NEW_CASES;
wenzelm
parents: 18050
diff changeset
     7
infix 1 THEN_ALL_NEW_CASES;
cb53a778455e added THEN_ALL_NEW_CASES;
wenzelm
parents: 18050
diff changeset
     8
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
     9
signature BASIC_RULE_CASES =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    10
sig
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    11
  type cases
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    12
  type cases_tactic
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    13
  val CASES: cases -> tactic -> cases_tactic
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    14
  val NO_CASES: tactic -> cases_tactic
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    15
  val SUBGOAL_CASES: ((term * int) -> cases_tactic) -> int -> cases_tactic
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    16
  val THEN_ALL_NEW_CASES: (int -> cases_tactic) * (int -> tactic) -> int -> cases_tactic
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    17
end
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    18
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    19
signature RULE_CASES =
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    20
sig
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
    21
  include BASIC_RULE_CASES
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    22
  datatype T = Case of
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    23
   {fixes: (binding * typ) list,
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    24
    assumes: (string * term list) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    25
    binds: (indexname * term option) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    26
    cases: (string * T) list}
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    27
  val strip_params: term -> (string * typ) list
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
    28
  val make_common: theory * term -> (string * string list) list -> cases
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
    29
  val make_nested: term -> theory * term -> (string * string list) list -> cases
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    30
  val apply: term list -> T -> T
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
    31
  val consume: thm list -> thm list -> ('a * int) * thm ->
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
    32
    (('a * (int * thm list)) * thm) Seq.seq
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
    33
  val add_consumes: int -> thm -> thm
24862
6b7258da912b export get_consumes;
wenzelm
parents: 24244
diff changeset
    34
  val get_consumes: thm -> int
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    35
  val consumes: int -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    36
  val consumes_default: int -> attribute
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
    37
  val get_constraints: thm -> int
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
    38
  val constraints: int -> attribute
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    39
  val name: string list -> thm -> thm
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    40
  val case_names: string list -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    41
  val case_conclusion: string * string list -> attribute
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
    42
  val is_inner_rule: thm -> bool
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
    43
  val inner_rule: attribute
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
    44
  val save: thm -> thm -> thm
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
    45
  val get: thm -> (string * string list) list * int
8427
wenzelm
parents: 8400
diff changeset
    46
  val rename_params: string list list -> thm -> thm
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    47
  val params: string list list -> attribute
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
    48
  val internalize_params: thm -> thm
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20218
diff changeset
    49
  val mutual_rule: Proof.context -> thm list -> (int list * thm) option
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20218
diff changeset
    50
  val strict_mutual_rule: Proof.context -> thm list -> int list * thm
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    51
end;
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    52
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33303
diff changeset
    53
structure Rule_Cases: RULE_CASES =
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    54
struct
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    55
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    56
(** cases **)
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    57
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    58
datatype T = Case of
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    59
 {fixes: (binding * typ) list,
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    60
  assumes: (string * term list) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    61
  binds: (indexname * term option) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    62
  cases: (string * T) list};
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    63
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    64
type cases = (string * T option) list;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    65
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    66
val case_conclN = "case";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    67
val case_hypsN = "hyps";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    68
val case_premsN = "prems";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    69
20087
979f012b5df3 Name.internal;
wenzelm
parents: 19917
diff changeset
    70
val strip_params = map (apfst (perhaps (try Name.dest_skolem))) o Logic.strip_params;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    71
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    72
local
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    73
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    74
fun abs xs t = Term.list_abs (xs, t);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    75
fun app us t = Term.betapplys (t, us);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    76
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    77
fun dest_binops cs tm =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    78
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    79
    val n = length cs;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    80
    fun dest 0 _ = []
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    81
      | dest 1 t = [t]
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    82
      | dest k (_ $ t $ u) = t :: dest (k - 1) u
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    83
      | dest _ _ = raise TERM ("Expected " ^ string_of_int n ^ " binop arguments", [tm]);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    84
  in cs ~~ dest n tm end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    85
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    86
fun extract_fixes NONE prop = (strip_params prop, [])
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    87
  | extract_fixes (SOME outline) prop =
19012
wenzelm
parents: 18909
diff changeset
    88
      chop (length (Logic.strip_params outline)) (strip_params prop);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    89
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    90
fun extract_assumes _ NONE prop = ([("", Logic.strip_assums_hyp prop)], [])
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    91
  | extract_assumes qual (SOME outline) prop =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    92
      let val (hyps, prems) =
19012
wenzelm
parents: 18909
diff changeset
    93
        chop (length (Logic.strip_assums_hyp outline)) (Logic.strip_assums_hyp prop)
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    94
      in ([(qual case_hypsN, hyps)], [(qual case_premsN, prems)]) end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    95
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    96
fun bindings args = map (apfst Binding.name) args;
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    97
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
    98
fun extract_case thy (case_outline, raw_prop) name concls =
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    99
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   100
    val props = Logic.dest_conjunctions (Drule.norm_hhf thy raw_prop);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   101
    val len = length props;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   102
    val nested = is_some case_outline andalso len > 1;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   103
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   104
    fun extract prop =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   105
      let
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   106
        val (fixes1, fixes2) = extract_fixes case_outline prop;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   107
        val abs_fixes = abs (fixes1 @ fixes2);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   108
        fun abs_fixes1 t =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   109
          if not nested then abs_fixes t
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   110
          else abs fixes1 (app (map (Term.dummy_pattern o #2) fixes2) (abs fixes2 t));
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   111
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 29269
diff changeset
   112
        val (assumes1, assumes2) = extract_assumes (Long_Name.qualify name) case_outline prop
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19423
diff changeset
   113
          |> pairself (map (apsnd (maps Logic.dest_conjunctions)));
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   114
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35408
diff changeset
   115
        val concl = Object_Logic.drop_judgment thy (Logic.strip_assums_concl prop);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   116
        val binds =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   117
          (case_conclN, concl) :: dest_binops concls concl
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   118
          |> map (fn (x, t) => ((x, 0), SOME (abs_fixes t)));
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   119
      in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   120
       ((fixes1, map (apsnd (map abs_fixes1)) assumes1),
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   121
        ((fixes2, map (apsnd (map abs_fixes)) assumes2), binds))
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   122
      end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   123
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   124
    val cases = map extract props;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   125
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   126
    fun common_case ((fixes1, assumes1), ((fixes2, assumes2), binds)) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   127
      Case {fixes = bindings (fixes1 @ fixes2),
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   128
        assumes = assumes1 @ assumes2, binds = binds, cases = []};
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   129
    fun inner_case (_, ((fixes2, assumes2), binds)) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   130
      Case {fixes = bindings fixes2, assumes = assumes2, binds = binds, cases = []};
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   131
    fun nested_case ((fixes1, assumes1), _) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   132
      Case {fixes = bindings fixes1, assumes = assumes1, binds = [],
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   133
        cases = map string_of_int (1 upto len) ~~ map inner_case cases};
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   134
  in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   135
    if len = 0 then NONE
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   136
    else if len = 1 then SOME (common_case (hd cases))
19046
bc5c6c9b114e removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19012
diff changeset
   137
    else if is_none case_outline orelse length (distinct (op =) (map fst cases)) > 1 then NONE
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   138
    else SOME (nested_case (hd cases))
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   139
  end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   140
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   141
fun make rule_struct (thy, prop) cases =
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   142
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   143
    val n = length cases;
21576
8c11b1ce2f05 simplified Logic.count_prems;
wenzelm
parents: 21395
diff changeset
   144
    val nprems = Logic.count_prems prop;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   145
    fun add_case (name, concls) (cs, i) =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   146
      ((case try (fn () =>
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   147
          (Option.map (curry Logic.nth_prem i) rule_struct, Logic.nth_prem (i, prop))) () of
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   148
        NONE => (name, NONE)
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   149
      | SOME p => (name, extract_case thy p name concls)) :: cs, i - 1);
34058
97fd820dd402 explicit lower bound for index
haftmann
parents: 33957
diff changeset
   150
  in fold_rev add_case (drop (Int.max (n - nprems, 0)) cases) ([], n) |> #1 end;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   151
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   152
in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   153
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   154
val make_common = make NONE;
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   155
fun make_nested rule_struct = make (SOME rule_struct);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   156
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   157
fun apply args =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   158
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   159
    fun appl (Case {fixes, assumes, binds, cases}) =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   160
      let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   161
        val assumes' = map (apsnd (map (app args))) assumes;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   162
        val binds' = map (apsnd (Option.map (app args))) binds;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   163
        val cases' = map (apsnd appl) cases;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   164
      in Case {fixes = fixes, assumes = assumes', binds = binds', cases = cases'} end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   165
  in appl end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   166
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   167
end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   168
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   169
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   170
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   171
(** tactics with cases **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   172
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   173
type cases_tactic = thm -> (cases * thm) Seq.seq;
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   174
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   175
fun CASES cases tac st = Seq.map (pair cases) (tac st);
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   176
fun NO_CASES tac = CASES [] tac;
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   177
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   178
fun SUBGOAL_CASES tac i st =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   179
  (case try Logic.nth_prem (i, Thm.prop_of st) of
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   180
    SOME goal => tac (goal, i) st
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   181
  | NONE => Seq.empty);
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   182
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   183
fun (tac1 THEN_ALL_NEW_CASES tac2) i st =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   184
  st |> tac1 i |> Seq.maps (fn (cases, st') =>
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   185
    CASES cases (Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st)) st');
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   186
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   187
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   188
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   189
(** consume facts **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   190
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   191
local
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   192
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   193
fun unfold_prems n defs th =
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   194
  if null defs then th
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 36354
diff changeset
   195
  else Conv.fconv_rule (Conv.prems_conv n (Raw_Simplifier.rewrite true defs)) th;
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   196
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   197
fun unfold_prems_concls defs th =
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   198
  if null defs orelse not (can Logic.dest_conjunction (Thm.concl_of th)) then th
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   199
  else
22900
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22568
diff changeset
   200
    Conv.fconv_rule
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 22900
diff changeset
   201
      (Conv.concl_conv ~1 (Conjunction.convs
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 36354
diff changeset
   202
        (Conv.prems_conv ~1 (Raw_Simplifier.rewrite true defs)))) th;
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   203
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   204
in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   205
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   206
fun consume defs facts ((xx, n), th) =
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   207
  let val m = Int.min (length facts, n) in
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   208
    th
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   209
    |> unfold_prems n defs
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   210
    |> unfold_prems_concls defs
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   211
    |> Drule.multi_resolve (take m facts)
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   212
    |> Seq.map (pair (xx, (n - m, drop m facts)))
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   213
  end;
10811
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   214
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   215
end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   216
10811
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   217
val consumes_tagN = "consumes";
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   218
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   219
fun lookup_consumes th =
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   220
  (case AList.lookup (op =) (Thm.get_tags th) consumes_tagN of
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   221
    NONE => NONE
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   222
  | SOME s =>
24244
d7ee11ba1534 Lexicon.read_indexname/nat/variable;
wenzelm
parents: 23657
diff changeset
   223
      (case Lexicon.read_nat s of SOME n => SOME n
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   224
      | _ => raise THM ("Malformed 'consumes' tag of theorem", 0, [th])));
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   225
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   226
fun get_consumes th = the_default 0 (lookup_consumes th);
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   227
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   228
fun put_consumes NONE th = th
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   229
  | put_consumes (SOME n) th = th
27865
27a8ad9612a3 moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents: 26923
diff changeset
   230
      |> Thm.untag_rule consumes_tagN
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   231
      |> Thm.tag_rule (consumes_tagN, string_of_int (if n < 0 then Thm.nprems_of th + n else n));
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   232
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   233
fun add_consumes k th = put_consumes (SOME (k + get_consumes th)) th;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   234
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   235
val save_consumes = put_consumes o lookup_consumes;
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   236
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   237
fun consumes n = Thm.rule_attribute (K (put_consumes (SOME n)));
21395
f34ac19659ae moved some fundamental concepts to General/basics.ML;
wenzelm
parents: 20289
diff changeset
   238
16148
5f15a14122dc make: T option -- actually remove undefined cases;
wenzelm
parents: 15973
diff changeset
   239
fun consumes_default n x =
21395
f34ac19659ae moved some fundamental concepts to General/basics.ML;
wenzelm
parents: 20289
diff changeset
   240
  if is_some (lookup_consumes (#2 x)) then x else consumes n x;
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   241
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   242
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   243
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   244
(** equality constraints **)
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   245
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   246
val constraints_tagN = "constraints";
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   247
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   248
fun lookup_constraints th =
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   249
  (case AList.lookup (op =) (Thm.get_tags th) constraints_tagN of
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   250
    NONE => NONE
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   251
  | SOME s =>
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   252
      (case Lexicon.read_nat s of SOME n => SOME n
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   253
      | _ => raise THM ("Malformed 'constraints' tag of theorem", 0, [th])));
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   254
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   255
fun get_constraints th = the_default 0 (lookup_constraints th);
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   256
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   257
fun put_constraints NONE th = th
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   258
  | put_constraints (SOME n) th = th
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   259
      |> Thm.untag_rule constraints_tagN
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   260
      |> Thm.tag_rule (constraints_tagN, string_of_int (if n < 0 then 0 else n));
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   261
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   262
val save_constraints = put_constraints o lookup_constraints;
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   263
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   264
fun constraints n = Thm.rule_attribute (K (put_constraints (SOME n)));
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   265
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   266
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   267
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   268
(** case names **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   269
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   270
val implode_args = space_implode ";";
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   271
val explode_args = space_explode ";";
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   272
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   273
val case_names_tagN = "case_names";
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   274
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   275
fun add_case_names NONE = I
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   276
  | add_case_names (SOME names) =
27865
27a8ad9612a3 moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents: 26923
diff changeset
   277
      Thm.untag_rule case_names_tagN
27a8ad9612a3 moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents: 26923
diff changeset
   278
      #> Thm.tag_rule (case_names_tagN, implode_args names);
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   279
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   280
fun lookup_case_names th =
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   281
  AList.lookup (op =) (Thm.get_tags th) case_names_tagN
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   282
  |> Option.map explode_args;
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   283
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   284
val save_case_names = add_case_names o lookup_case_names;
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   285
val name = add_case_names o SOME;
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
   286
fun case_names ss = Thm.rule_attribute (K (name ss));
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   287
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   288
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   289
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   290
(** case conclusions **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   291
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   292
val case_concl_tagN = "case_conclusion";
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   293
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   294
fun get_case_concl name (a, b) =
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   295
  if a = case_concl_tagN then
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   296
    (case explode_args b of c :: cs => if c = name then SOME cs else NONE)
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   297
  else NONE;
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   298
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 21576
diff changeset
   299
fun add_case_concl (name, cs) = Thm.map_tags (fn tags =>
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   300
  filter_out (is_some o get_case_concl name) tags @
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   301
    [(case_concl_tagN, implode_args (name :: cs))]);
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   302
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   303
fun get_case_concls th name =
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   304
  these (get_first (get_case_concl name) (Thm.get_tags th));
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   305
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   306
fun save_case_concls th =
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 21576
diff changeset
   307
  let val concls = Thm.get_tags th |> map_filter
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   308
    (fn (a, b) =>
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   309
      if a = case_concl_tagN then (case explode_args b of c :: cs => SOME (c, cs) | _ => NONE)
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   310
      else NONE)
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   311
  in fold add_case_concl concls end;
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   312
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   313
fun case_conclusion concl = Thm.rule_attribute (K (add_case_concl concl));
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   314
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   315
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   316
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   317
(** inner rule **)
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   318
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   319
val inner_rule_tagN = "inner_rule";
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   320
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   321
fun is_inner_rule th =
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   322
  AList.defined (op =) (Thm.get_tags th) inner_rule_tagN;
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   323
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   324
fun put_inner_rule inner =
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   325
  Thm.untag_rule inner_rule_tagN
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   326
  #> inner ? Thm.tag_rule (inner_rule_tagN, "");
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   327
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   328
val save_inner_rule = put_inner_rule o is_inner_rule;
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   329
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   330
val inner_rule = Thm.rule_attribute (K (put_inner_rule true));
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   331
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   332
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   333
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   334
(** case declarations **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   335
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   336
(* access hints *)
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   337
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   338
fun save th =
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   339
  save_consumes th #>
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   340
  save_constraints th #>
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   341
  save_case_names th #>
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   342
  save_case_concls th #>
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   343
  save_inner_rule th;
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   344
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   345
fun get th =
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   346
  let
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   347
    val n = get_consumes th;
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   348
    val cases =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   349
      (case lookup_case_names th of
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   350
        NONE => map (rpair [] o Library.string_of_int) (1 upto (Thm.nprems_of th - n))
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   351
      | SOME names => map (fn name => (name, get_case_concls th name)) names);
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   352
  in (cases, n) end;
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   353
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   354
8427
wenzelm
parents: 8400
diff changeset
   355
(* params *)
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   356
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   357
fun rename_params xss th =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   358
  th
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   359
  |> fold_index (fn (i, xs) => Thm.rename_params_rule (xs, i + 1)) xss
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   360
  |> save th;
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   361
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
   362
fun params xss = Thm.rule_attribute (K (rename_params xss));
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   363
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   364
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   365
(* internalize parameter names *)
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   366
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   367
fun internalize_params rule =
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   368
  let
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   369
    fun rename prem =
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   370
      let val xs =
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   371
        map (Name.internal o Name.clean o fst) (Logic.strip_params prem)
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   372
      in Logic.list_rename_params (xs, prem) end;
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   373
    fun rename_prems prop =
36354
bbd742107f56 eliminanated some unreferenced identifiers;
wenzelm
parents: 35625
diff changeset
   374
      let val (As, C) = Logic.strip_horn prop
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   375
      in Logic.list_implies (map rename As, C) end;
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   376
  in Thm.equal_elim (Thm.reflexive (Drule.cterm_fun rename_prems (Thm.cprop_of rule))) rule end;
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   377
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   378
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   379
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   380
(** mutual_rule **)
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   381
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   382
local
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   383
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   384
fun equal_cterms ts us =
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 34986
diff changeset
   385
  is_equal (list_ord (Term_Ord.fast_term_ord o pairself Thm.term_of) (ts, us));
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   386
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   387
fun prep_rule n th =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   388
  let
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   389
    val th' = Thm.permute_prems 0 n th;
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   390
    val prems = take (Thm.nprems_of th' - n) (Drule.cprems_of th');
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   391
    val th'' = Drule.implies_elim_list th' (map Thm.assume prems);
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   392
  in (prems, (n, th'')) end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   393
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   394
in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   395
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   396
fun mutual_rule _ [] = NONE
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   397
  | mutual_rule _ [th] = SOME ([0], th)
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   398
  | mutual_rule ctxt (ths as th :: _) =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   399
      let
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 30364
diff changeset
   400
        val ((_, ths'), ctxt') = Variable.import true ths ctxt;
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   401
        val rules as (prems, _) :: _ = map (prep_rule (get_consumes th)) ths';
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   402
        val (ns, rls) = split_list (map #2 rules);
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   403
      in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   404
        if not (forall (equal_cterms prems o #1) rules) then NONE
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   405
        else
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   406
          SOME (ns,
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   407
            rls
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 22900
diff changeset
   408
            |> Conjunction.intr_balanced
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   409
            |> Drule.implies_intr_list prems
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   410
            |> singleton (Variable.export ctxt' ctxt)
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   411
            |> save th
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   412
            |> put_consumes (SOME 0))
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   413
      end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   414
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   415
end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   416
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   417
fun strict_mutual_rule ctxt ths =
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   418
  (case mutual_rule ctxt ths of
18581
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   419
    NONE => error "Failed to join given rules into one mutual rule"
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   420
  | SOME res => res);
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   421
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   422
end;
17113
3b67c1809e1c added NO_CASES;
wenzelm
parents: 16390
diff changeset
   423
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33303
diff changeset
   424
structure Basic_Rule_Cases: BASIC_RULE_CASES = Rule_Cases;
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   425
open Basic_Rule_Cases;