author | Christian Urban <urbanc@in.tum.de> |
Wed, 28 Apr 2010 06:40:10 +0200 | |
changeset 1966 | b6b3374a402d |
parent 1965 | 4a3c05fe2bc5 |
child 1967 | 700024ec18da |
permissions | -rw-r--r-- |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory NewFv |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Nominal-General/Nominal2_Atoms" |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
"Abs" "Perm" "Rsp" "Nominal2_FSet" |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
datatype bmodes = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
8 |
BEmy of int |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
| BLst of ((term option * int) list) * (int list) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
10 |
| BSet of ((term option * int) list) * (int list) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
| BRes of ((term option * int) list) * (int list) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
ML {* |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
15 |
fun mk_singleton_atom x = HOLogic.mk_set @{typ atom} [mk_atom x]; |
1962
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents:
1960
diff
changeset
|
16 |
|
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
17 |
val noatoms = @{term "{} :: atom set"}; |
1963 | 18 |
|
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
19 |
fun mk_union sets = |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
20 |
fold (fn a => fn b => |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
21 |
if a = noatoms then b else |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
22 |
if b = noatoms then a else |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
23 |
if a = b then a else |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
24 |
HOLogic.mk_binop @{const_name sup} (a, b)) (rev sets) noatoms; |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
ML {* |
1963 | 28 |
fun is_atom thy ty = |
29 |
Sign.of_sort thy (ty, @{sort at}) |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
30 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
fun is_atom_set thy (Type ("fun", [t, @{typ bool}])) = is_atom thy t |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
32 |
| is_atom_set _ _ = false; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
fun is_atom_fset thy (Type ("FSet.fset", [t])) = is_atom thy t |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
| is_atom_fset _ _ = false; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
fun mk_atom_set t = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
val ty = fastype_of t; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
val atom_ty = HOLogic.dest_setT ty --> @{typ atom}; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
val img_ty = atom_ty --> ty --> @{typ "atom set"}; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
in |
1963 | 43 |
(Const (@{const_name image}, img_ty) $ mk_atom_ty atom_ty t) |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
44 |
end; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
45 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
46 |
fun mk_atom_fset t = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
47 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
val ty = fastype_of t; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
49 |
val atom_ty = dest_fsetT ty --> @{typ atom}; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
50 |
val fmap_ty = atom_ty --> ty --> @{typ "atom fset"}; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
51 |
val fset_to_set = @{term "fset_to_set :: atom fset \<Rightarrow> atom set"} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
52 |
in |
1963 | 53 |
fset_to_set $ (Const (@{const_name fmap}, fmap_ty) $ mk_atom_ty atom_ty t) |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
54 |
end; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
55 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
56 |
fun mk_diff a b = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
57 |
if b = noatoms then a else |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
58 |
if b = a then noatoms else |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
59 |
HOLogic.mk_binop @{const_name minus} (a, b); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
61 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
62 |
ML {* |
1963 | 63 |
fun atoms thy t = |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
64 |
let |
1963 | 65 |
val ty = fastype_of t; |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
66 |
in |
1963 | 67 |
if is_atom thy ty |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
68 |
then mk_singleton_atom t |
1963 | 69 |
else if is_atom_set thy ty |
70 |
then mk_atom_set t |
|
71 |
else if is_atom_fset thy ty |
|
72 |
then mk_atom_fset t |
|
73 |
else noatoms |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
74 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
75 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
77 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
fun setify x = |
1963 | 79 |
if fastype_of x = @{typ "atom list"} |
80 |
then @{term "set::atom list \<Rightarrow> atom set"} $ x |
|
81 |
else x |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
82 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
83 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
85 |
fun fv_bm_lsts thy dts args fv_frees bn_fvbn binds bodys = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
86 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
87 |
fun fv_body i = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
88 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
89 |
val x = nth args i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
90 |
val dt = nth dts i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
91 |
in |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
92 |
if Datatype_Aux.is_rec_type dt |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
93 |
then nth fv_frees (Datatype_Aux.body_index dt) $ x |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
94 |
else atoms thy x |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
95 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
96 |
val fv_bodys = mk_union (map fv_body bodys) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
97 |
val bound_vars = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
98 |
case binds of |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
99 |
[(SOME bn, i)] => setify (bn $ nth args i) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
100 |
| _ => mk_union (map fv_body (map snd binds)); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
101 |
val non_rec_vars = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
102 |
case binds of |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
103 |
[(SOME bn, i)] => |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
104 |
if i mem bodys |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
105 |
then ((the (AList.lookup (op=) bn_fvbn bn)) $ nth args i) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
106 |
else noatoms |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
107 |
| _ => noatoms |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
108 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
109 |
mk_union [mk_diff fv_bodys bound_vars, non_rec_vars] |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
110 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
111 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
112 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
113 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
114 |
fun fv_bn_bm thy dts args fv_frees bn_fvbn args_in_bn bm = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
case bm of |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
116 |
BEmy i => |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
117 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
118 |
val x = nth args i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
119 |
val dt = nth dts i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
120 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
121 |
case AList.lookup (op=) args_in_bn i of |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
122 |
NONE => if Datatype_Aux.is_rec_type dt |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
123 |
then nth fv_frees (Datatype_Aux.body_index dt) $ x |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
124 |
else atoms thy x |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
125 |
| SOME (SOME (f : term)) => (the (AList.lookup (op=) bn_fvbn f)) $ x |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
126 |
| SOME NONE => noatoms |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
127 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
128 |
| BLst (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
129 |
| BSet (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
130 |
| BRes (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
131 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
132 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
133 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
134 |
fun fv_bn thy (dt_info : Datatype_Aux.info) fv_frees |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
135 |
bn_fvbn bmll (fvbn, (_, ith_dtyp, args_in_bns)) = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
136 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
137 |
val {descr, sorts, ...} = dt_info; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
138 |
fun fv_bn_constr (cname, dts) (args_in_bn, bml) = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
139 |
let |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
140 |
val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
141 |
val names = Datatype_Prop.make_tnames Ts; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
142 |
val args = map Free (names ~~ Ts); |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
143 |
val c = Const (cname, Ts ---> (nth_dtyp descr sorts ith_dtyp)); |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
144 |
val fv_bn_bm = fv_bn_bm thy dts args fv_frees bn_fvbn args_in_bn |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
145 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
146 |
HOLogic.mk_Trueprop (HOLogic.mk_eq |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
147 |
(fvbn $ list_comb (c, args), mk_union (map fv_bn_bm bml))) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
148 |
end; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
149 |
val (_, (_, _, constrs)) = nth descr ith_dtyp; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
150 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
151 |
map2 fv_bn_constr constrs (args_in_bns ~~ bmll) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
152 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
153 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
154 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
155 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
156 |
fun fv_bns thy dt_info fv_frees bns bmlll = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
157 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
158 |
fun mk_fvbn_free (bn, ith, _) = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
159 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
160 |
val fvbn_name = "fv_" ^ (Long_Name.base_name (fst (dest_Const bn))); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
161 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
162 |
(fvbn_name, Free (fvbn_name, fastype_of (nth fv_frees ith))) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
163 |
end; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
164 |
val (fvbn_names, fvbn_frees) = split_list (map mk_fvbn_free bns); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
165 |
val bn_fvbn = (map (fn (bn, _, _) => bn) bns) ~~ fvbn_frees |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
166 |
val bmlls = map (fn (_, i, _) => nth bmlll i) bns; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
167 |
val eqs = map2 (fv_bn thy dt_info fv_frees bn_fvbn) bmlls (fvbn_frees ~~ bns); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
168 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
169 |
(bn_fvbn, (fvbn_names, eqs)) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
170 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
171 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
172 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
173 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
174 |
fun fv_bm thy dts args fv_frees bn_fvbn bm = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
175 |
case bm of |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
176 |
BEmy i => |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
177 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
178 |
val x = nth args i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
179 |
val dt = nth dts i; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
180 |
in |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
181 |
if Datatype_Aux.is_rec_type dt |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
182 |
then nth fv_frees (Datatype_Aux.body_index dt) $ x |
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
183 |
else atoms thy x |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
184 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
185 |
| BLst (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
186 |
| BSet (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
187 |
| BRes (x, y) => fv_bm_lsts thy dts args fv_frees bn_fvbn x y |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
188 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
189 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
190 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
191 |
fun fv thy (dt_info : Datatype_Aux.info) fv_frees bn_fvbn bmll (fv_free, ith_dtyp) = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
192 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
193 |
val {descr, sorts, ...} = dt_info; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
194 |
fun fv_constr (cname, dts) bml = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
195 |
let |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
196 |
val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
197 |
val names = Datatype_Prop.make_tnames Ts; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
198 |
val args = map Free (names ~~ Ts); |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
199 |
val c = Const (cname, Ts ---> (nth_dtyp descr sorts ith_dtyp)); |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
200 |
val fv_bn_bm = fv_bm thy dts args fv_frees bn_fvbn |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
201 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
202 |
HOLogic.mk_Trueprop (HOLogic.mk_eq |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
203 |
(fv_free $ list_comb (c, args), mk_union (map fv_bn_bm bml))) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
204 |
end; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
205 |
val (_, (_, _, constrs)) = nth descr ith_dtyp; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
206 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
207 |
map2 fv_constr constrs bmll |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
208 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
209 |
*} |
1963 | 210 |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
211 |
ML {* |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
212 |
val by_pat_completeness_auto = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
213 |
Proof.global_terminal_proof |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
214 |
(Method.Basic Pat_Completeness.pat_completeness, |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
215 |
SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none)))) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
216 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
217 |
fun termination_by method = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
218 |
Function.termination_proof NONE |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
219 |
#> Proof.global_terminal_proof (Method.Basic method, NONE) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
220 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
221 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
222 |
ML {* |
1963 | 223 |
fun define_raw_fv dt_info bns bmlll lthy = |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
224 |
let |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
225 |
val thy = ProofContext.theory_of lthy; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
226 |
val {descr, sorts, ...} = dt_info; |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
227 |
|
1966
b6b3374a402d
factured out common functionality of prefixing the dt-names with a string
Christian Urban <urbanc@in.tum.de>
parents:
1965
diff
changeset
|
228 |
val fv_names = prefix_dt_names descr sorts "fv_" |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
229 |
|
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
230 |
val fv_types = map (fn (i, _) => nth_dtyp descr sorts i --> @{typ "atom set"}) descr; |
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
231 |
val fv_frees = map Free (fv_names ~~ fv_types); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
232 |
val (bn_fvbn, (fv_bn_names, fv_bn_eqs)) = fv_bns thy dt_info fv_frees bns bmlll; |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
233 |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
234 |
val fvbns = map snd bn_fvbn; |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
235 |
val fv_nums = 0 upto (length fv_frees - 1) |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
236 |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
237 |
val fv_eqs = map2 (fv thy dt_info fv_frees bn_fvbn) bmlll (fv_frees ~~ fv_nums); |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
238 |
val fv_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
239 |
val fv_eqs_binds = map (fn x => (Attrib.empty_binding, x)) (flat fv_eqs @ flat fv_bn_eqs) |
1965
4a3c05fe2bc5
closed Datatype_Aux; replaced nth_dtyp by the function used in Perm.thy
Christian Urban <urbanc@in.tum.de>
parents:
1963
diff
changeset
|
240 |
|
1960
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
241 |
val lthy' = |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
242 |
lthy |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
243 |
|> Function.add_function fv_names fv_eqs_binds Function_Common.default_config |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
244 |
|> by_pat_completeness_auto |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
245 |
|> Local_Theory.restore |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
246 |
|> termination_by (Lexicographic_Order.lexicographic_order) |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
247 |
in |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
248 |
(fv_frees @ fvbns, lthy') |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
249 |
end |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
250 |
*} |
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
251 |
|
47e2e91705f3
Rewrote FV code and included the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
252 |
end |