author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 10 Mar 2010 09:58:43 +0100 | |
changeset 1387 | 9d70a0733786 |
parent 1386 | 0511e879a687 |
child 1388 | ad38ca4213f4 |
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 |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
2 |
imports "Nominal2_Atoms" "Abs" "Perm" "Rsp" |
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 |
|
1358 | 5 |
(* Bindings are a list of lists of lists of triples. |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
6 |
|
1358 | 7 |
The first list represents the datatypes defined. |
8 |
The second list represents the constructors. |
|
9 |
The internal list is a list of all the bndings that |
|
10 |
concern the constructor. |
|
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
11 |
|
1358 | 12 |
Every triple consists of a function, the binding and |
13 |
the body. |
|
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
14 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
15 |
Eg: |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
16 |
nominal_datatype |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
17 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
18 |
C1 |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
19 |
| C2 x y z bind x in z |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
20 |
| 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
|
21 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
22 |
yields: |
1172
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 |
[], |
1358 | 25 |
[(NONE, 0, 2)], |
26 |
[(SOME (Const f), 0, 2), (Some (Const g), 1, 2)]] |
|
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
27 |
|
1358 | 28 |
A SOME binding has to have a function which takes an appropriate |
29 |
argument and returns an atom set. A NONE binding has to be on an |
|
30 |
argument that is an atom or an atom set. |
|
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
31 |
|
1191
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
32 |
How the procedure works: |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
33 |
For each of the defined datatypes, |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
34 |
For each of the constructors, |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
35 |
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
|
36 |
|
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
37 |
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
|
38 |
bound variables. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
39 |
|
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
40 |
The variables are: |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
41 |
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
|
42 |
For an atom set, the atom set itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
43 |
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
|
44 |
(* TODO: This one is not implemented *) |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
45 |
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
|
46 |
in the database. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
47 |
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
|
48 |
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
|
49 |
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
|
50 |
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
|
51 |
For an atom set, the atom set itself. |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
52 |
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
|
53 |
(* TODO: This one is not implemented *) |
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
54 |
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
|
55 |
in the database. |
1358 | 56 |
|
57 |
For every argument that is a binding, we add a the same binding in its |
|
58 |
body. |
|
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
59 |
*) |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
60 |
|
1362
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
61 |
ML {* |
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
62 |
fun is_atom thy typ = |
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
63 |
Sign.of_sort thy (typ, @{sort at}) |
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
64 |
|
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
65 |
fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t |
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
66 |
| is_atom_set thy _ = false; |
1362
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
67 |
*} |
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
68 |
|
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
69 |
|
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
70 |
|
1358 | 71 |
(* Like map2, only if the second list is empty passes empty lists insted of error *) |
1302
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
72 |
ML {* |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
73 |
fun map2i _ [] [] = [] |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
74 |
| 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
|
75 |
| 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
|
76 |
| map2i _ _ _ = raise UnequalLengths; |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
77 |
*} |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
78 |
|
1358 | 79 |
(* Finds bindings with the same function and binding, and gathers all |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
80 |
bodys for such pairs |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
81 |
*) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
82 |
ML {* |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
83 |
fun gather_binds binds = |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
84 |
let |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
85 |
fun gather_binds_cons binds = |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
86 |
let |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
87 |
val common = map (fn (f, bi, _) => (f, bi)) binds |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
88 |
val nodups = distinct (op =) common |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
89 |
fun find_bodys (sf, sbi) = |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
90 |
filter (fn (f, bi, _) => f = sf andalso bi = sbi) binds |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
91 |
val bodys = map ((map (fn (_, _, bo) => bo)) o find_bodys) nodups |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
92 |
in |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
93 |
nodups ~~ bodys |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
94 |
end |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
95 |
in |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
96 |
map (map gather_binds_cons) binds |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
97 |
end |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
98 |
*} |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
99 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
100 |
ML {* |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
101 |
fun un_gather_binds_cons binds = |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
102 |
flat (map (fn (((f, bi), bos), pi) => map (fn bo => ((f, bi, bo), pi)) bos) binds) |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
103 |
*} |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
104 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
105 |
ML {* |
1175 | 106 |
open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *); |
1178 | 107 |
(* TODO: It is the same as one in 'nominal_atoms' *) |
1175 | 108 |
fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom}); |
109 |
val noatoms = @{term "{} :: atom set"}; |
|
110 |
fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]; |
|
111 |
fun mk_union sets = |
|
112 |
fold (fn a => fn b => |
|
113 |
if a = noatoms then b else |
|
114 |
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
|
115 |
if a = b then a else |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
116 |
HOLogic.mk_binop @{const_name sup} (a, b)) (rev sets) noatoms; |
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
117 |
val mk_inter = foldr1 (HOLogic.mk_binop @{const_name inf}) |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
118 |
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
|
119 |
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
|
120 |
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
|
121 |
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
|
122 |
HOLogic.mk_conj (a, b)) props @{term True}; |
1175 | 123 |
fun mk_diff a b = |
124 |
if b = noatoms then a else |
|
125 |
if b = a then noatoms else |
|
126 |
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
|
127 |
fun mk_atoms t = |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
128 |
let |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
129 |
val ty = fastype_of t; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
130 |
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
|
131 |
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
|
132 |
in |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
133 |
(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
|
134 |
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
|
135 |
(* 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
|
136 |
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
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
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
|
142 |
|
1175 | 143 |
*} |
144 |
||
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
145 |
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
|
146 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
147 |
ML {* |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
148 |
fun non_rec_binds l = |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
149 |
let |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
150 |
fun is_non_rec (SOME (f, false), _, _) = SOME f |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
151 |
| is_non_rec _ = NONE |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
152 |
in |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
153 |
distinct (op =) (map_filter is_non_rec (flat (flat l))) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
154 |
end |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
155 |
*} |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
156 |
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
157 |
(* We assume no bindings in the type on which bn is defined *) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
158 |
ML {* |
1379
5bb0149329ee
Separate lists for separate constructors, to match bn_eqs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1377
diff
changeset
|
159 |
fun fv_bn thy (dt_info : Datatype_Aux.info) fv_frees (bn, ith_dtyp, args_in_bns) = |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
160 |
let |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
161 |
val {descr, sorts, ...} = dt_info; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
162 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
163 |
val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
164 |
val fvbn = Free (fvbn_name, fastype_of (nth fv_frees ith_dtyp)); |
1379
5bb0149329ee
Separate lists for separate constructors, to match bn_eqs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1377
diff
changeset
|
165 |
fun fv_bn_constr (cname, dts) args_in_bn = |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
166 |
let |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
167 |
val Ts = map (typ_of_dtyp descr sorts) dts; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
168 |
val names = Datatype_Prop.make_tnames Ts; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
169 |
val args = map Free (names ~~ Ts); |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
170 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
171 |
fun fv_arg ((dt, x), arg_no) = |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
172 |
let |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
173 |
val ty = fastype_of x |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
174 |
in |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
175 |
if arg_no mem args_in_bn then |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
176 |
(if is_rec_type dt then |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
177 |
(if body_index dt = ith_dtyp then fvbn $ x else error "fv_bn: not good") |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
178 |
else @{term "{} :: atom set"}) else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
179 |
if is_atom thy ty then mk_single_atom x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
180 |
if is_atom_set thy ty then mk_atoms x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
181 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
182 |
@{term "{} :: atom set"} |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
183 |
end; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
184 |
val arg_nos = 0 upto (length dts - 1) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
185 |
in |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
186 |
HOLogic.mk_Trueprop (HOLogic.mk_eq |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
187 |
(fvbn $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
188 |
end; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
189 |
val (_, (_, _, constrs)) = nth descr ith_dtyp; |
1379
5bb0149329ee
Separate lists for separate constructors, to match bn_eqs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1377
diff
changeset
|
190 |
val eqs = map2i fv_bn_constr constrs args_in_bns |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
191 |
in |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
192 |
((bn, fvbn), (fvbn_name, eqs)) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
193 |
end |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
194 |
*} |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
195 |
|
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
196 |
ML {* |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
197 |
fun alpha_bn thy (dt_info : Datatype_Aux.info) alpha_frees ((bn, ith_dtyp, args_in_bns), is_rec) = |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
198 |
let |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
199 |
val {descr, sorts, ...} = dt_info; |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
200 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
201 |
val alpha_bn_name = "alpha_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
202 |
val alpha_bn_type = @{typ perm} --> nth_dtyp ith_dtyp --> nth_dtyp ith_dtyp --> @{typ bool} |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
203 |
val alpha_bn_free = Free(alpha_bn_name, alpha_bn_type); |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
204 |
fun alpha_bn_constr (cname, dts) args_in_bn = |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
205 |
let |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
206 |
val Ts = map (typ_of_dtyp descr sorts) dts; |
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
207 |
val pi = Free("pi", @{typ perm}) |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
208 |
val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
209 |
val names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
210 |
val args = map Free (names ~~ Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
211 |
val args2 = map Free (names2 ~~ Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
212 |
val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); |
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
213 |
val rhs = HOLogic.mk_Trueprop |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
214 |
(alpha_bn_free $ pi $ (list_comb (c, args)) $ (list_comb (c, args2))); |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
215 |
fun lhs_arg ((dt, arg_no), (arg, arg2)) = |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
216 |
(*...*) |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
217 |
@{term True} |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
218 |
val arg_nos = 0 upto (length dts - 1) |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
219 |
val lhss = mk_conjl (map lhs_arg (dts ~~ arg_nos ~~ (args ~~ args2))) |
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
220 |
val eq = Logic.mk_implies (HOLogic.mk_Trueprop lhss, rhs) |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
221 |
in |
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
222 |
eq |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
223 |
end |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
224 |
val (_, (_, _, constrs)) = nth descr ith_dtyp; |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
225 |
val eqs = map2i alpha_bn_constr constrs args_in_bns |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
226 |
in |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
227 |
((bn, alpha_bn_name), (alpha_bn_free, eqs)) |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
228 |
end |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
229 |
*} |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
230 |
|
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
231 |
ML {* fn x => split_list(flat x) *} |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
232 |
ML {* fn x => apsnd flat (split_list (map split_list x)) *} |
1206 | 233 |
(* TODO: Notice datatypes without bindings and replace alpha with equality *) |
1175 | 234 |
ML {* |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
235 |
fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall bns lthy = |
1178 | 236 |
let |
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
237 |
val thy = ProofContext.theory_of lthy; |
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1258
diff
changeset
|
238 |
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
|
239 |
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
|
240 |
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
|
241 |
"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
|
242 |
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
|
243 |
val fv_frees = map Free (fv_names ~~ fv_types); |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
244 |
val nr_bns = non_rec_binds bindsall; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
245 |
val rel_bns = filter (fn (bn, _, _) => bn mem nr_bns) bns; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
246 |
val (bn_fv_bns, fv_bn_names_eqs) = split_list (map (fv_bn thy dt_info fv_frees) rel_bns); |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
247 |
val (fv_bn_names, fv_bn_eqs) = split_list fv_bn_names_eqs; |
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
|
248 |
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
|
249 |
"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
|
250 |
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
|
251 |
val alpha_frees = map Free (alpha_names ~~ alpha_types); |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
252 |
(* We assume that a bn is either recursive or not *) |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
253 |
val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns; |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
254 |
|
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
255 |
val alpha_bnfrees = map (fst o snd) (map (alpha_bn thy dt_info alpha_frees) (bns ~~ bns_rec)) |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
256 |
|
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
257 |
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
|
258 |
let |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
259 |
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
|
260 |
val bindslen = length bindcs |
1323
acf262971303
Fixes for the fv problem and alpha problem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1308
diff
changeset
|
261 |
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
|
262 |
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
|
263 |
val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
264 |
val bind_pis_gath = bindcs ~~ pis; |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
265 |
val bind_pis = un_gather_binds_cons bind_pis_gath; |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
266 |
val bindcs = map fst bind_pis; |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
267 |
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
|
268 |
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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
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
|
273 |
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
|
274 |
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
|
275 |
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
|
276 |
if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else |
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
277 |
if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else |
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
278 |
if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atoms (nth args i) else |
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
279 |
(* TODO we do not know what to do with non-atomizable things *) |
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
280 |
@{term "{} :: atom set"} |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
281 |
| fv_bind args (SOME (f, _), i, _) = f $ (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
|
282 |
fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
283 |
fun find_nonrec_binder j (SOME (f, false), i, _) = if i = j then SOME f else NONE |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
284 |
| find_nonrec_binder _ _ = NONE |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
285 |
fun fv_arg ((dt, x), arg_no) = |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
286 |
case get_first (find_nonrec_binder arg_no) bindcs of |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
287 |
SOME f => |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
288 |
(case get_first (fn (x, y) => if x = f then SOME y else NONE) bn_fv_bns of |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
289 |
SOME fv_bn => fv_bn $ x |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
290 |
| NONE => error "bn specified in a non-rec binding but not in bn list") |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
291 |
| NONE => |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
292 |
let |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
293 |
val arg = |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
294 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
295 |
if ((is_atom thy) o fastype_of) x then mk_single_atom x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
296 |
if ((is_atom_set thy) o fastype_of) x then mk_atoms x else |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
297 |
(* TODO we do not know what to do with non-atomizable things *) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
298 |
@{term "{} :: atom set"} |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
299 |
(* If i = j then we generate it only once *) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
300 |
val relevant = filter (fn (_, i, j) => ((i = arg_no) orelse (j = arg_no))) bindcs; |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
301 |
val sub = fv_binds args relevant |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
302 |
in |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
303 |
mk_diff arg sub |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
304 |
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
|
305 |
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
|
306 |
(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
|
307 |
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
|
308 |
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
|
309 |
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
|
310 |
let |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
311 |
val rel_in_simp_binds = filter (fn ((NONE, i, _), _) => i = arg_no | _ => false) bind_pis; |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
312 |
val rel_in_comp_binds = filter (fn ((SOME _, i, _), _) => i = arg_no | _ => false) bind_pis; |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
313 |
val rel_has_binds = filter (fn ((SOME _, _, j), _) => j = arg_no | _ => false) bind_pis; |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
314 |
in |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
315 |
case (rel_in_simp_binds, rel_in_comp_binds, rel_has_binds) of |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
316 |
([], [], []) => |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
317 |
if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2) |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
318 |
else (HOLogic.mk_eq (arg, arg2)) |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
319 |
(* TODO: if more then check and accept *) |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
320 |
| (_, [], []) => @{term True} |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
321 |
| ([], [(((SOME (bn, _)), _, _), pi)], []) => |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
322 |
nth alpha_bnfrees (find_index (fn (b, _, _) => b = bn) bns) $ pi $ arg $ arg2 |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
323 |
| ([], [], relevant) => |
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
324 |
let |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
325 |
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
|
326 |
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
|
327 |
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
|
328 |
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
|
329 |
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
|
330 |
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
|
331 |
val fv = nth fv_frees (body_index dt); |
1359
3bf496a971c6
Fix permutation addition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1358
diff
changeset
|
332 |
val pi = foldr1 add_perm (distinct (op =) rpis); |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
333 |
val alpha_gen_pre = Const (@{const_name alpha_gen}, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
334 |
val alpha_gen = Syntax.check_term lthy alpha_gen_pre |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
335 |
(* val pi_supps = map ((curry op $) @{term "supp :: perm \<Rightarrow> atom set"}) rpis; |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
336 |
val pi_supps_eq = HOLogic.mk_eq (mk_inter pi_supps, @{term "{} :: atom set"}) *) |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
337 |
in |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
338 |
(* if length pi_supps > 1 then HOLogic.mk_conj (alpha_gen, pi_supps_eq) else*) |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
339 |
alpha_gen |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
340 |
end |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
341 |
| _ => error "Fv.alpha: not supported binding structure" |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
342 |
end |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
343 |
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
|
344 |
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
|
345 |
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
|
346 |
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
|
347 |
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
|
348 |
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
|
349 |
(fv_eq, alpha_eq) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
350 |
end; |
1302
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
351 |
fun fv_alpha_eq (i, (_, _, constrs)) binds = map2i (fv_alpha_constr i) constrs binds; |
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
352 |
val fveqs_alphaeqs = map2i fv_alpha_eq descr (gather_binds bindsall) |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
353 |
val (fv_eqs_perfv, alpha_eqs) = apsnd flat (split_list (map split_list fveqs_alphaeqs)) |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
354 |
val rel_bns_nos = map (fn (_, i, _) => i) rel_bns; |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
355 |
fun filter_fun (_, b) = b mem rel_bns_nos; |
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
356 |
val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1)) |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
357 |
val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs))) |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
358 |
val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs))) |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
359 |
val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs); |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
360 |
val fv_names_all = fv_names_fst @ fv_bn_names; |
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
|
361 |
val add_binds = map (fn x => (Attrib.empty_binding, x)) |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
362 |
(* Function_Fun.add_fun Function_Common.default_config ... true *) |
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
363 |
val _ = map tracing (map (Syntax.string_of_term @{context}) fv_eqs_all) |
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
|
364 |
val (fvs, lthy') = (Primrec.add_primrec |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
365 |
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy) |
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
366 |
val (fvs2, lthy'') = |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
367 |
if fv_eqs_snd = [] then (([], []), lthy') else |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
368 |
(Primrec.add_primrec |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
369 |
(map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy') |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
370 |
val (alphas, lthy''') = (Inductive.add_inductive_i |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
371 |
{quiet_mode = true, verbose = false, alt_name = Binding.empty, |
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
|
372 |
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
|
373 |
(map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) alpha_names alpha_types) [] |
1376
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
374 |
(add_binds alpha_eqs) [] lthy'') |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
375 |
val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) |
1178 | 376 |
in |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
377 |
((all_fvs, alphas), lthy''') |
1178 | 378 |
end |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
379 |
*} |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
380 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
381 |
(* |
1178 | 382 |
atom_decl name |
383 |
||
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
384 |
datatype lam = |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
385 |
VAR "name" |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
386 |
| APP "lam" "lam" |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
387 |
| LET "bp" "lam" |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
388 |
and bp = |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
389 |
BP "name" "lam" |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
390 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
391 |
primrec |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
392 |
bi::"bp \<Rightarrow> atom set" |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
393 |
where |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
394 |
"bi (BP x t) = {atom x}" |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
395 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
396 |
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Fv.lam") 2 *} |
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
|
397 |
|
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
398 |
|
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
399 |
local_setup {* |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
400 |
snd o define_fv_alpha (Datatype.the_info @{theory} "Fv.lam") |
1379
5bb0149329ee
Separate lists for separate constructors, to match bn_eqs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1377
diff
changeset
|
401 |
[[[], [], [(SOME (@{term bi}, false), 0, 1)]], [[]]] [(@{term bi}, 1, [[0]])] *} |
1375
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
402 |
print_theorems |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
403 |
*) |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
404 |
|
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
405 |
(* |
1178 | 406 |
datatype rtrm1 = |
407 |
rVr1 "name" |
|
408 |
| rAp1 "rtrm1" "rtrm1" |
|
409 |
| rLm1 "name" "rtrm1" --"name is bound in trm1" |
|
410 |
| rLt1 "bp" "rtrm1" "rtrm1" --"all variables in bp are bound in the 2nd trm1" |
|
411 |
and bp = |
|
412 |
BUnit |
|
413 |
| BVr "name" |
|
414 |
| BPr "bp" "bp" |
|
415 |
||
416 |
(* to be given by the user *) |
|
417 |
||
1358 | 418 |
primrec |
1178 | 419 |
bv1 |
420 |
where |
|
421 |
"bv1 (BUnit) = {}" |
|
422 |
| "bv1 (BVr x) = {atom x}" |
|
423 |
| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)" |
|
424 |
||
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
|
425 |
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
|
426 |
|
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
|
427 |
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
|
428 |
[[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]], |
1178 | 429 |
[[], [[]], [[], []]]] *} |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
430 |
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
|
431 |
*) |
1180 | 432 |
|
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
433 |
|
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
434 |
ML {* |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
435 |
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
|
436 |
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
|
437 |
(rtac @{thm iffI} THEN' RANGE [ |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
438 |
(eresolve_tac elims THEN_ALL_NEW |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
439 |
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
|
440 |
), |
1216
06ace3a1eedd
Fixed pseudo_injectivity for trm4
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1215
diff
changeset
|
441 |
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
|
442 |
*} |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
443 |
|
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
444 |
ML {* |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
445 |
fun build_alpha_inj_gl thm = |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
446 |
let |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
447 |
val prop = prop_of thm; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
448 |
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
|
449 |
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
|
450 |
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
|
451 |
in |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
452 |
if hyps = [] then concl |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
453 |
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
|
454 |
end; |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
455 |
*} |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
456 |
|
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
457 |
ML {* |
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
458 |
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
|
459 |
let |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
460 |
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
|
461 |
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
|
462 |
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
|
463 |
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
|
464 |
in |
1199
5770c73f2415
Automatic production and proving of pseudo-injectivity.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1196
diff
changeset
|
465 |
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
|
466 |
end |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
467 |
*} |
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
468 |
|
1207 | 469 |
ML {* |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
470 |
fun build_alpha_refl_gl alphas (x, y, z) = |
1207 | 471 |
let |
472 |
fun build_alpha alpha = |
|
473 |
let |
|
474 |
val ty = domain_type (fastype_of alpha); |
|
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
475 |
val var = Free(x, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
476 |
val var2 = Free(y, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
477 |
val var3 = Free(z, ty); |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
478 |
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
|
479 |
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
|
480 |
HOLogic.mk_all (z, ty, |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
481 |
HOLogic.mk_imp (alpha $ var2 $ var3, alpha $ var $ var3))) |
1207 | 482 |
in |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
483 |
((alpha $ var $ var), (symp, transp)) |
1208
cc86faf0d2a0
alpha-symmetric addons.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1207
diff
changeset
|
484 |
end; |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
485 |
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
|
486 |
val (sym_eqs, trans_eqs) = split_list eqs |
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
487 |
fun conj l = @{term Trueprop} $ foldr1 HOLogic.mk_conj l |
1207 | 488 |
in |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
489 |
(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
|
490 |
end |
1207 | 491 |
*} |
492 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
493 |
ML {* |
1333
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
494 |
fun reflp_tac induct inj ctxt = |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
495 |
rtac induct THEN_ALL_NEW |
1333
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
496 |
simp_tac ((mk_minimal_ss ctxt) addsimps inj) THEN_ALL_NEW |
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
497 |
split_conjs THEN_ALL_NEW REPEAT o rtac @{thm exI[of _ "0 :: perm"]} |
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
498 |
THEN_ALL_NEW split_conjs THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps |
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
499 |
@{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv |
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
500 |
add_0_left supp_zero_perm Int_empty_left}) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
501 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
502 |
|
1333
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
503 |
|
1301 | 504 |
lemma exi_neg: "\<exists>(pi :: perm). P pi \<Longrightarrow> (\<And>(p :: perm). P p \<Longrightarrow> Q (- p)) \<Longrightarrow> \<exists>pi. Q pi" |
505 |
apply (erule exE) |
|
506 |
apply (rule_tac x="-pi" in exI) |
|
507 |
by auto |
|
508 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
509 |
ML {* |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
510 |
fun symp_tac induct inj eqvt ctxt = |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
511 |
ind_tac induct THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
512 |
simp_tac ((mk_minimal_ss ctxt) addsimps inj) THEN_ALL_NEW split_conjs |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
513 |
THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
514 |
REPEAT o etac @{thm exi_neg} |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
515 |
THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
516 |
split_conjs THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
517 |
asm_full_simp_tac (HOL_ss addsimps @{thms supp_minus_perm minus_add[symmetric]}) THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
518 |
(rtac @{thm alpha_gen_compose_sym} THEN' RANGE [ |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
519 |
(asm_full_simp_tac (HOL_ss addsimps @{thms plus_perm_eq})), |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
520 |
(asm_full_simp_tac (HOL_ss addsimps (eqvt @ all_eqvts ctxt))) |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
521 |
]) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
522 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
523 |
|
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
524 |
ML {* |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
525 |
fun imp_elim_tac case_rules = |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
526 |
Subgoal.FOCUS (fn {concl, context, ...} => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
527 |
case term_of concl of |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
528 |
_ $ (_ $ asm $ _) => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
529 |
let |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
530 |
fun filter_fn case_rule = ( |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
531 |
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
|
532 |
((_ $ asmc) :: _) => |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
533 |
let |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
534 |
val thy = ProofContext.theory_of context |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
535 |
in |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
536 |
Pattern.matches thy (asmc, asm) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
537 |
end |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
538 |
| _ => false) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
539 |
val matching_rules = filter filter_fn case_rules |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
540 |
in |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
541 |
(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
|
542 |
end |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
543 |
| _ => no_tac |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
544 |
) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
545 |
*} |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
546 |
|
1301 | 547 |
|
548 |
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" |
|
549 |
apply (erule exE)+ |
|
550 |
apply (rule_tac x="pia + pi" in exI) |
|
551 |
by auto |
|
552 |
||
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
553 |
ML {* |
1339
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
554 |
fun is_ex (Const ("Ex", _) $ Abs _) = true |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
555 |
| is_ex _ = false; |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
556 |
*} |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
557 |
|
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
558 |
ML {* |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
559 |
fun eetac rule = Subgoal.FOCUS_PARAMS |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
560 |
(fn (focus) => |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
561 |
let |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
562 |
val concl = #concl focus |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
563 |
val prems = Logic.strip_imp_prems (term_of concl) |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
564 |
val exs = filter (fn x => is_ex (HOLogic.dest_Trueprop x)) prems |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
565 |
val cexs = map (SOME o (cterm_of (ProofContext.theory_of (#context focus)))) exs |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
566 |
val thins = map (fn cex => Drule.instantiate' [] [cex] Drule.thin_rl) cexs |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
567 |
in |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
568 |
(etac rule THEN' RANGE[ |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
569 |
atac, |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
570 |
eresolve_tac thins |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
571 |
]) 1 |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
572 |
end |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
573 |
) |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
574 |
*} |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
575 |
|
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
576 |
ML {* |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
577 |
fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt = |
1339
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
578 |
ind_tac induct THEN_ALL_NEW |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
579 |
(TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW |
1339
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
580 |
asm_full_simp_tac ((mk_minimal_ss ctxt) addsimps alpha_inj) THEN_ALL_NEW |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
581 |
split_conjs THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) THEN_ALL_NEW split_conjs |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
582 |
THEN_ALL_NEW (asm_full_simp_tac (HOL_ss addsimps (term_inj @ distinct))) |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
583 |
THEN_ALL_NEW split_conjs THEN_ALL_NEW |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
584 |
TRY o (etac @{thm alpha_gen_compose_trans} THEN' RANGE[atac]) THEN_ALL_NEW |
5256f256edd8
Comment out Weird and Phd until we have an idea how to handle multiple permutations. Transp that works for multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1338
diff
changeset
|
585 |
(asm_full_simp_tac (HOL_ss addsimps (all_eqvts ctxt @ eqvt @ term_inj @ distinct))) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
586 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
587 |
|
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
588 |
lemma transp_aux: |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
589 |
"(\<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
|
590 |
unfolding transp_def |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
591 |
by blast |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
592 |
|
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
593 |
ML {* |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
594 |
fun equivp_tac reflps symps transps = |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
595 |
simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def}) |
1221 | 596 |
THEN' rtac conjI THEN' rtac allI THEN' |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
597 |
resolve_tac reflps THEN' |
1221 | 598 |
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
|
599 |
resolve_tac symps THEN' |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
600 |
rtac @{thm transp_aux} THEN' resolve_tac transps |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
601 |
*} |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
602 |
|
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
603 |
ML {* |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
604 |
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
|
605 |
let |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
606 |
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
|
607 |
val (reflg, (symg, transg)) = build_alpha_refl_gl alphas (x, y, z) |
1333
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
608 |
fun reflp_tac' _ = reflp_tac term_induct alpha_inj ctxt 1; |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
609 |
fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt ctxt 1; |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
610 |
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
|
611 |
val reflt = Goal.prove ctxt' [] [] reflg reflp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
612 |
val symt = Goal.prove ctxt' [] [] symg symp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
613 |
val transt = Goal.prove ctxt' [] [] transg transp_tac'; |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
614 |
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
|
615 |
val reflts = HOLogic.conj_elims refltg |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
616 |
val symts = HOLogic.conj_elims symtg |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
617 |
val transts = HOLogic.conj_elims transtg |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
618 |
fun equivp alpha = |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
619 |
let |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
620 |
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
|
621 |
val goal = @{term Trueprop} $ (equivp $ alpha) |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
622 |
fun tac _ = equivp_tac reflts symts transts 1 |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
623 |
in |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
624 |
Goal.prove ctxt [] [] goal tac |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
625 |
end |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
626 |
in |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
627 |
map equivp alphas |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
628 |
end |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
629 |
*} |
1207 | 630 |
|
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
631 |
(* |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
632 |
Tests: |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
633 |
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
|
634 |
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
|
635 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
636 |
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
|
637 |
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
|
638 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
639 |
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
|
640 |
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
|
641 |
|
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
642 |
lemma alpha1_equivp: |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
643 |
"equivp alpha_rtrm1" |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
644 |
"equivp alpha_bp" |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
645 |
apply (tactic {* |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
646 |
(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
|
647 |
THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN' |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
648 |
resolve_tac (HOLogic.conj_elims @{thm alpha1_reflp_aux}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
649 |
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
|
650 |
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
|
651 |
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
|
652 |
) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
653 |
1 *}) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
654 |
done*) |
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
655 |
|
1308
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
656 |
ML {* |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
657 |
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
|
658 |
| 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
|
659 |
| dtyp_no_of_typ dts (Type (tname, Ts)) = |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
660 |
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
|
661 |
NONE => error "dtyp_no_of_typ: Illegal recursion" |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
662 |
| SOME i => i |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
663 |
*} |
80dabcaafc38
Moving wrappers out of Lift.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1303
diff
changeset
|
664 |
|
1207 | 665 |
end |