author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 74561 | 8e6c973003c8 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
1 |
(* Title: HOL/Nominal/nominal_thmdecls.ML |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
2 |
Author: Julien Narboux, TU Muenchen |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
3 |
Author: Christian Urban, TU Muenchen |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
4 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
5 |
Infrastructure for the lemma collection "eqvts". |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
6 |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
7 |
By attaching [eqvt] or [eqvt_force] to a lemma, it will get stored in |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
8 |
a data-slot in the context. Possible modifiers are [... add] and |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
9 |
[... del] for adding and deleting, respectively, the lemma from the |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32740
diff
changeset
|
10 |
data-slot. |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
11 |
*) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
12 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
13 |
signature NOMINAL_THMDECLS = |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
14 |
sig |
54990 | 15 |
val nominal_eqvt_debug: bool Config.T |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
16 |
val eqvt_add: attribute |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
17 |
val eqvt_del: attribute |
22715
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
urbanc
parents:
22562
diff
changeset
|
18 |
val eqvt_force_add: attribute |
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
urbanc
parents:
22562
diff
changeset
|
19 |
val eqvt_force_del: attribute |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
20 |
val setup: theory -> theory |
24571 | 21 |
val get_eqvt_thms: Proof.context -> thm list |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
22 |
end; |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
23 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
24 |
structure NominalThmDecls: NOMINAL_THMDECLS = |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
25 |
struct |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
26 |
|
33519 | 27 |
structure Data = Generic_Data |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
28 |
( |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
29 |
type T = thm list |
33519 | 30 |
val empty = [] |
31 |
val merge = Thm.merge_thms |
|
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
32 |
) |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
33 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
34 |
(* Exception for when a theorem does not conform with form of an equivariance lemma. *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
35 |
(* There are two forms: one is an implication (for relations) and the other is an *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
36 |
(* equality (for functions). In the implication-case, say P ==> Q, Q must be equal *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
37 |
(* to P except that every free variable of Q, say x, is replaced by pi o x. In the *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
38 |
(* equality case, say lhs = rhs, the lhs must be of the form pi o t and the rhs must *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
39 |
(* be equal to t except that every free variable, say x, is replaced by pi o x. In *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
40 |
(* the implicational case it is also checked that the variables and permutation fit *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
41 |
(* together, i.e. are of the right "pt_class", so that a stronger version of the *) |
24265
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
narboux
parents:
24039
diff
changeset
|
42 |
(* equality-lemma can be derived. *) |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
43 |
exception EQVT_FORM of string |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
44 |
|
69597 | 45 |
val nominal_eqvt_debug = Attrib.setup_config_bool \<^binding>\<open>nominal_eqvt_debug\<close> (K false); |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
46 |
|
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
47 |
fun tactic ctxt (msg, tac) = |
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
48 |
if Config.get ctxt nominal_eqvt_debug |
56491 | 49 |
then tac THEN' (K (print_tac ctxt ("after " ^ msg))) |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
50 |
else tac |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
51 |
|
30991 | 52 |
fun prove_eqvt_tac ctxt orig_thm pi pi' = |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
53 |
let |
42361 | 54 |
val thy = Proof_Context.theory_of ctxt |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
55 |
val T = fastype_of pi' |
69597 | 56 |
val mypifree = Thm.cterm_of ctxt (Const (\<^const_name>\<open>rev\<close>, T --> T) $ pi') |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
38864
diff
changeset
|
57 |
val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp" |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
58 |
in |
60754 | 59 |
EVERY1 [tactic ctxt ("iffI applied", resolve_tac ctxt @{thms iffI}), |
60 |
tactic ctxt ("remove pi with perm_boolE", dresolve_tac ctxt @{thms perm_boolE}), |
|
61 |
tactic ctxt ("solve with orig_thm", eresolve_tac ctxt [orig_thm]), |
|
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
62 |
tactic ctxt ("applies orig_thm instantiated with rev pi", |
60787 | 63 |
dresolve_tac ctxt [infer_instantiate ctxt [(#1 (dest_Var pi), mypifree)] orig_thm]), |
60754 | 64 |
tactic ctxt ("getting rid of the pi on the right", resolve_tac ctxt @{thms perm_boolI}), |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
65 |
tactic ctxt ("getting rid of all remaining perms", |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
42616
diff
changeset
|
66 |
full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps perm_pi_simp))] |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
67 |
end; |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
68 |
|
30088
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
69 |
fun get_derived_thm ctxt hyp concl orig_thm pi typi = |
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
70 |
let |
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
71 |
val pi' = Var (pi, typi); |
69597 | 72 |
val lhs = Const (\<^const_name>\<open>perm\<close>, typi --> HOLogic.boolT --> HOLogic.boolT) $ pi' $ hyp; |
30088
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
73 |
val ([goal_term, pi''], ctxt') = Variable.import_terms false |
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
74 |
[HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, concl)), pi'] ctxt |
32429
54758ca53fd6
modernized messages -- eliminated old Display.print_cterm;
wenzelm
parents:
32091
diff
changeset
|
75 |
val _ = writeln (Syntax.string_of_term ctxt' goal_term); |
30088
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
76 |
in |
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
77 |
Goal.prove ctxt' [] [] goal_term |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
78 |
(fn _ => prove_eqvt_tac ctxt' orig_thm pi' pi'') |> |
42361 | 79 |
singleton (Proof_Context.export ctxt' ctxt) |
30088
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
80 |
end |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
81 |
|
30991 | 82 |
(* replaces in t every variable, say x, with pi o x *) |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
83 |
fun apply_pi trm (pi, typi) = |
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
84 |
let |
30991 | 85 |
fun replace n ty = |
86 |
let |
|
69597 | 87 |
val c = Const (\<^const_name>\<open>perm\<close>, typi --> ty --> ty) |
30991 | 88 |
val v1 = Var (pi, typi) |
89 |
val v2 = Var (n, ty) |
|
90 |
in |
|
91 |
c $ v1 $ v2 |
|
92 |
end |
|
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
93 |
in |
30991 | 94 |
map_aterms (fn Var (n, ty) => replace n ty | t => t) trm |
95 |
end |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
96 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
97 |
(* returns *the* pi which is in front of all variables, provided there *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
98 |
(* exists such a pi; otherwise raises EQVT_FORM *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
99 |
fun get_pi t thy = |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
100 |
let fun get_pi_aux s = |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
101 |
(case s of |
69597 | 102 |
(Const (\<^const_name>\<open>perm\<close> ,typrm) $ |
103 |
(Var (pi,typi as Type(\<^type_name>\<open>list\<close>, [Type (\<^type_name>\<open>Product_Type.prod\<close>, [Type (tyatm,[]),_])]))) $ |
|
22846 | 104 |
(Var (n,ty))) => |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
105 |
let |
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
22846
diff
changeset
|
106 |
(* FIXME: this should be an operation the library *) |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30280
diff
changeset
|
107 |
val class_name = (Long_Name.map_base_name (fn s => "pt_"^s) tyatm) |
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
22846
diff
changeset
|
108 |
in |
24271 | 109 |
if (Sign.of_sort thy (ty,[class_name])) |
22846 | 110 |
then [(pi,typi)] |
111 |
else raise |
|
112 |
EQVT_FORM ("Could not find any permutation or an argument is not an instance of "^class_name) |
|
113 |
end |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
114 |
| Abs (_,_,t1) => get_pi_aux t1 |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
115 |
| (t1 $ t2) => get_pi_aux t1 @ get_pi_aux t2 |
22846 | 116 |
| _ => []) |
117 |
in |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
118 |
(* collect first all pi's in front of variables in t and then use distinct *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
119 |
(* to ensure that all pi's must have been the same, i.e. distinct returns *) |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
120 |
(* a singleton-list *) |
22846 | 121 |
(case (distinct (op =) (get_pi_aux t)) of |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
122 |
[(pi,typi)] => (pi, typi) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22286
diff
changeset
|
123 |
| _ => raise EQVT_FORM "All permutation should be the same") |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
124 |
end; |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
125 |
|
26652 | 126 |
(* Either adds a theorem (orig_thm) to or deletes one from the equivariance *) |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
127 |
(* lemma list depending on flag. To be added the lemma has to satisfy a *) |
22846 | 128 |
(* certain form. *) |
24265
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
narboux
parents:
24039
diff
changeset
|
129 |
|
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
narboux
parents:
24039
diff
changeset
|
130 |
fun eqvt_add_del_aux flag orig_thm context = |
22846 | 131 |
let |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
132 |
val thy = Context.theory_of context |
59582 | 133 |
val thms_to_be_added = |
134 |
(case Thm.prop_of orig_thm of |
|
22846 | 135 |
(* case: eqvt-lemma is of the implicational form *) |
69597 | 136 |
(Const(\<^const_name>\<open>Pure.imp\<close>, _) $ (Const (\<^const_name>\<open>Trueprop\<close>,_) $ hyp) $ (Const (\<^const_name>\<open>Trueprop\<close>,_) $ concl)) => |
22846 | 137 |
let |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
138 |
val (pi,typi) = get_pi concl thy |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
139 |
in |
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
140 |
if (apply_pi hyp (pi,typi) = concl) |
22846 | 141 |
then |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
142 |
(warning ("equivariance lemma of the relational form"); |
30088
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
143 |
[orig_thm, |
fe6eac03b816
Replaced Logic.unvarify by Variable.import_terms to make declaration of
berghofe
parents:
29585
diff
changeset
|
144 |
get_derived_thm (Context.proof_of context) hyp concl orig_thm pi typi]) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22286
diff
changeset
|
145 |
else raise EQVT_FORM "Type Implication" |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
146 |
end |
22846 | 147 |
(* case: eqvt-lemma is of the equational form *) |
69597 | 148 |
| (Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ |
149 |
(Const (\<^const_name>\<open>perm\<close>,typrm) $ Var (pi,typi) $ lhs) $ rhs)) => |
|
24039
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents:
22846
diff
changeset
|
150 |
(if (apply_pi lhs (pi,typi)) = rhs |
22437
b3526cd2a336
removes some unused code that used to try to derive a simpler version of the eqvt lemmas
narboux
parents:
22418
diff
changeset
|
151 |
then [orig_thm] |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22286
diff
changeset
|
152 |
else raise EQVT_FORM "Type Equality") |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22286
diff
changeset
|
153 |
| _ => raise EQVT_FORM "Type unknown") |
22846 | 154 |
in |
26652 | 155 |
fold (fn thm => Data.map (flag thm)) thms_to_be_added context |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
156 |
end |
22846 | 157 |
handle EQVT_FORM s => |
61268 | 158 |
error (Thm.string_of_thm (Context.proof_of context) orig_thm ^ |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
159 |
" does not comply with the form of an equivariance lemma (" ^ s ^").") |
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
160 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
161 |
|
30991 | 162 |
val eqvt_add = Thm.declaration_attribute (eqvt_add_del_aux (Thm.add_thm)); |
163 |
val eqvt_del = Thm.declaration_attribute (eqvt_add_del_aux (Thm.del_thm)); |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
164 |
|
30991 | 165 |
val eqvt_force_add = Thm.declaration_attribute (Data.map o Thm.add_thm); |
166 |
val eqvt_force_del = Thm.declaration_attribute (Data.map o Thm.del_thm); |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
167 |
|
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
168 |
val get_eqvt_thms = Context.Proof #> Data.get; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22286
diff
changeset
|
169 |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
170 |
val setup = |
69597 | 171 |
Attrib.setup \<^binding>\<open>eqvt\<close> (Attrib.add_del eqvt_add eqvt_del) |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
172 |
"equivariance theorem declaration" #> |
69597 | 173 |
Attrib.setup \<^binding>\<open>eqvt_force\<close> (Attrib.add_del eqvt_force_add eqvt_force_del) |
38807
378ffc903bed
eliminated Unsynchronized.ref in favour of configuration option;
wenzelm
parents:
38558
diff
changeset
|
174 |
"equivariance theorem declaration (without checking the form of the lemma)" #> |
69597 | 175 |
Global_Theory.add_thms_dynamic (\<^binding>\<open>eqvts\<close>, Data.get); |
30986
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents:
30528
diff
changeset
|
176 |
|
22245
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
177 |
|
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents:
diff
changeset
|
178 |
end; |