author | Christian Urban <urbanc@in.tum.de> |
Sun, 11 Apr 2010 18:11:23 +0200 | |
changeset 1806 | 90095f23fc60 |
parent 1774 | c34347ec7ab3 |
child 1830 | 8db45a106569 |
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 |
1806 | 2 |
imports "../Nominal-General/Nominal2_Atoms" |
3 |
"Abs" "Perm" "Rsp" "Nominal2_FSet" |
|
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
|
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
6 |
(* The bindings data structure: |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
7 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
8 |
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
|
9 |
|
1358 | 10 |
The first list represents the datatypes defined. |
11 |
The second list represents the constructors. |
|
12 |
The internal list is a list of all the bndings that |
|
13 |
concern the constructor. |
|
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
14 |
|
1358 | 15 |
Every triple consists of a function, the binding and |
16 |
the body. |
|
1169
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 |
Eg: |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
19 |
nominal_datatype |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
20 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
21 |
C1 |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
22 |
| C2 x y z bind x in z |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
23 |
| 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
|
24 |
|
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
25 |
yields: |
1172
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
26 |
[ |
9a609fefcf24
Simplified format of bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1169
diff
changeset
|
27 |
[], |
1358 | 28 |
[(NONE, 0, 2)], |
29 |
[(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
|
30 |
|
1358 | 31 |
A SOME binding has to have a function which takes an appropriate |
32 |
argument and returns an atom set. A NONE binding has to be on an |
|
33 |
argument that is an atom or an atom set. |
|
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
34 |
*) |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
35 |
|
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
36 |
(* |
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1508
diff
changeset
|
37 |
An overview of the generation of free variables: |
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
38 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
39 |
1) fv_bn functions are generated only for the non-recursive binds. |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
40 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
41 |
An fv_bn for a constructor is a union of values for the arguments: |
1191
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
42 |
|
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
43 |
For an argument x that is in the bn function |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
44 |
- if it is a recursive argument bn' we return: fv_bn' x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
45 |
- otherwise empty |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
46 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
47 |
For an argument x that is not in the bn function |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
48 |
- for atom we return: {atom x} |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
49 |
- for atom set we return: atom ` x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
50 |
- for a recursive call to type ty' we return: fv_ty' x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
51 |
with fv of the appropriate type |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
52 |
- otherwise empty |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
53 |
|
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
54 |
2) fv_ty functions generated for all types being defined: |
1191
15362b433d64
Description of the fv procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1185
diff
changeset
|
55 |
|
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
56 |
fv_ty for a constructor is a union of values for the arguments. |
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
57 |
|
1508
883b38196dba
Simplified the description.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1505
diff
changeset
|
58 |
For an argument that is bound in a shallow binding we return empty. |
883b38196dba
Simplified the description.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1505
diff
changeset
|
59 |
|
883b38196dba
Simplified the description.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1505
diff
changeset
|
60 |
For an argument x that bound in a non-recursive deep binding |
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
61 |
we return: fv_bn x. |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
62 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
63 |
Otherwise we return the free variables of the argument minus the |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
64 |
bound variables of the argument. |
1358 | 65 |
|
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
66 |
The free variables for an argument x are: |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
67 |
- for an atom: {atom x} |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
68 |
- for atom set: atom ` x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
69 |
- for recursive call to type ty' return: fv_ty' x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
70 |
- for nominal datatype ty' return: fv_ty' x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
71 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
72 |
The bound variables are a union of results of all bindings that |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
73 |
involve the given argument. For a paricular binding: |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
74 |
|
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
75 |
- for a binding function bn: bn x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
76 |
- for a recursive argument of type ty': fv_fy' x |
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
77 |
- for nominal datatype ty' return: fv_ty' x |
1169
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
78 |
*) |
b9d02e0800e9
Description of intended bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1168
diff
changeset
|
79 |
|
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1508
diff
changeset
|
80 |
(* |
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1508
diff
changeset
|
81 |
An overview of the generation of alpha-equivalence: |
1513
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
82 |
|
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
83 |
1) alpha_bn relations are generated for binding functions. |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
84 |
|
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
85 |
An alpha_bn for a constructor is true if a conjunction of |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
86 |
propositions for each argument holds. |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
87 |
|
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
88 |
For an argument a proposition is build as follows from |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
89 |
th: |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
90 |
|
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
91 |
- for a recursive argument in the bn function, we return: alpha_bn argl argr |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
92 |
- for a recursive argument for type ty not in bn, we return: alpha_ty argl argr |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
93 |
- for other arguments in the bn function we return: True |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
94 |
- for other arguments not in the bn function we return: argl = argr |
44840614ea0c
Description of generation of alpha_bn.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
95 |
|
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
96 |
2) alpha_ty relations are generated for all the types being defined: |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
97 |
|
1516
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
98 |
For each constructor we gather all the arguments that are bound, |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
99 |
and for each of those we add a permutation. We associate those |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
100 |
permutations with the bindings. Note that two bindings can have |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
101 |
the same permutation if the arguments being bound are the same. |
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
102 |
|
1516
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
103 |
An alpha_ty for a constructor is true if there exist permutations |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
104 |
as above such that a conjunction of propositions for all arguments holds. |
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
105 |
|
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
106 |
For an argument we allow bindings where only one of the following |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
107 |
holds: |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
108 |
|
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
109 |
- Argument is bound in some shallow bindings: We return true |
1516
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
110 |
- Argument of type ty is bound recursively in some other |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
111 |
arguments [i1, .. in] with one binding function bn. |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
112 |
We return: |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
113 |
|
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
114 |
(bn argl, (argl, argl_i1, ..., argl_in)) \<approx>gen |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
115 |
\<lambda>(argl,argl1,..,argln) (argr,argr1,..,argrn). |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
116 |
(alpha_ty argl argr) \<and> (alpha_i1 argl1 argr1) \<and> .. \<and> (alpha_in argln argrn) |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
117 |
\<lambda>(arg,arg1,..,argn). (fv_ty arg) \<union> (fv_i1 arg1) \<union> .. \<union> (fv_in argn) |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
118 |
pi |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
119 |
(bn argr, (argr, argr_i1, ..., argr_in)) |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
120 |
|
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
121 |
- Argument is bound in some deep non-recursive bindings. |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
122 |
We return: alpha_bn argl argr |
1516
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
123 |
- Argument of type ty has some shallow bindings [b1..bn] and/or |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
124 |
non-recursive bindings [f1 a1, .., fm am], where the bindings |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
125 |
have the permutations p1..pl. We return: |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
126 |
|
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
127 |
(b1l \<union>..\<union> bnl \<union> f1 a1l \<union>..\<union> fn anl, argl) \<approx>gen |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
128 |
alpha_ty fv_ty (p1 +..+ pl) |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
129 |
(b1r \<union>..\<union> bnr \<union> f1 a1r \<union>..\<union> fn anr, argr) |
e3a82a3529ce
Continued description of alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1514
diff
changeset
|
130 |
|
1514
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
131 |
- Argument has some recursive bindings. The bindings were |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
132 |
already treated in 2nd case so we return: True |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
133 |
- Argument has no bindings and is not bound. |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
134 |
If it is recursive for type ty, we return: alpha_ty argl argr |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
135 |
Otherwise we return: argl = argr |
b52b72676e20
First part of the description of alpha_ty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1513
diff
changeset
|
136 |
|
1510
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1508
diff
changeset
|
137 |
*) |
be911e869fde
Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1508
diff
changeset
|
138 |
|
1362
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
139 |
ML {* |
1669 | 140 |
datatype alpha_type = AlphaGen | AlphaRes | AlphaLst; |
141 |
*} |
|
142 |
||
143 |
ML {* |
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
144 |
fun atyp_const AlphaGen = @{const_name alpha_gen} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
145 |
| atyp_const AlphaRes = @{const_name alpha_res} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
146 |
| atyp_const AlphaLst = @{const_name alpha_lst} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
147 |
*} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
148 |
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
149 |
(* TODO: make sure that parser checks that bindings are compatible *) |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
150 |
ML {* |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
151 |
fun alpha_const_for_binds [] = atyp_const AlphaGen |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
152 |
| alpha_const_for_binds ((NONE, _, _, at) :: t) = atyp_const at |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
153 |
| alpha_const_for_binds ((SOME (_, _), _, _, at) :: _) = atyp_const at |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
154 |
*} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
155 |
|
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
156 |
ML {* |
1362
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
157 |
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
|
158 |
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
|
159 |
|
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
160 |
fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t |
1669 | 161 |
| is_atom_set _ _ = false; |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
162 |
|
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
163 |
fun is_atom_fset thy (Type ("FSet.fset", [t])) = is_atom thy t |
1669 | 164 |
| is_atom_fset _ _ = false; |
1362
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
165 |
*} |
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
166 |
|
e72d9d9eada3
Term5 written as nominal_datatype is the recursive let.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1359
diff
changeset
|
167 |
|
1358 | 168 |
(* 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
|
169 |
ML {* |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
170 |
fun map2i _ [] [] = [] |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
171 |
| 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
|
172 |
| 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
|
173 |
| map2i _ _ _ = raise UnequalLengths; |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
174 |
*} |
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
175 |
|
1358 | 176 |
(* 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
|
177 |
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
|
178 |
*) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
179 |
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
|
180 |
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
|
181 |
let |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
182 |
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
|
183 |
let |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
184 |
val common = map (fn (f, bi, _, aty) => (f, bi, aty)) binds |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
185 |
val nodups = distinct (op =) common |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
186 |
fun find_bodys (sf, sbi, sty) = |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
187 |
filter (fn (f, bi, _, aty) => f = sf andalso bi = sbi andalso aty = sty) binds |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
188 |
val bodys = map ((map (fn (_, _, bo, _) => bo)) o find_bodys) nodups |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
189 |
in |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
190 |
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
|
191 |
end |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
192 |
in |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
193 |
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
|
194 |
end |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
195 |
*} |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
196 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
197 |
ML {* |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
198 |
fun un_gather_binds_cons binds = |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
199 |
flat (map (fn (((f, bi, aty), bos), pi) => map (fn bo => ((f, bi, bo, aty), pi)) bos) binds) |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
200 |
*} |
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
201 |
|
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
202 |
ML {* |
1175 | 203 |
open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *); |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
204 |
*} |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
205 |
ML {* |
1178 | 206 |
(* TODO: It is the same as one in 'nominal_atoms' *) |
1175 | 207 |
fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom}); |
208 |
val noatoms = @{term "{} :: atom set"}; |
|
209 |
fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x]; |
|
210 |
fun mk_union sets = |
|
211 |
fold (fn a => fn b => |
|
212 |
if a = noatoms then b else |
|
213 |
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
|
214 |
if a = b then a else |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
215 |
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
|
216 |
val mk_inter = foldr1 (HOLogic.mk_binop @{const_name inf}) |
1175 | 217 |
fun mk_diff a b = |
218 |
if b = noatoms then a else |
|
219 |
if b = a then noatoms else |
|
220 |
HOLogic.mk_binop @{const_name minus} (a, b); |
|
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
221 |
fun mk_atom_set t = |
1185
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
222 |
let |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
223 |
val ty = fastype_of t; |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
224 |
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
|
225 |
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
|
226 |
in |
7566b899ca6a
Code for handling atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1180
diff
changeset
|
227 |
(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
|
228 |
end; |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
229 |
fun mk_atom_fset t = |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
230 |
let |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
231 |
val ty = fastype_of t; |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
232 |
val atom_ty = dest_fsetT ty --> @{typ atom}; |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
233 |
val fmap_ty = atom_ty --> ty --> @{typ "atom fset"}; |
1656
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
234 |
val fset_to_set = @{term "fset_to_set :: atom fset \<Rightarrow> atom set"} |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
235 |
in |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
236 |
fset_to_set $ ((Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)) |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
237 |
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
|
238 |
(* 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
|
239 |
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
|
240 |
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
|
241 |
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
|
242 |
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
|
243 |
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
|
244 |
end; |
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
245 |
*} |
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
|
246 |
|
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
247 |
(* Given [fv1, fv2, fv3] creates %(x, y, z). fv1 x u fv2 y u fv3 z *) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
248 |
ML {* |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
249 |
fun mk_compound_fv fvs = |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
250 |
let |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
251 |
val nos = (length fvs - 1) downto 0; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
252 |
val fvs_applied = map (fn (fv, no) => fv $ Bound no) (fvs ~~ nos); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
253 |
val fvs_union = mk_union fvs_applied; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
254 |
val (tyh :: tys) = rev (map (domain_type o fastype_of) fvs); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
255 |
fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t)) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
256 |
in |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
257 |
fold fold_fun tys (Abs ("", tyh, fvs_union)) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
258 |
end; |
1175 | 259 |
*} |
260 |
||
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
261 |
(* Given [R1, R2, R3] creates %(x,x'). %(y,y'). %(z,z'). R x x' \<and> R y y' \<and> R z z' *) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
262 |
ML {* |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
263 |
fun mk_compound_alpha Rs = |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
264 |
let |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
265 |
val nos = (length Rs - 1) downto 0; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
266 |
val nos2 = (2 * length Rs - 1) downto length Rs; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
267 |
val Rs_applied = map (fn (R, (no2, no)) => R $ Bound no2 $ Bound no) (Rs ~~ (nos2 ~~ nos)); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
268 |
val Rs_conj = mk_conjl Rs_applied; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
269 |
val (tyh :: tys) = rev (map (domain_type o fastype_of) Rs); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
270 |
fun fold_fun ty t = HOLogic.mk_split (Abs ("", ty, t)) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
271 |
val abs_rhs = fold fold_fun tys (Abs ("", tyh, Rs_conj)) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
272 |
in |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
273 |
fold fold_fun tys (Abs ("", tyh, abs_rhs)) |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
274 |
end; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
275 |
*} |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
276 |
|
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
277 |
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
|
278 |
|
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
|
279 |
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
|
280 |
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
|
281 |
let |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
282 |
fun is_non_rec (SOME (f, false), _, _, _) = SOME f |
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 |
| 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
|
284 |
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
|
285 |
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
|
286 |
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
|
287 |
*} |
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 |
|
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 |
(* We assume no bindings in the type on which bn is defined *) |
1505
e12ebfa745f6
Update the description of the generation of fv function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1498
diff
changeset
|
290 |
(* TODO: currently works only with current fv_bn function *) |
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
|
291 |
ML {* |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
292 |
fun fv_bn thy (dt_info : Datatype_Aux.info) fv_frees bn_fvbn (fvbn, (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
|
293 |
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
|
294 |
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
|
295 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
1379
5bb0149329ee
Separate lists for separate constructors, to match bn_eqs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1377
diff
changeset
|
296 |
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
|
297 |
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
|
298 |
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
|
299 |
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
|
300 |
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
|
301 |
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
|
302 |
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
|
303 |
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
|
304 |
val ty = fastype_of x |
1622
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
305 |
(* val _ = tracing ("B 1" ^ PolyML.makestring args_in_bn);*) |
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
306 |
(* val _ = tracing ("B 2" ^ PolyML.makestring bn_fvbn);*) |
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
|
307 |
in |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
308 |
case AList.lookup (op=) args_in_bn arg_no of |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
309 |
SOME NONE => @{term "{} :: atom set"} |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
310 |
| SOME (SOME (f : term)) => (the (AList.lookup (op=) bn_fvbn f)) $ x |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
311 |
| NONE => |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
312 |
if is_atom thy ty then mk_single_atom x else |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
313 |
if is_atom_set thy ty then mk_atom_set x else |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
314 |
if is_atom_fset thy ty then mk_atom_fset x else |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
315 |
if is_rec_type dt then nth fv_frees (body_index dt) $ x else |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
316 |
@{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
|
317 |
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
|
318 |
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
|
319 |
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
|
320 |
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
|
321 |
(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
|
322 |
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
|
323 |
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
|
324 |
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
|
325 |
in |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
326 |
((bn, fvbn), eqs) |
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
|
327 |
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
|
328 |
*} |
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1366
diff
changeset
|
329 |
|
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
330 |
ML {* print_depth 100 *} |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
331 |
ML {* |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
332 |
fun fv_bns thy dt_info fv_frees rel_bns = |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
333 |
let |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
334 |
fun mk_fvbn_free (bn, ith, _) = |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
335 |
let |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
336 |
val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
337 |
in |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
338 |
(fvbn_name, Free (fvbn_name, fastype_of (nth fv_frees ith))) |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
339 |
end; |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
340 |
val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free rel_bns); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
341 |
val bn_fvbn = (map (fn (bn, _, _) => bn) rel_bns) ~~ fvbn_frees |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
342 |
val (l1, l2) = split_list (map (fv_bn thy dt_info fv_frees bn_fvbn) (fvbn_frees ~~ rel_bns)); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
343 |
in |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
344 |
(l1, (fvbn_names ~~ l2)) |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
345 |
end |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
346 |
*} |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
347 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
348 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
349 |
ML {* |
1669 | 350 |
fun alpha_bn (dt_info : Datatype_Aux.info) alpha_frees bn_alphabn ((bn, ith_dtyp, args_in_bns), (alpha_bn_free, _ (*is_rec*) )) = |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
351 |
let |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
352 |
val {descr, sorts, ...} = dt_info; |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
353 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
354 |
fun alpha_bn_constr (cname, dts) args_in_bn = |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
355 |
let |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
356 |
val Ts = map (typ_of_dtyp descr sorts) dts; |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
357 |
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
|
358 |
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
|
359 |
val args = map Free (names ~~ Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
360 |
val args2 = map Free (names2 ~~ Ts); |
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
361 |
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
|
362 |
val rhs = HOLogic.mk_Trueprop |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
363 |
(alpha_bn_free $ (list_comb (c, args)) $ (list_comb (c, args2))); |
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
364 |
fun lhs_arg ((dt, arg_no), (arg, arg2)) = |
1669 | 365 |
case AList.lookup (op=) args_in_bn arg_no of |
366 |
SOME NONE => @{term True} |
|
367 |
| SOME (SOME f) => (the (AList.lookup (op=) bn_alphabn f)) $ arg $ arg2 |
|
368 |
| NONE => |
|
369 |
if is_rec_type dt then (nth alpha_frees (body_index dt)) $ arg $ arg2 |
|
370 |
else HOLogic.mk_eq (arg, arg2) |
|
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
371 |
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
|
372 |
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
|
373 |
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
|
374 |
in |
1387
9d70a0733786
rhs of alpha_bn, and template for the arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1386
diff
changeset
|
375 |
eq |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
376 |
end |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
377 |
val (_, (_, _, constrs)) = nth descr ith_dtyp; |
1386
0511e879a687
alpha_bn_constr template
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1385
diff
changeset
|
378 |
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
|
379 |
in |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
380 |
((bn, alpha_bn_free), eqs) |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
381 |
end |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
382 |
*} |
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
383 |
|
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
384 |
ML {* |
1669 | 385 |
fun alpha_bns dt_info alpha_frees rel_bns bns_rec = |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
386 |
let |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
387 |
val {descr, sorts, ...} = dt_info; |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
388 |
fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
389 |
fun mk_alphabn_free (bn, ith, _) = |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
390 |
let |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
391 |
val alphabn_name = "alpha_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
392 |
val alphabn_type = nth_dtyp ith --> nth_dtyp ith --> @{typ bool}; |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
393 |
val alphabn_free = Free(alphabn_name, alphabn_type); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
394 |
in |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
395 |
(alphabn_name, alphabn_free) |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
396 |
end; |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
397 |
val (alphabn_names, alphabn_frees) = split_list (map mk_alphabn_free rel_bns); |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
398 |
val bn_alphabn = (map (fn (bn, _, _) => bn) rel_bns) ~~ alphabn_frees; |
1669 | 399 |
val pair = split_list (map (alpha_bn dt_info alpha_frees bn_alphabn) |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
400 |
(rel_bns ~~ (alphabn_frees ~~ bns_rec))) |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
401 |
in |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
402 |
(alphabn_names, pair) |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
403 |
end |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
404 |
*} |
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
405 |
|
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
406 |
|
1397
6e2dfe52b271
Allows multiple bindings with same lhs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1394
diff
changeset
|
407 |
(* Checks that a list of bindings contains only compatible ones *) |
6e2dfe52b271
Allows multiple bindings with same lhs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1394
diff
changeset
|
408 |
ML {* |
6e2dfe52b271
Allows multiple bindings with same lhs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1394
diff
changeset
|
409 |
fun bns_same l = |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
410 |
length (distinct (op =) (map (fn ((b, _, _, atyp), _) => (b, atyp)) l)) = 1 |
1397
6e2dfe52b271
Allows multiple bindings with same lhs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1394
diff
changeset
|
411 |
*} |
6e2dfe52b271
Allows multiple bindings with same lhs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1394
diff
changeset
|
412 |
|
1679 | 413 |
ML {* |
414 |
fun setify x = |
|
415 |
if fastype_of x = @{typ "atom list"} then |
|
416 |
Const (@{const_name set}, @{typ "atom list \<Rightarrow> atom set"}) $ x else x |
|
417 |
*} |
|
418 |
||
1206 | 419 |
(* TODO: Notice datatypes without bindings and replace alpha with equality *) |
1175 | 420 |
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
|
421 |
fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall bns lthy = |
1178 | 422 |
let |
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
423 |
val thy = ProofContext.theory_of lthy; |
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1258
diff
changeset
|
424 |
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
|
425 |
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
|
426 |
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
|
427 |
"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
|
428 |
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
|
429 |
val fv_frees = map Free (fv_names ~~ fv_types); |
1622
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
430 |
(* TODO: We need a transitive closure, but instead we do this hack considering |
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
431 |
all binding functions as recursive or not *) |
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
432 |
val nr_bns = |
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
433 |
if (non_rec_binds bindsall) = [] then [] |
006d81399f6a
Compute Fv for non-recursive bn functions calling other bn functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1615
diff
changeset
|
434 |
else map (fn (bn, _, _) => bn) bns; |
1464
1850361efb8f
Revert 7c8cd6eae8e2, now all proofs in Term5 go through, both recursive and not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1462
diff
changeset
|
435 |
val rel_bns = filter (fn (bn, _, _) => bn mem nr_bns) bns; |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
436 |
val (bn_fv_bns, fv_bn_names_eqs) = fv_bns thy dt_info fv_frees rel_bns; |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
437 |
val fvbns = map snd bn_fv_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
|
438 |
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
|
439 |
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
|
440 |
"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
|
441 |
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
|
442 |
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
|
443 |
(* 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
|
444 |
val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns; |
1615
0ea578c6dae3
Parsing bn functions that call other bn functions and transmitting this information to fv/alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1609
diff
changeset
|
445 |
val (alpha_bn_names, (bn_alpha_bns, alpha_bn_eqs)) = |
1669 | 446 |
alpha_bns dt_info alpha_frees bns bns_rec |
1389
d0ba4829a76c
include alpha in the definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1388
diff
changeset
|
447 |
val alpha_bn_frees = map snd bn_alpha_bns; |
d0ba4829a76c
include alpha in the definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1388
diff
changeset
|
448 |
val alpha_bn_types = map fastype_of alpha_bn_frees; |
1679 | 449 |
|
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
450 |
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
|
451 |
let |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
452 |
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
|
453 |
val bindslen = length bindcs |
1323
acf262971303
Fixes for the fv problem and alpha problem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1308
diff
changeset
|
454 |
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
|
455 |
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
|
456 |
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
|
457 |
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
|
458 |
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
|
459 |
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
|
460 |
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
|
461 |
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
|
462 |
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
|
463 |
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
|
464 |
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
|
465 |
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
|
466 |
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
|
467 |
val arg_nos = 0 upto (length dts - 1) |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
468 |
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
|
469 |
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
|
470 |
if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
471 |
if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
472 |
if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else |
1366
2bf82fa871e7
Proper recognition of atoms and atom sets.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1362
diff
changeset
|
473 |
(* 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
|
474 |
@{term "{} :: atom set"} |
1679 | 475 |
| 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
|
476 |
fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) |
1679 | 477 |
fun fv_binds_as_set args relevant = mk_union (map (setify o fv_bind args) relevant) |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
478 |
fun find_nonrec_binder j (SOME (f, false), i, _, _) = if i = j then SOME f else NONE |
1464
1850361efb8f
Revert 7c8cd6eae8e2, now all proofs in Term5 go through, both recursive and not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1462
diff
changeset
|
479 |
| 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
|
480 |
fun fv_arg ((dt, x), arg_no) = |
1464
1850361efb8f
Revert 7c8cd6eae8e2, now all proofs in Term5 go through, both recursive and not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1462
diff
changeset
|
481 |
case get_first (find_nonrec_binder arg_no) bindcs of |
1850361efb8f
Revert 7c8cd6eae8e2, now all proofs in Term5 go through, both recursive and not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1462
diff
changeset
|
482 |
SOME f => |
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
|
483 |
(case get_first (fn (x, y) => if x = f then SOME y else NONE) bn_fv_bns of |
1464
1850361efb8f
Revert 7c8cd6eae8e2, now all proofs in Term5 go through, both recursive and not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1462
diff
changeset
|
484 |
SOME fv_bn => fv_bn $ 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
|
485 |
| 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
|
486 |
| 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
|
487 |
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
|
488 |
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
|
489 |
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
|
490 |
if ((is_atom thy) o fastype_of) x then mk_single_atom x else |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
491 |
if ((is_atom_set thy) o fastype_of) x then mk_atom_set x else |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
492 |
if ((is_atom_fset thy) o fastype_of) x then mk_atom_fset x else |
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
|
493 |
(* TODO we do not know what to do with non-atomizable things *) |
1454
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1452
diff
changeset
|
494 |
@{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
|
495 |
(* If i = j then we generate it only once *) |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
496 |
val relevant = filter (fn (_, i, j, _) => ((i = arg_no) orelse (j = arg_no))) bindcs; |
1679 | 497 |
val sub = fv_binds_as_set 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
|
498 |
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
|
499 |
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
|
500 |
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
|
501 |
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
|
502 |
(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
|
503 |
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
|
504 |
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
|
505 |
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
|
506 |
let |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
507 |
val rel_in_simp_binds = filter (fn ((NONE, i, _, _), _) => i = arg_no | _ => false) bind_pis; |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
508 |
val rel_in_comp_binds = filter (fn ((SOME _, i, _, _), _) => i = arg_no | _ => false) bind_pis; |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
509 |
val rel_has_binds = filter (fn ((NONE, _, j, _), _) => j = arg_no |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
510 |
| ((SOME (_, false), _, j, _), _) => j = arg_no |
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
511 |
| _ => false) bind_pis; |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
512 |
val rel_has_rec_binds = filter |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
513 |
(fn ((SOME (_, true), _, 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
|
514 |
in |
1472
4fa5365cd871
Fix in alpha; support of the recursive Let works :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1468
diff
changeset
|
515 |
case (rel_in_simp_binds, rel_in_comp_binds, rel_has_binds, rel_has_rec_binds) of |
4fa5365cd871
Fix in alpha; support of the recursive Let works :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1468
diff
changeset
|
516 |
([], [], [], []) => |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
517 |
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
|
518 |
else (HOLogic.mk_eq (arg, arg2)) |
1472
4fa5365cd871
Fix in alpha; support of the recursive Let works :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1468
diff
changeset
|
519 |
| (_, [], [], []) => @{term True} |
4fa5365cd871
Fix in alpha; support of the recursive Let works :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1468
diff
changeset
|
520 |
| ([], [], [], _) => @{term True} |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
521 |
| ([], ((((SOME (bn, is_rec)), _, _, atyp), _) :: _), [], []) => |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
522 |
if not (bns_same rel_in_comp_binds) then error "incompatible bindings for an argument" else |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
523 |
if is_rec then |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
524 |
let |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
525 |
val (rbinds, rpis) = split_list rel_in_comp_binds |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
526 |
val bound_in_nos = map (fn (_, _, i, _) => i) rbinds |
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
527 |
val bound_in_ty_nos = map (fn i => body_index (nth dts i)) bound_in_nos; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
528 |
val bound_args = arg :: map (nth args) bound_in_nos; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
529 |
val bound_args2 = arg2 :: map (nth args2) bound_in_nos; |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
530 |
val lhs_binds = fv_binds args rbinds |
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
531 |
val lhs_arg = foldr1 HOLogic.mk_prod bound_args |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
532 |
val lhs = mk_pair (lhs_binds, lhs_arg); |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
533 |
val rhs_binds = fv_binds args2 rbinds; |
1468
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
534 |
val rhs_arg = foldr1 HOLogic.mk_prod bound_args2; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
535 |
val rhs = mk_pair (rhs_binds, rhs_arg); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
536 |
val fvs = map (nth fv_frees) ((body_index dt) :: bound_in_ty_nos); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
537 |
val fv = mk_compound_fv fvs; |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
538 |
val alphas = map (nth alpha_frees) ((body_index dt) :: bound_in_ty_nos); |
416c9c5a1126
Generate compound FV and Alpha for recursive bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1464
diff
changeset
|
539 |
val alpha = mk_compound_alpha alphas; |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
540 |
val pi = foldr1 add_perm (distinct (op =) rpis); |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
541 |
val alpha_gen_pre = Const (atyp_const atyp, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
1462
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
542 |
val alpha_gen = Syntax.check_term lthy alpha_gen_pre |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
543 |
in |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
544 |
alpha_gen |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
545 |
end |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
546 |
else |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
547 |
let |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
548 |
val alpha_bn_const = |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
549 |
nth alpha_bn_frees (find_index (fn (b, _, _) => b = bn) bns) |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
550 |
in |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
551 |
alpha_bn_const $ arg $ arg2 |
7c59dd8e5435
The old recursive alpha works fine.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1457
diff
changeset
|
552 |
end |
1472
4fa5365cd871
Fix in alpha; support of the recursive Let works :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1468
diff
changeset
|
553 |
| ([], [], relevant, []) => |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
554 |
let |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
555 |
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
|
556 |
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
|
557 |
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
|
558 |
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
|
559 |
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
|
560 |
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
|
561 |
val fv = nth fv_frees (body_index dt); |
1359
3bf496a971c6
Fix permutation addition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1358
diff
changeset
|
562 |
val pi = foldr1 add_perm (distinct (op =) rpis); |
1670
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
563 |
val alpha_const = alpha_const_for_binds rbinds; |
ed89a26b7074
Fv/Alpha now takes into account Alpha_Type given from the parser.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1669
diff
changeset
|
564 |
val alpha_gen_pre = Const (alpha_const, dummyT) $ lhs $ alpha $ fv $ pi $ rhs; |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
565 |
val alpha_gen = Syntax.check_term lthy alpha_gen_pre |
1288
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
566 |
in |
1357
42b7abf779ec
Gather bindings with same binder, and generate only one permutation for them.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1339
diff
changeset
|
567 |
alpha_gen |
1383
8399236f901b
Use alpha_bns in normal alpha defs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1382
diff
changeset
|
568 |
end |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
569 |
| _ => 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
|
570 |
end |
0203cd5cfd6c
The new alpha-equivalence and testing in Trm2 and Trm5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
571 |
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
|
572 |
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
|
573 |
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
|
574 |
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
|
575 |
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
|
576 |
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
|
577 |
(fv_eq, alpha_eq) |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
578 |
end; |
1302
d3101a5b9c4d
Length fix for nested recursions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1301
diff
changeset
|
579 |
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
|
580 |
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
|
581 |
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
|
582 |
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
|
583 |
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
|
584 |
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
|
585 |
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
|
586 |
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
|
587 |
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
|
588 |
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
|
589 |
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
|
590 |
(* Function_Fun.add_fun Function_Common.default_config ... true *) |
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
|
591 |
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
|
592 |
(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
|
593 |
val (fvs2, lthy'') = |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
594 |
if fv_eqs_snd = [] then (([], []), lthy') else |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
595 |
(Primrec.add_primrec |
56efa1e854bf
Separate primrecs in Fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1375
diff
changeset
|
596 |
(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
|
597 |
val (alphas, lthy''') = (Inductive.add_inductive_i |
1325
0be098c61d00
Add the supp intersection conditions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1323
diff
changeset
|
598 |
{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
|
599 |
coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false} |
1389
d0ba4829a76c
include alpha in the definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1388
diff
changeset
|
600 |
(map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) (alpha_names @ alpha_bn_names) |
d0ba4829a76c
include alpha in the definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1388
diff
changeset
|
601 |
(alpha_types @ alpha_bn_types)) [] |
d0ba4829a76c
include alpha in the definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1388
diff
changeset
|
602 |
(add_binds (alpha_eqs @ flat alpha_bn_eqs)) [] lthy'') |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
603 |
val ordered_fvs = fv_frees @ fvbns; |
1385
51b8e6dd72d5
exported template for alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1383
diff
changeset
|
604 |
val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) |
1178 | 605 |
in |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
606 |
(((all_fvs, ordered_fvs), alphas), lthy''') |
1178 | 607 |
end |
1168
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
608 |
*} |
5c1e16806901
Code for generating the fv function, no bindings yet.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
609 |
|
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
|
610 |
|
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
611 |
|
1207 | 612 |
ML {* |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
613 |
fun build_alpha_sym_trans_gl alphas (x, y, z) = |
1207 | 614 |
let |
615 |
fun build_alpha alpha = |
|
616 |
let |
|
617 |
val ty = domain_type (fastype_of alpha); |
|
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
618 |
val var = Free(x, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
619 |
val var2 = Free(y, ty); |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
620 |
val var3 = Free(z, ty); |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
621 |
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
|
622 |
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
|
623 |
HOLogic.mk_all (z, ty, |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
624 |
HOLogic.mk_imp (alpha $ var2 $ var3, alpha $ var $ var3))) |
1207 | 625 |
in |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
626 |
(symp, transp) |
1208
cc86faf0d2a0
alpha-symmetric addons.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1207
diff
changeset
|
627 |
end; |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
628 |
val eqs = map build_alpha alphas |
1209
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
629 |
val (sym_eqs, trans_eqs) = split_list eqs |
7b1a3df239cd
Some progress about transp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1208
diff
changeset
|
630 |
fun conj l = @{term Trueprop} $ foldr1 HOLogic.mk_conj l |
1207 | 631 |
in |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
632 |
(conj sym_eqs, conj trans_eqs) |
1196
4efbaba9d754
Constructing alpha_inj goal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1193
diff
changeset
|
633 |
end |
1207 | 634 |
*} |
635 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
636 |
ML {* |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
637 |
fun build_alpha_refl_gl fv_alphas_lst alphas = |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
638 |
let |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
639 |
val (fvs_alphas, _) = split_list fv_alphas_lst; |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
640 |
val (_, alpha_ts) = split_list fvs_alphas; |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
641 |
val tys = map (domain_type o fastype_of) alpha_ts; |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
642 |
val names = Datatype_Prop.make_tnames tys; |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
643 |
val args = map Free (names ~~ tys); |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
644 |
fun find_alphas ty x = |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
645 |
domain_type (fastype_of x) = ty; |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
646 |
fun refl_eq_arg (ty, arg) = |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
647 |
let |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
648 |
val rel_alphas = filter (find_alphas ty) alphas; |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
649 |
in |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
650 |
map (fn x => x $ arg $ arg) rel_alphas |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
651 |
end; |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
652 |
(* Flattening loses the induction structure *) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
653 |
val eqs = map (foldr1 HOLogic.mk_conj) (map refl_eq_arg (tys ~~ args)) |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
654 |
in |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
655 |
(names, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj eqs)) |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
656 |
end |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
657 |
*} |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
658 |
|
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
659 |
ML {* |
1669 | 660 |
fun reflp_tac induct eq_iff = |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
661 |
rtac induct THEN_ALL_NEW |
1656
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
662 |
simp_tac (HOL_basic_ss addsimps eq_iff) THEN_ALL_NEW |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
663 |
split_conj_tac THEN_ALL_NEW REPEAT o rtac @{thm exI[of _ "0 :: perm"]} |
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
664 |
THEN_ALL_NEW split_conj_tac THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps |
1672
94b8b70f7bc0
Initial proof modifications for alpha_res
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1670
diff
changeset
|
665 |
@{thms alphas fresh_star_def fresh_zero_perm permute_zero ball_triv |
1482
a98c15866300
Trying to find a compose lemma for 2 arguments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1480
diff
changeset
|
666 |
add_0_left supp_zero_perm Int_empty_left split_conv}) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
667 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
668 |
|
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
669 |
ML {* |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
670 |
fun build_alpha_refl fv_alphas_lst alphas induct eq_iff ctxt = |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
671 |
let |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
672 |
val (names, gl) = build_alpha_refl_gl fv_alphas_lst alphas; |
1669 | 673 |
val refl_conj = Goal.prove ctxt names [] gl (fn _ => reflp_tac induct eq_iff 1); |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
674 |
in |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
675 |
HOLogic.conj_elims refl_conj |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
676 |
end |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
677 |
*} |
1333
c6e521a2a0b1
reflp for multiple quantifiers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1325
diff
changeset
|
678 |
|
1301 | 679 |
lemma exi_neg: "\<exists>(pi :: perm). P pi \<Longrightarrow> (\<And>(p :: perm). P p \<Longrightarrow> Q (- p)) \<Longrightarrow> \<exists>pi. Q pi" |
680 |
apply (erule exE) |
|
681 |
apply (rule_tac x="-pi" in exI) |
|
682 |
by auto |
|
683 |
||
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
684 |
ML {* |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
685 |
fun symp_tac induct inj eqvt ctxt = |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
686 |
rel_indtac induct THEN_ALL_NEW |
1656
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
687 |
simp_tac (HOL_basic_ss addsimps inj) THEN_ALL_NEW split_conj_tac |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
688 |
THEN_ALL_NEW |
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
689 |
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
|
690 |
THEN_ALL_NEW |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
691 |
split_conj_tac THEN_ALL_NEW |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
692 |
asm_full_simp_tac (HOL_ss addsimps @{thms supp_minus_perm minus_add[symmetric]}) THEN_ALL_NEW |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1672
diff
changeset
|
693 |
TRY o (resolve_tac @{thms alphas_compose_sym2} ORELSE' resolve_tac @{thms alphas_compose_sym}) THEN_ALL_NEW |
1487 | 694 |
(asm_full_simp_tac (HOL_ss addsimps (eqvt @ all_eqvts ctxt))) |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
695 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
696 |
|
1301 | 697 |
|
698 |
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" |
|
699 |
apply (erule exE)+ |
|
700 |
apply (rule_tac x="pia + pi" in exI) |
|
701 |
by auto |
|
702 |
||
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
|
703 |
|
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
|
704 |
ML {* |
1656
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
705 |
fun eetac rule = |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
706 |
Subgoal.FOCUS_PARAMS (fn focus => |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
707 |
let |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
708 |
val concl = #concl focus |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
709 |
val prems = Logic.strip_imp_prems (term_of concl) |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
710 |
val exs = filter (fn x => is_ex (HOLogic.dest_Trueprop x)) prems |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
711 |
val cexs = map (SOME o (cterm_of (ProofContext.theory_of (#context focus)))) exs |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
712 |
val thins = map (fn cex => Drule.instantiate' [] [cex] Drule.thin_rl) cexs |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
713 |
in |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
714 |
(etac rule THEN' RANGE[atac, eresolve_tac thins]) 1 |
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
715 |
end |
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
|
716 |
) |
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
|
717 |
*} |
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
|
718 |
|
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
|
719 |
ML {* |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
720 |
fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt = |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
721 |
rel_indtac induct THEN_ALL_NEW |
1217
74e2b9b95add
Fixes for auxiliary datatypes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1216
diff
changeset
|
722 |
(TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW |
1656
c9d3dda79fe3
Removed remaining cheats + some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1653
diff
changeset
|
723 |
asm_full_simp_tac (HOL_basic_ss addsimps alpha_inj) THEN_ALL_NEW |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
724 |
split_conj_tac THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) THEN_ALL_NEW split_conj_tac |
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
|
725 |
THEN_ALL_NEW (asm_full_simp_tac (HOL_ss addsimps (term_inj @ distinct))) |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
726 |
THEN_ALL_NEW split_conj_tac THEN_ALL_NEW |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1672
diff
changeset
|
727 |
TRY o (eresolve_tac @{thms alphas_compose_trans2} ORELSE' eresolve_tac @{thms alphas_compose_trans}) 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
|
728 |
(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
|
729 |
*} |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
730 |
|
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
731 |
lemma transpI: |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
732 |
"(\<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
|
733 |
unfolding transp_def |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
734 |
by blast |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
735 |
|
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
736 |
ML {* |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
737 |
fun equivp_tac reflps symps transps = |
1609
c9bc3b61046c
Modification to Core Haskell to make it accepted with an empty binding function.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1581
diff
changeset
|
738 |
(*let val _ = tracing (PolyML.makestring (reflps, symps, transps)) in *) |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
739 |
simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def}) |
1221 | 740 |
THEN' rtac conjI THEN' rtac allI THEN' |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
741 |
resolve_tac reflps THEN' |
1221 | 742 |
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
|
743 |
resolve_tac symps THEN' |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1560
diff
changeset
|
744 |
rtac @{thm transpI} THEN' resolve_tac transps |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
745 |
*} |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
746 |
|
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
747 |
ML {* |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
748 |
fun build_equivps alphas reflps 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
|
749 |
let |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
750 |
val ([x, y, z], ctxt') = Variable.variant_fixes ["x","y","z"] ctxt; |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
751 |
val (symg, transg) = build_alpha_sym_trans_gl alphas (x, y, z) |
1334
80441e27dfd6
Code for solving symp goals with multiple existentials.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1333
diff
changeset
|
752 |
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
|
753 |
fun transp_tac' _ = transp_tac ctxt alpha_induct alpha_inj term_inj distinct cases eqvt 1; |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
754 |
val symp_loc = Goal.prove ctxt' [] [] symg symp_tac'; |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
755 |
val transp_loc = Goal.prove ctxt' [] [] transg transp_tac'; |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
756 |
val [symp, transp] = Variable.export ctxt' ctxt [symp_loc, transp_loc] |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
757 |
val symps = HOLogic.conj_elims symp |
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
758 |
val transps = HOLogic.conj_elims transp |
1214
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
759 |
fun equivp alpha = |
0f92257edeee
A tactic for final equivp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1213
diff
changeset
|
760 |
let |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
761 |
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
|
762 |
val goal = @{term Trueprop} $ (equivp $ alpha) |
1559
d375804ce6ba
Prove reflp for all relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1553
diff
changeset
|
763 |
fun tac _ = equivp_tac reflps symps transps 1 |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
764 |
in |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
765 |
Goal.prove ctxt [] [] goal tac |
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
766 |
end |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
767 |
in |
1215
aec9576b3950
Testing auto equivp code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1214
diff
changeset
|
768 |
map equivp alphas |
1213
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
769 |
end |
43bd70786f9f
More equivp infrastructure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1209
diff
changeset
|
770 |
*} |
1207 | 771 |
|
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
772 |
lemma not_in_union: "c \<notin> a \<union> b \<equiv> (c \<notin> a \<and> c \<notin> b)" |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
773 |
by auto |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
774 |
|
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
775 |
ML {* |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
776 |
fun supports_tac perm = |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
777 |
simp_tac (HOL_ss addsimps @{thms supports_def not_in_union} @ perm) THEN_ALL_NEW ( |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
778 |
REPEAT o rtac allI THEN' REPEAT o rtac impI THEN' split_conj_tac THEN' |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
779 |
asm_full_simp_tac (HOL_ss addsimps @{thms fresh_def[symmetric] |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
780 |
swap_fresh_fresh fresh_atom swap_at_base_simps(3) swap_atom_image_fresh |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
781 |
supp_fset_to_set supp_fmap_atom})) |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
782 |
*} |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
783 |
|
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
784 |
ML {* |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
785 |
fun mk_supp ty x = |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
786 |
Const (@{const_name supp}, ty --> @{typ "atom set"}) $ x |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
787 |
*} |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
788 |
|
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
789 |
ML {* |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
790 |
fun mk_supports_eq thy cnstr = |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
791 |
let |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
792 |
val (tys, ty) = (strip_type o fastype_of) cnstr |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
793 |
val names = Datatype_Prop.make_tnames tys |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
794 |
val frees = map Free (names ~~ tys) |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
795 |
val rhs = list_comb (cnstr, frees) |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
796 |
|
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
797 |
fun mk_supp_arg (x, ty) = |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
798 |
if is_atom thy ty then mk_supp @{typ atom} (mk_atom ty $ x) else |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
799 |
if is_atom_set thy ty then mk_supp @{typ "atom set"} (mk_atom_set x) else |
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
800 |
if is_atom_fset thy ty then mk_supp @{typ "atom set"} (mk_atom_fset x) |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
801 |
else mk_supp ty x |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
802 |
val lhss = map mk_supp_arg (frees ~~ tys) |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
803 |
val supports = Const(@{const_name "supports"}, @{typ "atom set"} --> ty --> @{typ bool}) |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
804 |
val eq = HOLogic.mk_Trueprop (supports $ mk_union lhss $ rhs) |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
805 |
in |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
806 |
(names, eq) |
1422
a26cb19375e8
Remove "_raw" from lifted theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1415
diff
changeset
|
807 |
end |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
808 |
*} |
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
809 |
|
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
810 |
ML {* |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
811 |
fun prove_supports ctxt perms cnst = |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
812 |
let |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
813 |
val (names, eq) = mk_supports_eq (ProofContext.theory_of ctxt) cnst |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
814 |
in |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
815 |
Goal.prove ctxt names [] eq (fn _ => supports_tac perms 1) |
1427
b355cab42841
mk_supports_eq and supports_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1422
diff
changeset
|
816 |
end |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
817 |
*} |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
818 |
|
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
819 |
ML {* |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
820 |
fun mk_fs tys = |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
821 |
let |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
822 |
val names = Datatype_Prop.make_tnames tys |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
823 |
val frees = map Free (names ~~ tys) |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
824 |
val supps = map2 mk_supp tys frees |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
825 |
val fin_supps = map (fn x => @{term "finite :: atom set \<Rightarrow> bool"} $ x) supps |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
826 |
in |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
827 |
(names, HOLogic.mk_Trueprop (mk_conjl fin_supps)) |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
828 |
end |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
829 |
*} |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
830 |
|
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
831 |
ML {* |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
832 |
fun fs_tac induct supports = rel_indtac induct THEN_ALL_NEW ( |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
833 |
rtac @{thm supports_finite} THEN' resolve_tac supports) THEN_ALL_NEW |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1516
diff
changeset
|
834 |
asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom supp_atom_image supp_fset_to_set |
1547
57f7af5d7564
Use fs typeclass in showing finite support + some cheat cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
835 |
supp_fmap_atom finite_insert finite.emptyI finite_Un finite_supp}) |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
836 |
*} |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
837 |
|
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
838 |
ML {* |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
839 |
fun prove_fs ctxt induct supports tys = |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
840 |
let |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
841 |
val (names, eq) = mk_fs tys |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
842 |
in |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
843 |
Goal.prove ctxt names [] eq (fn _ => fs_tac induct supports 1) |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
844 |
end |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
845 |
*} |
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
846 |
|
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
847 |
ML {* |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
848 |
fun mk_supp x = Const (@{const_name supp}, fastype_of x --> @{typ "atom set"}) $ x; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
849 |
|
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
850 |
fun mk_supp_neq arg (fv, alpha) = |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
851 |
let |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
852 |
val collect = Const ("Collect", @{typ "(atom \<Rightarrow> bool) \<Rightarrow> atom \<Rightarrow> bool"}); |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
853 |
val ty = fastype_of arg; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
854 |
val perm = Const ("Nominal2_Base.pt_class.permute", @{typ perm} --> ty --> ty); |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
855 |
val finite = @{term "finite :: atom set \<Rightarrow> bool"} |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
856 |
val rhs = collect $ Abs ("a", @{typ atom}, |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
857 |
HOLogic.mk_not (finite $ |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
858 |
(collect $ Abs ("b", @{typ atom}, |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
859 |
HOLogic.mk_not (alpha $ (perm $ (@{term swap} $ Bound 1 $ Bound 0) $ arg) $ arg))))) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
860 |
in |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
861 |
HOLogic.mk_eq (fv $ arg, rhs) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
862 |
end; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
863 |
|
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
864 |
fun supp_eq fv_alphas_lst = |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
865 |
let |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
866 |
val (fvs_alphas, ls) = split_list fv_alphas_lst; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
867 |
val (fv_ts, _) = split_list fvs_alphas; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
868 |
val tys = map (domain_type o fastype_of) fv_ts; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
869 |
val names = Datatype_Prop.make_tnames tys; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
870 |
val args = map Free (names ~~ tys); |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
871 |
fun supp_eq_arg ((fv, arg), l) = |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
872 |
mk_conjl |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
873 |
((HOLogic.mk_eq (fv $ arg, mk_supp arg)) :: |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
874 |
(map (mk_supp_neq arg) l)) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
875 |
val eqs = mk_conjl (map supp_eq_arg ((fv_ts ~~ args) ~~ ls)) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
876 |
in |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
877 |
(names, HOLogic.mk_Trueprop eqs) |
1428
4029105011ca
Show that the new types are in finite support typeclass.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1427
diff
changeset
|
878 |
end |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
879 |
*} |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
880 |
|
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
881 |
ML {* |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
882 |
fun combine_fv_alpha_bns (fv_ts_nobn, fv_ts_bn) (alpha_ts_nobn, alpha_ts_bn) bn_nos = |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
883 |
if length fv_ts_bn < length alpha_ts_bn then |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
884 |
(fv_ts_nobn ~~ alpha_ts_nobn) ~~ (replicate (length fv_ts_nobn) []) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
885 |
else let |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
886 |
val fv_alpha_nos = 0 upto (length fv_ts_nobn - 1); |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
887 |
fun filter_fn i (x, j) = if j = i then SOME x else NONE; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
888 |
val fv_alpha_bn_nos = (fv_ts_bn ~~ alpha_ts_bn) ~~ bn_nos; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
889 |
val fv_alpha_bn_all = map (fn i => map_filter (filter_fn i) fv_alpha_bn_nos) fv_alpha_nos; |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
890 |
in |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
891 |
(fv_ts_nobn ~~ alpha_ts_nobn) ~~ fv_alpha_bn_all |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
892 |
end |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
893 |
*} |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
894 |
|
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
895 |
(* TODO: this is a hack, it assumes that only one type of Abs's is present |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
896 |
in the type and chooses this supp_abs. Additionally single atoms are |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
897 |
treated properly. *) |
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
898 |
ML {* |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
899 |
fun choose_alpha_abs eqiff = |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
900 |
let |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
901 |
fun exists_subterms f ts = true mem (map (exists_subterm f) ts); |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
902 |
val terms = map prop_of eqiff; |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
903 |
fun check cname = exists_subterms (fn x => fst(dest_Const x) = cname handle _ => false) terms |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
904 |
val no = |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
905 |
if check @{const_name alpha_lst} then 2 else |
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
906 |
if check @{const_name alpha_res} then 1 else |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
907 |
if check @{const_name alpha_gen} then 0 else |
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
908 |
error "Failure choosing supp_abs" |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
909 |
in |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
910 |
nth @{thms supp_abs[symmetric]} no |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
911 |
end |
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
912 |
*} |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
913 |
lemma supp_abs_atom: "supp (Abs {atom a} (x :: 'a :: fs)) = supp x - {atom a}" |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
914 |
by (rule supp_abs(1)) |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
915 |
|
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
916 |
lemma supp_abs_sum: |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
917 |
"supp (Abs x (a :: 'a :: fs)) \<union> supp (Abs x (b :: 'b :: fs)) = supp (Abs x (a, b))" |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
918 |
"supp (Abs_res x (a :: 'a :: fs)) \<union> supp (Abs_res x (b :: 'b :: fs)) = supp (Abs_res x (a, b))" |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
919 |
"supp (Abs_lst y (a :: 'a :: fs)) \<union> supp (Abs_lst y (b :: 'b :: fs)) = supp (Abs_lst y (a, b))" |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
920 |
apply (simp_all add: supp_abs supp_Pair) |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
921 |
apply blast+ |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
922 |
done |
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
923 |
|
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
924 |
|
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
925 |
ML {* |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
926 |
fun supp_eq_tac ind fv perm eqiff ctxt = |
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
927 |
rel_indtac ind THEN_ALL_NEW |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
928 |
asm_full_simp_tac (HOL_basic_ss addsimps fv) THEN_ALL_NEW |
1685
721d92623c9d
Lets finally abstract lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1680
diff
changeset
|
929 |
asm_full_simp_tac (HOL_basic_ss addsimps @{thms supp_abs_atom[symmetric]}) THEN_ALL_NEW |
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
930 |
asm_full_simp_tac (HOL_basic_ss addsimps [choose_alpha_abs eqiff]) THEN_ALL_NEW |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
931 |
simp_tac (HOL_basic_ss addsimps @{thms supp_abs_sum}) THEN_ALL_NEW |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
932 |
simp_tac (HOL_basic_ss addsimps @{thms supp_def}) THEN_ALL_NEW |
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
933 |
simp_tac (HOL_basic_ss addsimps (@{thms permute_abs} @ perm)) THEN_ALL_NEW |
1675
d24f59f78a86
Generalize Abs_eq_iff.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
934 |
simp_tac (HOL_basic_ss addsimps (@{thms Abs_eq_iff} @ eqiff)) THEN_ALL_NEW |
1744
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1685
diff
changeset
|
935 |
simp_tac (HOL_basic_ss addsimps @{thms alphas3 alphas2}) THEN_ALL_NEW |
1672
94b8b70f7bc0
Initial proof modifications for alpha_res
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1670
diff
changeset
|
936 |
simp_tac (HOL_basic_ss addsimps @{thms alphas}) THEN_ALL_NEW |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
937 |
asm_full_simp_tac (HOL_basic_ss addsimps (@{thm supp_Pair} :: sym_eqvts ctxt)) THEN_ALL_NEW |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
938 |
asm_full_simp_tac (HOL_basic_ss addsimps (@{thm Pair_eq} :: all_eqvts ctxt)) THEN_ALL_NEW |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
939 |
simp_tac (HOL_basic_ss addsimps @{thms supp_at_base[symmetric,simplified supp_def]}) THEN_ALL_NEW |
1625
6ad74d73e1b1
Support proof modification for Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1622
diff
changeset
|
940 |
simp_tac (HOL_basic_ss addsimps @{thms Collect_disj_eq[symmetric]}) THEN_ALL_NEW |
6ad74d73e1b1
Support proof modification for Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1622
diff
changeset
|
941 |
simp_tac (HOL_basic_ss addsimps @{thms infinite_Un[symmetric]}) THEN_ALL_NEW |
6ad74d73e1b1
Support proof modification for Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1622
diff
changeset
|
942 |
simp_tac (HOL_basic_ss addsimps @{thms Collect_disj_eq[symmetric]}) THEN_ALL_NEW |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
943 |
simp_tac (HOL_basic_ss addsimps @{thms de_Morgan_conj[symmetric]}) THEN_ALL_NEW |
1625
6ad74d73e1b1
Support proof modification for Core Haskell.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1622
diff
changeset
|
944 |
simp_tac (HOL_basic_ss addsimps @{thms ex_simps(1,2)[symmetric]}) THEN_ALL_NEW |
1553
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
945 |
simp_tac (HOL_ss addsimps @{thms Collect_const finite.emptyI}) |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
946 |
*} |
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1547
diff
changeset
|
947 |
|
1653
a2142526bb01
Removed another cheat and cleaned the code a bit.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1651
diff
changeset
|
948 |
|
1650
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
949 |
|
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
950 |
ML {* |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
951 |
fun build_eqvt_gl pi frees fnctn ctxt = |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
952 |
let |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
953 |
val typ = domain_type (fastype_of fnctn); |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
954 |
val arg = the (AList.lookup (op=) frees typ); |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
955 |
in |
1680
4abf7d631ef0
Equivariance when bn functions are lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1679
diff
changeset
|
956 |
([HOLogic.mk_eq ((perm_arg (fnctn $ arg) $ pi $ (fnctn $ arg)), (fnctn $ (perm_arg arg $ pi $ arg)))], ctxt) |
1650
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
957 |
end |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
958 |
*} |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
959 |
|
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
960 |
ML {* |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
961 |
fun prove_eqvt tys ind simps funs ctxt = |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
962 |
let |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
963 |
val ([pi], ctxt') = Variable.variant_fixes ["p"] ctxt; |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
964 |
val pi = Free (pi, @{typ perm}); |
1680
4abf7d631ef0
Equivariance when bn functions are lists.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1679
diff
changeset
|
965 |
val tac = asm_full_simp_tac (HOL_ss addsimps (@{thms atom_eqvt permute_list.simps} @ simps @ all_eqvts ctxt')) |
1650
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
966 |
val ths_loc = prove_by_induct tys (build_eqvt_gl pi) ind tac funs ctxt' |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
967 |
val ths = Variable.export ctxt' ctxt ths_loc |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
968 |
val add_eqvt = Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add) |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
969 |
in |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
970 |
(ths, snd (Local_Theory.note ((Binding.empty, [add_eqvt]), ths) ctxt)) |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
971 |
end |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
972 |
*} |
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1625
diff
changeset
|
973 |
|
1651
f731e9aff866
Proper bn_rsp, for bn functions calling each other.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1650
diff
changeset
|
974 |
end |