| author | wenzelm | 
| Mon, 16 Apr 2012 21:37:08 +0200 | |
| changeset 47497 | c78c6e1ec75d | 
| parent 47302 | 70239da25ef6 | 
| permissions | -rw-r--r-- | 
| 47269 | 1  | 
theory MyList  | 
2  | 
imports Main  | 
|
3  | 
begin  | 
|
4  | 
||
| 47302 | 5  | 
datatype 'a list = Nil | Cons 'a "'a list"  | 
| 47269 | 6  | 
|
7  | 
fun app :: "'a list => 'a list => 'a list" where  | 
|
8  | 
"app Nil ys = ys" |  | 
|
9  | 
"app (Cons x xs) ys = Cons x (app xs ys)"  | 
|
10  | 
||
11  | 
fun rev :: "'a list => 'a list" where  | 
|
12  | 
"rev Nil = Nil" |  | 
|
13  | 
"rev (Cons x xs) = app (rev xs) (Cons x Nil)"  | 
|
14  | 
||
15  | 
value "rev(Cons True (Cons False Nil))"  | 
|
16  | 
||
17  | 
end  |