-
Notifications
You must be signed in to change notification settings - Fork 0
/
Requalify.thy
225 lines (162 loc) · 6.44 KB
/
Requalify.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
(*
* Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
*
* SPDX-License-Identifier: BSD-2-Clause
*)
(*
Requalify constants, types and facts into the current naming
*)
theory Requalify
imports Main
keywords "requalify_facts" :: thy_decl and "requalify_types" :: thy_decl and "requalify_consts" :: thy_decl and
"global_naming" :: thy_decl
begin
ML \<open>
local
fun all_facts_of ctxt =
let
val thy = Proof_Context.theory_of ctxt;
val global_facts = Global_Theory.facts_of thy;
in
Facts.dest_static false [] global_facts
end;
fun global_fact ctxt nm =
let
val facts = Proof_Context.facts_of ctxt;
val {name, thms, ...} = (Facts.retrieve (Context.Proof ctxt) facts (nm, Position.none));
fun tl' (_ :: xs) = xs
| tl' _ = []
fun matches suf (gnm, gthms) =
let
val gsuf = Long_Name.explode gnm |> tl' |> tl' |> Long_Name.implode;
in suf = gsuf andalso eq_list (Thm.equiv_thm (Proof_Context.theory_of ctxt)) (thms, gthms)
end
in
case Long_Name.dest_local name of NONE => (name, thms) | SOME suf =>
(case (find_first (matches suf) (all_facts_of ctxt)) of
SOME x => x
| NONE => raise Fail ("Couldn't find global equivalent of local fact: " ^ nm))
end
fun syntax_alias global_alias local_alias b (name : string) =
Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
let val b' = Morphism.binding phi b
in Context.mapping (global_alias b' name) (local_alias b' name) end);
val alias_fact = syntax_alias Global_Theory.alias_fact Proof_Context.alias_fact;
val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias;
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;
in
fun gen_requalify get_proper_nm parse_nm check_nm alias =
(Parse.opt_target -- Scan.repeat1 (Parse.position (Scan.ahead parse_nm -- Parse.name))
>> (fn (target,bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun read_entry ((entry, t), pos) lthy =
let
val local_nm = get_proper_nm lthy t;
val _ = check_nm lthy (entry, (local_nm, pos));
val b = Binding.make (Long_Name.base_name t, pos)
val lthy' = lthy
|> alias b local_nm
in lthy' end
in fold read_entry bs lthy end)))
local
val get_const_nm = ((fst o dest_Const) oo (Proof_Context.read_const {proper = true, strict = false}))
val get_type_nm = ((fst o dest_Type) oo (Proof_Context.read_type_name {proper = true, strict = false}))
val get_fact_nm = (fst oo global_fact)
fun check_fact lthy (_, (nm, pos)) = Proof_Context.get_fact lthy (Facts.Named ((nm,pos), NONE))
in
val _ =
Outer_Syntax.command @{command_keyword requalify_consts} "alias const with current naming"
(gen_requalify get_const_nm Parse.const (fn lthy => fn (e, _) => get_const_nm lthy e) const_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_types} "alias type with current naming"
(gen_requalify get_type_nm Parse.typ (fn lthy => fn (e, _) => get_type_nm lthy e) type_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_facts} "alias fact with current naming"
(gen_requalify get_fact_nm Parse.thm check_fact alias_fact)
val _ =
Outer_Syntax.command @{command_keyword global_naming} "change global naming of context block"
(Parse.binding >> (fn naming =>
Toplevel.local_theory NONE NONE
(Local_Theory.map_background_naming (Name_Space.parent_path #> Name_Space.qualified_path true naming))))
end
end
\<close>
(*Tests and examples *)
locale Requalify_Locale
begin
typedecl requalify_type
definition "requalify_const == (undefined :: requalify_type)"
end
typedecl requalify_type
definition "requalify_const == (undefined :: requalify_type)"
context Requalify_Locale begin global_naming Requalify_Locale2
requalify_consts requalify_const
requalify_types requalify_type
requalify_facts requalify_const_def
end
term "requalify_const :: requalify_type"
term "Requalify_Locale2.requalify_const :: Requalify_Locale2.requalify_type"
lemma "Requalify_Locale2.requalify_const = (undefined :: Requalify_Locale2.requalify_type)"
by (simp add: Requalify_Locale2.requalify_const_def)
consts requalify_test_f :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const Requalify.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
context Requalify_Locale begin
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const Requalify.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f2)?
apply (rule f1)
apply (simp add: requalify_const_def)
done
end
context Requalify_Locale begin global_naming global
requalify_consts Requalify.requalify_const
requalify_types Requalify.requalify_type
requalify_facts Requalify.requalify_const_def
end
context Requalify_Locale begin
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const global.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
end
context begin interpretation Requalify_Locale .
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const global.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
end
locale Requalify_Locale3
begin
typedecl requalify_type2
definition "requalify_const2 == (undefined :: requalify_type2)"
end
context begin interpretation Requalify_Locale3 .
requalify_consts requalify_const2
requalify_types requalify_type2
requalify_facts requalify_const2_def
end
lemma "(requalify_const2 :: requalify_type2) = undefined"
by (simp add: requalify_const2_def)
context Requalify_Locale3 begin
lemma "(requalify_const2 :: requalify_type2) = undefined"
by (simp add: requalify_const2_def)
end
end