src/Pure/Isar/rule_cases.ML
author wenzelm
Fri, 11 Dec 2015 13:44:20 +0100
changeset 61834 2154e6c8d52d
parent 60576 51f1d213079c
child 61853 fb7756087101
permissions -rw-r--r--
clarified modules;
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
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     7
signature RULE_CASES =
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
     8
sig
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
     9
  datatype T = Case of
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    10
   {fixes: (binding * typ) list,
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    11
    assumes: (string * term list) list,
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    12
    binds: (indexname * term option) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    13
    cases: (string * T) list}
61834
2154e6c8d52d clarified modules;
wenzelm
parents: 60576
diff changeset
    14
  type cases = (string * T option) list
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
    15
  val case_hypsN: string
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    16
  val strip_params: term -> (string * typ) list
60576
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    17
  type info = ((string * string list) * string list) list
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    18
  val make_common: Proof.context -> term -> info -> cases
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    19
  val make_nested: Proof.context -> term -> term -> info -> cases
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    20
  val apply: term list -> T -> T
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
    21
  val consume: Proof.context -> thm list -> thm list -> ('a * int) * thm ->
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
    22
    (('a * (int * thm list)) * thm) Seq.seq
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
    23
  val get_consumes: thm -> int
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
    24
  val put_consumes: int option -> thm -> thm
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
    25
  val add_consumes: int -> thm -> thm
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
    26
  val default_consumes: int -> thm -> thm
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    27
  val consumes: int -> attribute
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
    28
  val get_constraints: thm -> int
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
    29
  val constraints: int -> attribute
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    30
  val name: string list -> thm -> thm
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    31
  val case_names: string list -> attribute
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
    32
  val cases_hyp_names: string list -> string list list -> attribute
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    33
  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
    34
  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
    35
  val inner_rule: attribute
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
    36
  val is_cases_open: thm -> bool
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
    37
  val cases_open: attribute
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
    38
  val internalize_params: thm -> thm
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
    39
  val save: thm -> thm -> thm
60576
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    40
  val get: thm -> info * int
8427
wenzelm
parents: 8400
diff changeset
    41
  val rename_params: string list list -> thm -> thm
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
    42
  val params: string list list -> attribute
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20218
diff changeset
    43
  val mutual_rule: Proof.context -> thm list -> (int list * thm) option
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20218
diff changeset
    44
  val strict_mutual_rule: Proof.context -> thm list -> int list * thm
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    45
end;
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    46
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33303
diff changeset
    47
structure Rule_Cases: RULE_CASES =
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    48
struct
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
    49
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    50
(** cases **)
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    51
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    52
datatype T = Case of
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    53
 {fixes: (binding * typ) list,
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    54
  assumes: (string * term list) list,
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    55
  binds: (indexname * term option) list,
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    56
  cases: (string * T) list};
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    57
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    58
type cases = (string * T option) list;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    59
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    60
val case_conclN = "case";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    61
val case_hypsN = "hyps";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    62
val case_premsN = "prems";
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    63
20087
979f012b5df3 Name.internal;
wenzelm
parents: 19917
diff changeset
    64
val strip_params = map (apfst (perhaps (try Name.dest_skolem))) o Logic.strip_params;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    65
60576
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    66
type info = ((string * string list) * string list) list;
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
    67
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    68
local
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    69
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    70
fun app us t = Term.betapplys (t, us);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    71
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    72
fun dest_binops cs tm =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    73
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    74
    val n = length cs;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    75
    fun dest 0 _ = []
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    76
      | dest 1 t = [t]
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    77
      | dest k (_ $ t $ u) = t :: dest (k - 1) u
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    78
      | dest _ _ = raise TERM ("Expected " ^ string_of_int n ^ " binop arguments", [tm]);
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    79
  in cs ~~ dest n tm end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    80
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    81
fun extract_fixes NONE prop = (strip_params prop, [])
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    82
  | extract_fixes (SOME outline) prop =
19012
wenzelm
parents: 18909
diff changeset
    83
      chop (length (Logic.strip_params outline)) (strip_params prop);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    84
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    85
fun extract_assumes _ NONE prop = ([("", Logic.strip_assums_hyp prop)], [])
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    86
  | extract_assumes hyp_names (SOME outline) prop =
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
    87
      let
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
    88
        val (hyps, prems) =
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
    89
          chop (length (Logic.strip_assums_hyp outline)) (Logic.strip_assums_hyp prop)
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    90
        fun extract ((hyp_name, hyp) :: rest) = (hyp_name, hyp :: map snd rest);
53380
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
    91
        val (hyps1, hyps2) = chop (length hyp_names) hyps;
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
    92
        val pairs1 = if null hyps1 then [] else hyp_names ~~ hyps1;
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
    93
        val pairs = pairs1 @ map (pair case_hypsN) hyps2;
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
    94
        val named_hyps = map extract (partition_eq (eq_fst op =) pairs);
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    95
      in (named_hyps, [(case_premsN, prems)]) end;
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
    96
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    97
fun bindings args = map (apfst Binding.name) args;
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
    98
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
    99
fun extract_case ctxt (case_outline, raw_prop) hyp_names concls =
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   100
  let
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59953
diff changeset
   101
    val props = Logic.dest_conjunctions (Drule.norm_hhf (Proof_Context.theory_of ctxt) raw_prop);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   102
    val len = length props;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   103
    val nested = is_some case_outline andalso len > 1;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   104
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   105
    fun extract prop =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   106
      let
34916
f625d8d6fcf3 Eliminated is_open option of Rule_Cases.make_nested/make_common;
berghofe
parents: 34058
diff changeset
   107
        val (fixes1, fixes2) = extract_fixes case_outline prop;
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45525
diff changeset
   108
        val abs_fixes = fold_rev Term.abs (fixes1 @ fixes2);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   109
        fun abs_fixes1 t =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   110
          if not nested then abs_fixes t
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45525
diff changeset
   111
          else
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45525
diff changeset
   112
            fold_rev Term.abs fixes1
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 45525
diff changeset
   113
              (app (map (Term.dummy_pattern o #2) fixes2) (fold_rev Term.abs fixes2 t));
53380
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
   114
        val (assumes1, assumes2) =
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
   115
          extract_assumes hyp_names case_outline prop
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   116
          |> apply2 (map (apsnd (maps Logic.dest_conjunctions)));
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   117
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59953
diff changeset
   118
        val concl = Object_Logic.drop_judgment ctxt (Logic.strip_assums_concl prop);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   119
        val binds =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   120
          (case_conclN, concl) :: dest_binops concls concl
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   121
          |> map (fn (x, t) => ((x, 0), SOME (abs_fixes t)));
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   122
      in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   123
       ((fixes1, map (apsnd (map abs_fixes1)) assumes1),
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   124
        ((fixes2, map (apsnd (map abs_fixes)) assumes2), binds))
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   125
      end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   126
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   127
    val cases = map extract props;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   128
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   129
    fun common_case ((fixes1, assumes1), ((fixes2, assumes2), binds)) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   130
      Case {fixes = bindings (fixes1 @ fixes2),
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   131
        assumes = assumes1 @ assumes2, binds = binds, cases = []};
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   132
    fun inner_case (_, ((fixes2, assumes2), binds)) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   133
      Case {fixes = bindings fixes2, assumes = assumes2, binds = binds, cases = []};
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   134
    fun nested_case ((fixes1, assumes1), _) =
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 41228
diff changeset
   135
      Case {fixes = bindings fixes1, assumes = assumes1, binds = [],
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   136
        cases = map string_of_int (1 upto len) ~~ map inner_case cases};
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   137
  in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   138
    if len = 0 then NONE
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   139
    else if len = 1 then SOME (common_case (hd cases))
19046
bc5c6c9b114e removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 19012
diff changeset
   140
    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
   141
    else SOME (nested_case (hd cases))
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   142
  end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   143
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59953
diff changeset
   144
fun make ctxt rule_struct prop cases =
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   145
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   146
    val n = length cases;
21576
8c11b1ce2f05 simplified Logic.count_prems;
wenzelm
parents: 21395
diff changeset
   147
    val nprems = Logic.count_prems prop;
53380
08f3491c50bf cases: formal binding of 'assumes', with position provided via invoke_case;
wenzelm
parents: 46219
diff changeset
   148
    fun add_case ((name, hyp_names), concls) (cs, i) =
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   149
      ((case try (fn () =>
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   150
          (Option.map (curry Logic.nth_prem i) rule_struct, Logic.nth_prem (i, prop))) () of
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   151
        NONE => (name, NONE)
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60456
diff changeset
   152
      | SOME p => (name, extract_case ctxt p hyp_names concls)) :: cs, i - 1);
34058
97fd820dd402 explicit lower bound for index
haftmann
parents: 33957
diff changeset
   153
  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
   154
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   155
in
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   156
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59953
diff changeset
   157
fun make_common ctxt = make ctxt NONE;
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59953
diff changeset
   158
fun make_nested ctxt rule_struct = make ctxt (SOME rule_struct);
18608
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   159
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   160
fun apply args =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   161
  let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   162
    fun appl (Case {fixes, assumes, binds, cases}) =
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   163
      let
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   164
        val assumes' = map (apsnd (map (app args))) assumes;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   165
        val binds' = map (apsnd (Option.map (app args))) binds;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   166
        val cases' = map (apsnd appl) cases;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   167
      in Case {fixes = fixes, assumes = assumes', binds = binds', cases = cases'} end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   168
  in appl end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   169
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   170
end;
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   171
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   172
9cdcc2a5c8b3 support nested cases;
wenzelm
parents: 18581
diff changeset
   173
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   174
(** annotations and operations on rules **)
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   175
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   176
(* consume facts *)
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   177
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   178
local
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   179
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   180
fun unfold_prems ctxt n defs th =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   181
  if null defs then th
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   182
  else Conv.fconv_rule (Conv.prems_conv n (Raw_Simplifier.rewrite ctxt true defs)) th;
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   183
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   184
fun unfold_prems_concls ctxt defs th =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   185
  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
   186
  else
22900
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22568
diff changeset
   187
    Conv.fconv_rule
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 22900
diff changeset
   188
      (Conv.concl_conv ~1 (Conjunction.convs
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   189
        (Conv.prems_conv ~1 (Raw_Simplifier.rewrite ctxt true defs)))) th;
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   190
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   191
in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   192
60576
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
   193
fun consume ctxt defs facts ((a, n), th) =
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   194
  let val m = Int.min (length facts, n) in
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   195
    th
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   196
    |> unfold_prems ctxt n defs
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53380
diff changeset
   197
    |> unfold_prems_concls ctxt defs
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58014
diff changeset
   198
    |> Drule.multi_resolve (SOME ctxt) (take m facts)
60576
51f1d213079c tuned signature;
wenzelm
parents: 60574
diff changeset
   199
    |> Seq.map (pair (a, (n - m, drop m facts)))
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   200
  end;
10811
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   201
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   202
end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   203
10811
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   204
val consumes_tagN = "consumes";
98f52cb93d93 support ?case binding;
wenzelm
parents: 10530
diff changeset
   205
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   206
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
   207
  (case AList.lookup (op =) (Thm.get_tags th) consumes_tagN of
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   208
    NONE => NONE
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   209
  | SOME s =>
55639
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   210
      (case Lexicon.read_nat s of
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   211
        SOME n => SOME n
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   212
      | _ => raise THM ("Malformed 'consumes' tag of theorem", 0, [th])));
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   213
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   214
fun get_consumes th = the_default 0 (lookup_consumes th);
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   215
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   216
fun put_consumes NONE th = th
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   217
  | 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
   218
      |> 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
   219
      |> 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
   220
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   221
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
   222
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   223
fun default_consumes n th =
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   224
  if is_some (lookup_consumes th) then th else put_consumes (SOME n) th;
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   225
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   226
val save_consumes = put_consumes o lookup_consumes;
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   227
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   228
fun consumes n = Thm.mixed_attribute (apsnd (put_consumes (SOME n)));
21395
f34ac19659ae moved some fundamental concepts to General/basics.ML;
wenzelm
parents: 20289
diff changeset
   229
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   230
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   231
(* equality constraints *)
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   232
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   233
val constraints_tagN = "constraints";
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   234
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   235
fun lookup_constraints th =
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   236
  (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
   237
    NONE => NONE
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   238
  | SOME s =>
55639
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   239
      (case Lexicon.read_nat s of
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   240
        SOME n => SOME n
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   241
      | _ => 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
   242
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   243
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
   244
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   245
fun put_constraints NONE th = th
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   246
  | put_constraints (SOME n) th = th
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   247
      |> Thm.untag_rule constraints_tagN
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   248
      |> 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
   249
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   250
val save_constraints = put_constraints o lookup_constraints;
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   251
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   252
fun constraints n = Thm.mixed_attribute (apsnd (put_constraints (SOME n)));
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   253
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   254
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   255
(* case names *)
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   256
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   257
val implode_args = space_implode ";";
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   258
val explode_args = space_explode ";";
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   259
18237
2edb6a1f9c14 consume: proper treatment of defs;
wenzelm
parents: 18229
diff changeset
   260
val case_names_tagN = "case_names";
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   261
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   262
fun add_case_names NONE = I
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   263
  | 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
   264
      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
   265
      #> Thm.tag_rule (case_names_tagN, implode_args names);
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   266
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   267
fun lookup_case_names th =
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   268
  AList.lookup (op =) (Thm.get_tags th) case_names_tagN
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   269
  |> Option.map explode_args;
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   270
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   271
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
   272
val name = add_case_names o SOME;
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   273
fun case_names cs = Thm.mixed_attribute (apsnd (name cs));
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   274
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   275
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   276
(* hyp names *)
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   277
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   278
val implode_hyps = implode_args o map (suffix "," o space_implode ",");
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   279
val explode_hyps = map (space_explode "," o unsuffix ",") o explode_args;
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   280
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   281
val cases_hyp_names_tagN = "cases_hyp_names";
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   282
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   283
fun add_cases_hyp_names NONE = I
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   284
  | add_cases_hyp_names (SOME namess) =
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   285
      Thm.untag_rule cases_hyp_names_tagN
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   286
      #> Thm.tag_rule (cases_hyp_names_tagN, implode_hyps namess);
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   287
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   288
fun lookup_cases_hyp_names th =
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   289
  AList.lookup (op =) (Thm.get_tags th) cases_hyp_names_tagN
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   290
  |> Option.map explode_hyps;
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   291
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   292
val save_cases_hyp_names = add_cases_hyp_names o lookup_cases_hyp_names;
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   293
fun cases_hyp_names cs hs =
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   294
  Thm.mixed_attribute (apsnd (name cs #> add_cases_hyp_names (SOME hs)));
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 45328
diff changeset
   295
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   296
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   297
(* case conclusions *)
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   298
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   299
val case_concl_tagN = "case_conclusion";
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   300
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   301
fun get_case_concl name (a, b) =
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   302
  if a = case_concl_tagN then
58014
47ecbef7274b tuned -- more total;
wenzelm
parents: 56231
diff changeset
   303
    (case explode_args b of
47ecbef7274b tuned -- more total;
wenzelm
parents: 56231
diff changeset
   304
      c :: cs => if c = name then SOME cs else NONE
47ecbef7274b tuned -- more total;
wenzelm
parents: 56231
diff changeset
   305
    | [] => NONE)
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   306
  else NONE;
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   307
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 21576
diff changeset
   308
fun add_case_concl (name, cs) = Thm.map_tags (fn tags =>
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   309
  filter_out (is_some o get_case_concl name) tags @
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   310
    [(case_concl_tagN, implode_args (name :: cs))]);
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   311
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   312
fun get_case_concls th name =
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   313
  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
   314
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   315
fun save_case_concls th =
21646
c07b5b0e8492 thm/prf: separate official name vs. additional tags;
wenzelm
parents: 21576
diff changeset
   316
  let val concls = Thm.get_tags th |> map_filter
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   317
    (fn (a, b) =>
55639
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   318
      if a = case_concl_tagN
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   319
      then (case explode_args b of c :: cs => SOME (c, cs) | _ => NONE)
23657
2332c79f4dc8 thm tag: Markup.property list;
wenzelm
parents: 23584
diff changeset
   320
      else NONE)
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   321
  in fold add_case_concl concls end;
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   322
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   323
fun case_conclusion concl = Thm.mixed_attribute (apsnd (add_case_concl concl));
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   324
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   325
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   326
(* inner rule *)
33303
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 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
   329
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   330
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
   331
  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
   332
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   333
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
   334
  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
   335
  #> 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
   336
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   337
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
   338
45525
59561859c452 tagging is not stable under morphisms and need to be replayed dynamically (mixed_attribute);
wenzelm
parents: 45375
diff changeset
   339
val inner_rule = Thm.mixed_attribute (apsnd (put_inner_rule true));
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   340
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   341
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   342
(* parameter names *)
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   343
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   344
val cases_open_tagN = "cases_open";
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   345
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   346
fun is_cases_open th =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   347
  AList.defined (op =) (Thm.get_tags th) cases_open_tagN;
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   348
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   349
fun put_cases_open cases_open =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   350
  Thm.untag_rule cases_open_tagN
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   351
  #> cases_open ? Thm.tag_rule (cases_open_tagN, "");
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   352
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   353
val save_cases_open = put_cases_open o is_cases_open;
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   354
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   355
val cases_open = Thm.mixed_attribute (apsnd (put_cases_open true));
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   356
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   357
fun internalize_params rule =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   358
  if is_cases_open rule then rule
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   359
  else
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   360
    let
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   361
      fun rename prem =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   362
        let val xs =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   363
          map (Name.internal o Name.clean o fst) (Logic.strip_params prem)
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   364
        in Logic.list_rename_params xs prem end;
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   365
      fun rename_prems prop =
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   366
        let val (As, C) = Logic.strip_horn prop
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   367
        in Logic.list_implies (map rename As, C) end;
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   368
    in Thm.renamed_prop (rename_prems (Thm.prop_of rule)) rule end;
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   369
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   370
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   371
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   372
(** case declarations **)
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   373
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   374
(* access hints *)
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   375
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   376
fun save th =
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   377
  save_consumes th #>
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34916
diff changeset
   378
  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
   379
  save_case_names th #>
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   380
  save_cases_hyp_names th #>
33303
1e1210f31207 separate "inner_rule" tag indicates parts of induction rules -- avoids unclear overlap with "internal" tag;
wenzelm
parents: 31794
diff changeset
   381
  save_case_concls th #>
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   382
  save_inner_rule th #>
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60313
diff changeset
   383
  save_cases_open th;
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   384
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   385
fun get th =
12761
b0c39f9837af save: be slightly more about absent tags;
wenzelm
parents: 12166
diff changeset
   386
  let
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   387
    val n = get_consumes th;
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   388
    val cases =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   389
      (case lookup_case_names th of
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   390
        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
   391
      | SOME names => map (fn name => (name, get_case_concls th name)) names);
55639
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   392
    val cases_hyps =
e4e8cbd9d780 tuned whitespace;
wenzelm
parents: 54742
diff changeset
   393
      (case lookup_cases_hyp_names th of
44045
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   394
        NONE => replicate (length cases) []
2814ff2a6e3e infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
nipkow
parents: 42496
diff changeset
   395
      | SOME names => names);
60574
wenzelm
parents: 60565
diff changeset
   396
    fun regroup (name, concls) hyps = ((name, hyps), concls);
wenzelm
parents: 60565
diff changeset
   397
  in (map2 regroup cases cases_hyps, n) end;
10530
df079a585e6d added consumes, consumes_default;
wenzelm
parents: 10407
diff changeset
   398
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   399
8427
wenzelm
parents: 8400
diff changeset
   400
(* params *)
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   401
18229
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   402
fun rename_params xss th =
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   403
  th
e5a624483a23 added type cases/cases_tactic, and CASES, SUBGOAL_CASES;
wenzelm
parents: 18188
diff changeset
   404
  |> 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
   405
  |> save th;
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   406
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18608
diff changeset
   407
fun params xss = Thm.rule_attribute (K (rename_params xss));
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   408
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   409
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   410
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   411
(** mutual_rule **)
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   412
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   413
local
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
fun equal_cterms ts us =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58950
diff changeset
   416
  is_equal (list_ord (Term_Ord.fast_term_ord o apply2 Thm.term_of) (ts, us));
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   417
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   418
fun prep_rule n th =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   419
  let
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   420
    val th' = Thm.permute_prems 0 n th;
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   421
    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
   422
    val th'' = Drule.implies_elim_list th' (map Thm.assume prems);
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   423
  in (prems, (n, th'')) end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   424
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   425
in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   426
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   427
fun mutual_rule _ [] = NONE
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   428
  | mutual_rule _ [th] = SOME ([0], th)
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   429
  | mutual_rule ctxt (ths as th :: _) =
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   430
      let
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 30364
diff changeset
   431
        val ((_, ths'), ctxt') = Variable.import true ths ctxt;
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   432
        val rules as (prems, _) :: _ = map (prep_rule (get_consumes th)) ths';
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   433
        val (ns, rls) = split_list (map #2 rules);
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   434
      in
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   435
        if not (forall (equal_cterms prems o #1) rules) then NONE
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   436
        else
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   437
          SOME (ns,
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   438
            rls
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 22900
diff changeset
   439
            |> Conjunction.intr_balanced
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   440
            |> Drule.implies_intr_list prems
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   441
            |> singleton (Variable.export ctxt' ctxt)
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   442
            |> save th
18477
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   443
            |> put_consumes (SOME 0))
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   444
      end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   445
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   446
end;
bf2a02c82a55 consume: expand defs in prems of concls;
wenzelm
parents: 18375
diff changeset
   447
19917
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   448
fun strict_mutual_rule ctxt ths =
8b4869928f72 mutual_rule: proper context;
wenzelm
parents: 19482
diff changeset
   449
  (case mutual_rule ctxt ths of
18581
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   450
    NONE => error "Failed to join given rules into one mutual rule"
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   451
  | SOME res => res);
6538fdcc98dc added strict_mutual_rule;
wenzelm
parents: 18477
diff changeset
   452
8364
0eb9ee70c8f8 added Isar/rule_cases.ML;
wenzelm
parents:
diff changeset
   453
end;