author | kleing |
Tue, 13 May 2003 08:59:21 +0200 | |
changeset 14024 | 213dcc39358f |
parent 12338 | de0f4a63baa5 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/Sequence.thy |
3275 | 2 |
ID: $Id$ |
12218 | 3 |
Author: Olaf Müller |
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
3071 | 5 |
|
12218 | 6 |
Sequences over flat domains with lifted elements. |
3071 | 7 |
*) |
8 |
||
9 |
Sequence = Seq + |
|
10 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12218
diff
changeset
|
11 |
default type |
3071 | 12 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12218
diff
changeset
|
13 |
types 'a Seq = ('a::type lift)seq |
3071 | 14 |
|
3952 | 15 |
consts |
3071 | 16 |
|
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
4559
diff
changeset
|
17 |
Consq ::"'a => 'a Seq -> 'a Seq" |
3071 | 18 |
Filter ::"('a => bool) => 'a Seq -> 'a Seq" |
19 |
Map ::"('a => 'b) => 'a Seq -> 'b Seq" |
|
20 |
Forall ::"('a => bool) => 'a Seq => bool" |
|
21 |
Last ::"'a Seq -> 'a lift" |
|
22 |
Dropwhile, |
|
23 |
Takewhile ::"('a => bool) => 'a Seq -> 'a Seq" |
|
24 |
Zip ::"'a Seq -> 'b Seq -> ('a * 'b) Seq" |
|
25 |
Flat ::"('a Seq) seq -> 'a Seq" |
|
26 |
||
27 |
Filter2 ::"('a => bool) => 'a Seq -> 'a Seq" |
|
28 |
||
29 |
syntax |
|
30 |
||
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
4559
diff
changeset
|
31 |
"@Consq" ::"'a => 'a Seq => 'a Seq" ("(_/>>_)" [66,65] 65) |
4283 | 32 |
(* list Enumeration *) |
33 |
"_totlist" :: args => 'a Seq ("[(_)!]") |
|
34 |
"_partlist" :: args => 'a Seq ("[(_)?]") |
|
35 |
||
3071 | 36 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12028
diff
changeset
|
37 |
syntax (xsymbols) |
3071 | 38 |
|
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
4559
diff
changeset
|
39 |
"@Consq" ::"'a => 'a Seq => 'a Seq" ("(_\\<leadsto>_)" [66,65] 65) |
4283 | 40 |
|
3071 | 41 |
|
42 |
translations |
|
43 |
||
10835 | 44 |
"a>>s" == "Consq a$s" |
4283 | 45 |
"[x, xs!]" == "x>>[xs!]" |
46 |
"[x!]" == "x>>nil" |
|
47 |
"[x, xs?]" == "x>>[xs?]" |
|
48 |
"[x?]" == "x>>UU" |
|
3071 | 49 |
|
50 |
defs |
|
51 |
||
52 |
||
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
4559
diff
changeset
|
53 |
Consq_def "Consq a == LAM s. Def a ## s" |
3071 | 54 |
|
10835 | 55 |
Filter_def "Filter P == sfilter$(flift2 P)" |
3071 | 56 |
|
10835 | 57 |
Map_def "Map f == smap$(flift2 f)" |
3071 | 58 |
|
59 |
Forall_def "Forall P == sforall (flift2 P)" |
|
60 |
||
61 |
Last_def "Last == slast" |
|
62 |
||
10835 | 63 |
Dropwhile_def "Dropwhile P == sdropwhile$(flift2 P)" |
3071 | 64 |
|
10835 | 65 |
Takewhile_def "Takewhile P == stakewhile$(flift2 P)" |
3071 | 66 |
|
67 |
Flat_def "Flat == sflat" |
|
68 |
||
69 |
Zip_def |
|
10835 | 70 |
"Zip == (fix$(LAM h t1 t2. case t1 of |
3071 | 71 |
nil => nil |
72 |
| x##xs => (case t2 of |
|
73 |
nil => UU |
|
74 |
| y##ys => (case x of |
|
12028 | 75 |
UU => UU |
3071 | 76 |
| Def a => (case y of |
12028 | 77 |
UU => UU |
10835 | 78 |
| Def b => Def (a,b)##(h$xs$ys))))))" |
3071 | 79 |
|
10835 | 80 |
Filter2_def "Filter2 P == (fix$(LAM h t. case t of |
3071 | 81 |
nil => nil |
12028 | 82 |
| x##xs => (case x of UU => UU | Def y => (if P y |
10835 | 83 |
then x##(h$xs) |
84 |
else h$xs))))" |
|
3071 | 85 |
|
86 |
rules |
|
87 |
||
88 |
||
89 |
(* for test purposes should be deleted FIX !! *) |
|
90 |
adm_all "adm f" |
|
91 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12218
diff
changeset
|
92 |
end |