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