author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 02 Mar 2010 21:10:04 +0100 | |
changeset 1323 | acf262971303 |
parent 1308 | 80dabcaafc38 |
child 1325 | 0be098c61d00 |
permissions | -rw-r--r-- |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Fv |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
2 |
imports "Nominal2_Atoms" "Abs" "Perm" (* For testing *) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
5 |
(* Bindings are given as a list which has a length being equal |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
6 |
to the length of the number of constructors. |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
7 |
|
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
8 |
Each element is a list whose length is equal to the number |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
9 |
of arguents. |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
10 |
|
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
11 |
Every element specifies bindings of this argument given as |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
12 |
a tuple: function, bound argument. |
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
13 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
14 |
Eg: |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
15 |
nominal_datatype |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
16 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
17 |
C1 |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
18 |
| C2 x y z bind x in z |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
19 |
| C3 x y z bind f x in z bind g y in z |
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
20 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
21 |
yields: |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
22 |
[ |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
23 |
[], |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
24 |
[[], [], [(NONE, 0)]], |
1173
9cb99a28b40e
Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1172
diff
changeset
|
25 |
[[], [], [(SOME (Const f), 0), (Some (Const g), 1)]]] |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
26 |
|
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
27 |
A SOME binding has to have a function returning an atom set, |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
28 |
and a NONE binding has to be on an argument that is an atom |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
29 |
or an atom set. |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
30 |
|
1191
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
31 |
How the procedure works: |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
32 |
For each of the defined datatypes, |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
33 |
For each of the constructors, |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
34 |
It creates a union of free variables for each argument. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
35 |
|
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
36 |
For an argument the free variables are the variables minus |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
37 |
bound variables. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
38 |
|
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
39 |
The variables are: |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
40 |
For an atom, a singleton set with the atom itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
41 |
For an atom set, the atom set itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
42 |
For a recursive argument, the appropriate fv function applied to it. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
43 |
(* TODO: This one is not implemented *) |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
44 |
For other arguments it should be an appropriate fv function stored |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
45 |
in the database. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
46 |
The bound variables are a union of results of all bindings that |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
47 |
involve the given argument. For a paricular binding the result is: |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
48 |
For a function applied to an argument this function with the argument. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
49 |
For an atom, a singleton set with the atom itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
50 |
For an atom set, the atom set itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
51 |
For a recursive argument, the appropriate fv function applied to it. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
52 |
(* TODO: This one is not implemented *) |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
53 |
For other arguments it should be an appropriate fv function stored |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
54 |
in the database. |
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
55 |
*) |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
56 |
|
1302
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
57 |
ML {* |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
58 |
fun map2i _ [] [] = [] |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
59 |
| map2i f (x :: xs) (y :: ys) = f x y :: map2i f xs ys |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
60 |
| map2i f (x :: xs) [] = f x [] :: map2i f xs [] |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
61 |
| map2i _ _ _ = raise UnequalLengths; |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
62 |
*} |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
63 |
|
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
64 |
(* TODO: should be const_name union *) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
65 |
ML {* |
1175 | 66 |
open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *); |
1178 | 67 |
(* TODO: It is the same as one in 'nominal_atoms' *) |
1175 | 68 |
fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom}); |
69 |
val noatoms = @{term "{} :: atom set"}; |
|
70 |
fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]; |
|
71 |
fun mk_union sets = |
|
72 |
fold (fn a => fn b => |
|
73 |
if a = noatoms then b else |
|
74 |
if b = noatoms then a else |
|
1323
acf262971303
Fixes for the fv problem and alpha problem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1308
diff
changeset
|
75 |
if a = b then a else |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
76 |
HOLogic.mk_binop "Set.union" (a, b)) (rev sets) noatoms; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
77 |
fun mk_conjl props = |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
78 |
fold (fn a => fn b => |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
79 |
if a = @{term True} then b else |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
80 |
if b = @{term True} then a else |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
81 |
HOLogic.mk_conj (a, b)) props @{term True}; |
1175 | 82 |
fun mk_diff a b = |
83 |
if b = noatoms then a else |
|
84 |
if b = a then noatoms else |
|
85 |
HOLogic.mk_binop @{const_name minus} (a, b); |
|
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
86 |
fun mk_atoms t = |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
87 |
let |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
88 |
val ty = fastype_of t; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
89 |
val atom_ty = HOLogic.dest_setT ty --> @{typ atom}; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
90 |
val img_ty = atom_ty --> ty --> @{typ "atom set"}; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
91 |
in |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
92 |
(Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t) |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
93 |
end; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
94 |
(* Copy from Term *) |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
95 |
fun is_funtype (Type ("fun", [_, _])) = true |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
96 |
| is_funtype _ = false; |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
97 |
(* Similar to one in USyntax *) |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
98 |
fun mk_pair (fst, snd) = |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
99 |
let val ty1 = fastype_of fst |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
100 |
val ty2 = fastype_of snd |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
101 |
val c = HOLogic.pair_const ty1 ty2 |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
102 |
in c $ fst $ snd |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
103 |
end; |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
104 |
|
1175 | 105 |
*} |
106 |
||
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
107 |
ML {* fun add_perm (p1, p2) = Const(@{const_name plus}, @{typ "perm \<Rightarrow> perm \<Rightarrow> perm"}) $ p1 $ p2 *} |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
108 |
|
1206 | 109 |
(* TODO: Notice datatypes without bindings and replace alpha with equality *) |
1175 | 110 |
ML {* |
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1258
diff
changeset
|
111 |
fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall lthy = |
1178 | 112 |
let |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
113 |
val thy = ProofContext.theory_of lthy; |
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1258
diff
changeset
|
114 |
val {descr, sorts, ...} = dt_info; |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
116 |
val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
117 |
"fv_" ^ name_of_typ (nth_dtyp i)) descr); |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
118 |
val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr; |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
119 |
val fv_frees = map Free (fv_names ~~ fv_types); |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
120 |
val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) => |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
121 |
"alpha_" ^ name_of_typ (nth_dtyp i)) descr); |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
122 |
val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr; |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
123 |
val alpha_frees = map Free (alpha_names ~~ alpha_types); |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
124 |
fun fv_alpha_constr ith_dtyp (cname, dts) bindcs = |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
125 |
let |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
126 |
val Ts = map (typ_of_dtyp descr sorts) dts; |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
127 |
val bindslen = length bindcs |
1323
acf262971303
Fixes for the fv problem and alpha problem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1308
diff
changeset
|
128 |
val pi_strs_same = replicate bindslen "pi" |
acf262971303
Fixes for the fv problem and alpha problem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1308
diff
changeset
|
129 |
val pi_strs = Name.variant_list [] pi_strs_same; |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
130 |
val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
131 |
val bind_pis = bindcs ~~ pis; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
132 |
val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
133 |
val args = map Free (names ~~ Ts); |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
134 |
val names2 = Name.variant_list (pi_strs @ names) (Datatype_Prop.make_tnames Ts); |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
135 |
val args2 = map Free (names2 ~~ Ts); |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
136 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
137 |
val fv_c = nth fv_frees ith_dtyp; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
138 |
val alpha = nth alpha_frees ith_dtyp; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
139 |
val arg_nos = 0 upto (length dts - 1) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
140 |
fun fv_bind args (NONE, i, _) = |
1177
6f01720fe520
Add bindings of recursive types by free_variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1176
diff
changeset
|
141 |
if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else |
6f01720fe520
Add bindings of recursive types by free_variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1176
diff
changeset
|
142 |
(* TODO we assume that all can be 'atomized' *) |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
143 |
if (is_funtype o fastype_of) (nth args i) then mk_atoms (nth args i) else |
1177
6f01720fe520
Add bindings of recursive types by free_variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1176
diff
changeset
|
144 |
mk_single_atom (nth args i) |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
145 |
| fv_bind args (SOME f, i, _) = f $ (nth args i); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
146 |
fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
147 |
fun fv_arg ((dt, x), arg_no) = |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
148 |
let |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
149 |
val arg = |
1175 | 150 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
1174 | 151 |
(* TODO: we just assume everything can be 'atomized' *) |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
152 |
if (is_funtype o fastype_of) x then mk_atoms x else |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
153 |
HOLogic.mk_set @{typ atom} [mk_atom (fastype_of x) $ x]; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
154 |
(* If i = j then we generate it only once *) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
155 |
val relevant = filter (fn (_, i, j) => ((i = arg_no) orelse (j = arg_no))) bindcs; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
156 |
val sub = fv_binds args relevant |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
157 |
in |
1173
9cb99a28b40e
Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1172
diff
changeset
|
158 |
mk_diff arg sub |
9cb99a28b40e
Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1172
diff
changeset
|
159 |
end; |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
160 |
val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
161 |
(fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
162 |
val alpha_rhs = |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
163 |
HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2))); |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
164 |
fun alpha_arg ((dt, arg_no), (arg, arg2)) = |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
165 |
let |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
166 |
val relevant = filter (fn ((_, i, j), _) => i = arg_no orelse j = arg_no) bind_pis; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
167 |
in |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
168 |
if relevant = [] then ( |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
169 |
if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
170 |
else (HOLogic.mk_eq (arg, arg2))) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
171 |
else |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
172 |
if is_rec_type dt then let |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
173 |
(* THE HARD CASE *) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
174 |
val (rbinds, rpis) = split_list relevant |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
175 |
val lhs_binds = fv_binds args rbinds |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
176 |
val lhs = mk_pair (lhs_binds, arg); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
177 |
val rhs_binds = fv_binds args2 rbinds; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
178 |
val rhs = mk_pair (rhs_binds, arg2); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
179 |
val alpha = nth alpha_frees (body_index dt); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
180 |
val fv = nth fv_frees (body_index dt); |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
181 |
(* TODO: check that pis have empty intersection *) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
182 |
val pi = foldr1 add_perm rpis; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
183 |
val alpha_gen_pre = Const (@{const_name alpha_gen}, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
184 |
in |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
185 |
Syntax.check_term lthy alpha_gen_pre |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
186 |
(* TODO Add some test that is makes sense *) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
187 |
end else @{term "True"} |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
188 |
end |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
189 |
val alphas = map alpha_arg (dts ~~ arg_nos ~~ (args ~~ args2)) |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
190 |
val alpha_lhss = mk_conjl alphas |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
191 |
val alpha_lhss_ex = |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
192 |
fold (fn pi_str => fn t => HOLogic.mk_exists (pi_str, @{typ perm}, t)) pi_strs alpha_lhss |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
193 |
val alpha_eq = Logic.mk_implies (HOLogic.mk_Trueprop alpha_lhss_ex, alpha_rhs) |
1173
9cb99a28b40e
Some optimizations and fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1172
diff
changeset
|
194 |
in |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
195 |
(fv_eq, alpha_eq) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
196 |
end; |
1302
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
197 |
fun fv_alpha_eq (i, (_, _, constrs)) binds = map2i (fv_alpha_constr i) constrs binds; |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
198 |
val (fv_eqs, alpha_eqs) = split_list (flat (map2i fv_alpha_eq descr bindsall)) |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
199 |
val add_binds = map (fn x => (Attrib.empty_binding, x)) |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
200 |
val (fvs, lthy') = (Primrec.add_primrec |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
201 |
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) (add_binds fv_eqs) lthy) |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
202 |
val (alphas, lthy'') = (Inductive.add_inductive_i |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
203 |
{quiet_mode = false, verbose = true, alt_name = Binding.empty, |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
204 |
coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false} |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
205 |
(map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) alpha_names alpha_types) [] |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
206 |
(add_binds alpha_eqs) [] lthy') |
1178 | 207 |
in |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
208 |
((fvs, alphas), lthy'') |
1178 | 209 |
end |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
210 |
*} |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
211 |
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
212 |
(* tests |
1178 | 213 |
atom_decl name |
214 |
||
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
215 |
datatype ty = |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
216 |
Var "name set" |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
217 |
|
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
218 |
ML {* Syntax.check_term @{context} (mk_atoms @{term "a :: name set"}) *} |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
219 |
|
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
220 |
local_setup {* define_fv_alpha "Fv.ty" [[[[]]]] *} |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
221 |
print_theorems |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
222 |
|
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
223 |
|
1178 | 224 |
datatype rtrm1 = |
225 |
rVr1 "name" |
|
226 |
| rAp1 "rtrm1" "rtrm1" |
|
227 |
| rLm1 "name" "rtrm1" --"name is bound in trm1" |
|
228 |
| rLt1 "bp" "rtrm1" "rtrm1" --"all variables in bp are bound in the 2nd trm1" |
|
229 |
and bp = |
|
230 |
BUnit |
|
231 |
| BVr "name" |
|
232 |
| BPr "bp" "bp" |
|
233 |
||
234 |
(* to be given by the user *) |
|
235 |
||
236 |
primrec |
|
237 |
bv1 |
|
238 |
where |
|
239 |
"bv1 (BUnit) = {}" |
|
240 |
| "bv1 (BVr x) = {atom x}" |
|
241 |
| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)" |
|
242 |
||
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
243 |
setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Fv.rtrm1", "Fv.bp"] *} |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
244 |
|
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
245 |
local_setup {* define_fv_alpha "Fv.rtrm1" |
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
246 |
[[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]], |
1178 | 247 |
[[], [[]], [[], []]]] *} |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
248 |
print_theorems |
1193
a228acf2907e
Full alpha equivalence + testing in terms. Some differ but it seems the generated version is more correct.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1192
diff
changeset
|
249 |
*) |
1180 | 250 |
|
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
251 |
|
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
252 |
ML {* |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
253 |
fun alpha_inj_tac dist_inj intrs elims = |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
254 |
SOLVED' (asm_full_simp_tac (HOL_ss addsimps intrs)) ORELSE' |
1216
06ace3a1eedd
Fixed pseudo_injectivity for trm4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1215
diff
changeset
|
255 |
(rtac @{thm iffI} THEN' RANGE [ |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
256 |
(eresolve_tac elims THEN_ALL_NEW |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
257 |
asm_full_simp_tac (HOL_ss addsimps dist_inj) |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
258 |
), |
1216
06ace3a1eedd
Fixed pseudo_injectivity for trm4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1215
diff
changeset
|
259 |
asm_full_simp_tac (HOL_ss addsimps intrs)]) |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
260 |
*} |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
261 |
|
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
262 |
ML {* |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
263 |
fun build_alpha_inj_gl thm = |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
264 |
let |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
265 |
val prop = prop_of thm; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
266 |
val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop); |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
267 |
val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop); |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
268 |
fun list_conj l = foldr1 HOLogic.mk_conj l; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
269 |
in |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
270 |
if hyps = [] then concl |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
271 |
else HOLogic.mk_eq (concl, list_conj hyps) |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
272 |
end; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
273 |
*} |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
274 |
|
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
275 |
ML {* |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
276 |
fun build_alpha_inj intrs dist_inj elims ctxt = |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
277 |
let |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
278 |
val ((_, thms_imp), ctxt') = Variable.import false intrs ctxt; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
279 |
val gls = map (HOLogic.mk_Trueprop o build_alpha_inj_gl) thms_imp; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
280 |
fun tac _ = alpha_inj_tac dist_inj intrs elims 1; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
281 |
val thms = map (fn gl => Goal.prove ctxt' [] [] gl tac) gls; |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
282 |
in |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
283 |
Variable.export ctxt' ctxt thms |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
284 |
end |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
285 |
*} |
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
286 |
|
1207 | 287 |
ML {* |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
288 |
fun build_alpha_refl_gl alphas (x, y, z) = |
1207 | 289 |
let |
290 |
fun build_alpha alpha = |
|
291 |
let |
|
292 |
val ty = domain_type (fastype_of alpha); |
|
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
293 |
val var = Free(x, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
294 |
val var2 = Free(y, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
295 |
val var3 = Free(z, ty); |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
296 |
val symp = HOLogic.mk_imp (alpha $ var $ var2, alpha $ var2 $ var); |
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
297 |
val transp = HOLogic.mk_imp (alpha $ var $ var2, |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
298 |
HOLogic.mk_all (z, ty, |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
299 |
HOLogic.mk_imp (alpha $ var2 $ var3, alpha $ var $ var3))) |
1207 | 300 |
in |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
301 |
((alpha $ var $ var), (symp, transp)) |
1208
cc86faf0d2a0
alpha-symmetric addons.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1207
diff
changeset
|
302 |
end; |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
303 |
val (refl_eqs, eqs) = split_list (map build_alpha alphas) |
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
304 |
val (sym_eqs, trans_eqs) = split_list eqs |
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
305 |
fun conj l = @{term Trueprop} $ foldr1 HOLogic.mk_conj l |
1207 | 306 |
in |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
307 |
(conj refl_eqs, (conj sym_eqs, conj trans_eqs)) |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
308 |
end |
1207 | 309 |
*} |
310 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
311 |
ML {* |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
312 |
fun reflp_tac induct inj = |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
313 |
rtac induct THEN_ALL_NEW |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
314 |
asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW |
1301 | 315 |
(* TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW*) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
316 |
(rtac @{thm exI[of _ "0 :: perm"]} THEN' |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
317 |
asm_full_simp_tac (HOL_ss addsimps |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
318 |
@{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv})) |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
319 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
320 |
|
1301 | 321 |
lemma exi_neg: "\<exists>(pi :: perm). P pi \<Longrightarrow> (\<And>(p :: perm). P p \<Longrightarrow> Q (- p)) \<Longrightarrow> \<exists>pi. Q pi" |
322 |
apply (erule exE) |
|
323 |
apply (rule_tac x="-pi" in exI) |
|
324 |
by auto |
|
325 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
326 |
ML {* |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
327 |
fun symp_tac induct inj eqvt = |
1301 | 328 |
(((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
329 |
asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1302
diff
changeset
|
330 |
(REPEAT o etac conjE THEN' etac @{thm exi_neg} THEN' REPEAT o etac conjE THEN' |
1301 | 331 |
(TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI))) THEN_ALL_NEW |
332 |
(asm_full_simp_tac HOL_ss) THEN_ALL_NEW |
|
333 |
(etac @{thm alpha_gen_compose_sym} THEN' |
|
334 |
(asm_full_simp_tac (HOL_ss addsimps (@{thm atom_eqvt} :: eqvt))))) |
|
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
335 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
336 |
|
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
337 |
ML {* |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
338 |
fun imp_elim_tac case_rules = |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
339 |
Subgoal.FOCUS (fn {concl, context, ...} => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
340 |
case term_of concl of |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
341 |
_ $ (_ $ asm $ _) => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
342 |
let |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
343 |
fun filter_fn case_rule = ( |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
344 |
case Logic.strip_assums_hyp (prop_of case_rule) of |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
345 |
((_ $ asmc) :: _) => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
346 |
let |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
347 |
val thy = ProofContext.theory_of context |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
348 |
in |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
349 |
Pattern.matches thy (asmc, asm) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
350 |
end |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
351 |
| _ => false) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
352 |
val matching_rules = filter filter_fn case_rules |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
353 |
in |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
354 |
(rtac impI THEN' rotate_tac (~1) THEN' eresolve_tac matching_rules) 1 |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
355 |
end |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
356 |
| _ => no_tac |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
357 |
) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
358 |
*} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
359 |
|
1301 | 360 |
|
361 |
lemma exi_sum: "\<exists>(pi :: perm). P pi \<Longrightarrow> \<exists>(pi :: perm). Q pi \<Longrightarrow> (\<And>(p :: perm) (pi :: perm). P p \<Longrightarrow> Q pi \<Longrightarrow> R (pi + p)) \<Longrightarrow> \<exists>pi. R pi" |
|
362 |
apply (erule exE)+ |
|
363 |
apply (rule_tac x="pia + pi" in exI) |
|
364 |
by auto |
|
365 |
||
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
366 |
ML {* |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
367 |
fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt = |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
368 |
((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
369 |
(TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
370 |
( |
1301 | 371 |
asm_full_simp_tac (HOL_ss addsimps alpha_inj @ term_inj @ distinct) |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1302
diff
changeset
|
372 |
THEN_ALL_NEW (REPEAT o etac conjE THEN' etac @{thm exi_sum} THEN' RANGE [atac]) THEN_ALL_NEW |
1301 | 373 |
(REPEAT o etac conjE THEN' (TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI))) |
374 |
THEN_ALL_NEW (asm_full_simp_tac HOL_ss) THEN_ALL_NEW |
|
375 |
(etac @{thm alpha_gen_compose_trans} THEN' RANGE[atac]) THEN_ALL_NEW |
|
376 |
(asm_full_simp_tac (HOL_ss addsimps (@{thm atom_eqvt} :: eqvt))) |
|
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
377 |
) |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
378 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
379 |
|
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
380 |
lemma transp_aux: |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
381 |
"(\<And>xa ya. R xa ya \<longrightarrow> (\<forall>z. R ya z \<longrightarrow> R xa z)) \<Longrightarrow> transp R" |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
382 |
unfolding transp_def |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
383 |
by blast |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
384 |
|
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
385 |
ML {* |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
386 |
fun equivp_tac reflps symps transps = |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
387 |
simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def}) |
1221 | 388 |
THEN' rtac conjI THEN' rtac allI THEN' |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
389 |
resolve_tac reflps THEN' |
1221 | 390 |
rtac conjI THEN' rtac allI THEN' rtac allI THEN' |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
391 |
resolve_tac symps THEN' |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
392 |
rtac @{thm transp_aux} THEN' resolve_tac transps |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
393 |
*} |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
394 |
|
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
395 |
ML {* |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
396 |
fun build_equivps alphas term_induct alpha_induct term_inj alpha_inj distinct cases eqvt ctxt = |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
397 |
let |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
398 |
val ([x, y, z], ctxt') = Variable.variant_fixes ["x","y","z"] ctxt; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
399 |
val (reflg, (symg, transg)) = build_alpha_refl_gl alphas (x, y, z) |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
400 |
fun reflp_tac' _ = reflp_tac term_induct alpha_inj 1; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
401 |
fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt 1; |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
402 |
fun transp_tac' _ = transp_tac ctxt alpha_induct alpha_inj term_inj distinct cases eqvt 1; |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
403 |
val reflt = Goal.prove ctxt' [] [] reflg reflp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
404 |
val symt = Goal.prove ctxt' [] [] symg symp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
405 |
val transt = Goal.prove ctxt' [] [] transg transp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
406 |
val [refltg, symtg, transtg] = Variable.export ctxt' ctxt [reflt, symt, transt] |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
407 |
val reflts = HOLogic.conj_elims refltg |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
408 |
val symts = HOLogic.conj_elims symtg |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
409 |
val transts = HOLogic.conj_elims transtg |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
410 |
fun equivp alpha = |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
411 |
let |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
412 |
val equivp = Const (@{const_name equivp}, fastype_of alpha --> @{typ bool}) |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
413 |
val goal = @{term Trueprop} $ (equivp $ alpha) |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
414 |
fun tac _ = equivp_tac reflts symts transts 1 |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
415 |
in |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
416 |
Goal.prove ctxt [] [] goal tac |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
417 |
end |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
418 |
in |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
419 |
map equivp alphas |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
420 |
end |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
421 |
*} |
1207 | 422 |
|
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
423 |
(* |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
424 |
Tests: |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
425 |
prove alpha1_reflp_aux: {* fst (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
426 |
by (tactic {* reflp_tac @{thm rtrm1_bp.induct} @{thms alpha1_inj} 1 *}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
427 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
428 |
prove alpha1_symp_aux: {* (fst o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
429 |
by (tactic {* symp_tac @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms alpha1_eqvt} 1 *}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
430 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
431 |
prove alpha1_transp_aux: {* (snd o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
432 |
by (tactic {* transp_tac @{context} @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms rtrm1.inject bp.inject} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} 1 *}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
433 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
434 |
lemma alpha1_equivp: |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
435 |
"equivp alpha_rtrm1" |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
436 |
"equivp alpha_bp" |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
437 |
apply (tactic {* |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
438 |
(simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
439 |
THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN' |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
440 |
resolve_tac (HOLogic.conj_elims @{thm alpha1_reflp_aux}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
441 |
THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
442 |
resolve_tac (HOLogic.conj_elims @{thm alpha1_symp_aux}) THEN' rtac @{thm transp_aux} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
443 |
THEN' resolve_tac (HOLogic.conj_elims @{thm alpha1_transp_aux}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
444 |
) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
445 |
1 *}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
446 |
done*) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
447 |
|
1308
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
448 |
ML {* |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
449 |
fun dtyp_no_of_typ _ (TFree (n, _)) = error "dtyp_no_of_typ: Illegal free" |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
450 |
| dtyp_no_of_typ _ (TVar _) = error "dtyp_no_of_typ: Illegal schematic" |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
451 |
| dtyp_no_of_typ dts (Type (tname, Ts)) = |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
452 |
case try (find_index (curry op = tname o fst)) dts of |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
453 |
NONE => error "dtyp_no_of_typ: Illegal recursion" |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
454 |
| SOME i => i |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
455 |
*} |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
456 |
|
1207 | 457 |
end |