| 19829 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Authors:    Klaus Aehlig, Tobias Nipkow
 | 
|  |      3 | 
 | 
|  |      4 | Test of normalization function
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | theory NormalForm
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
| 19971 |     11 | lemma "p \<longrightarrow> True" by normalization
 | 
| 19829 |     12 | 
 | 
|  |     13 | (* FIXME Eventually the code generator should be able to handle this
 | 
|  |     14 | by re-generating the existing code for "or":
 | 
|  |     15 | 
 | 
|  |     16 | declare disj_assoc[code]
 | 
|  |     17 | 
 | 
|  |     18 | normal_form "(P | Q) | R"
 | 
|  |     19 | 
 | 
|  |     20 | *)
 | 
|  |     21 | 
 | 
|  |     22 | 
 | 
| 19971 |     23 | lemma "0 + (n::nat) = n" by normalization
 | 
|  |     24 | lemma "0 + Suc(n) = Suc n" by normalization
 | 
|  |     25 | lemma "Suc(n) + Suc m = n + Suc(Suc m)" by normalization
 | 
|  |     26 | lemma "~((0::nat) < (0::nat))" by normalization
 | 
|  |     27 | 
 | 
|  |     28 | 
 | 
| 19829 |     29 | datatype n = Z | S n
 | 
|  |     30 | consts
 | 
|  |     31 |  add :: "n \<Rightarrow> n \<Rightarrow> n"
 | 
|  |     32 |  add2 :: "n \<Rightarrow> n \<Rightarrow> n"
 | 
|  |     33 |  mul :: "n \<Rightarrow> n \<Rightarrow> n"
 | 
|  |     34 |  mul2 :: "n \<Rightarrow> n \<Rightarrow> n"
 | 
|  |     35 |  exp :: "n \<Rightarrow> n \<Rightarrow> n"
 | 
|  |     36 | primrec
 | 
|  |     37 | "add Z = id"
 | 
|  |     38 | "add (S m) = S o add m"
 | 
|  |     39 | primrec
 | 
|  |     40 | "add2 Z n = n"
 | 
|  |     41 | "add2 (S m) n = S(add2 m n)"
 | 
|  |     42 | 
 | 
|  |     43 | lemma [code]: "add2 (add2 n m) k = add2 n (add2 m k)"
 | 
|  |     44 | by(induct n, auto)
 | 
|  |     45 | lemma [code]: "add2 n (S m) =  S(add2 n m)"
 | 
|  |     46 | by(induct n, auto)
 | 
|  |     47 | lemma [code]: "add2 n Z = n"
 | 
|  |     48 | by(induct n, auto)
 | 
| 19971 |     49 | 
 | 
|  |     50 | lemma "add2 (add2 n m) k = add2 n (add2 m k)" by normalization
 | 
|  |     51 | lemma "add2 (add2 (S n) (S m)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization
 | 
|  |     52 | lemma "add2 (add2 (S n) (add2 (S m) Z)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization
 | 
| 19829 |     53 | 
 | 
|  |     54 | primrec
 | 
|  |     55 | "mul Z = (%n. Z)"
 | 
|  |     56 | "mul (S m) = (%n. add (mul m n) n)"
 | 
|  |     57 | primrec
 | 
|  |     58 | "mul2 Z n = Z"
 | 
|  |     59 | "mul2 (S m) n = add2 n (mul2 m n)"
 | 
|  |     60 | primrec
 | 
|  |     61 | "exp m Z = S Z"
 | 
|  |     62 | "exp m (S n) = mul (exp m n) m"
 | 
|  |     63 | 
 | 
| 19971 |     64 | lemma "mul2 (S(S(S(S(S Z))))) (S(S(S Z))) = S(S(S(S(S(S(S(S(S(S(S(S(S(S(S Z))))))))))))))" by normalization
 | 
|  |     65 | lemma "mul (S(S(S(S(S Z))))) (S(S(S Z))) = S(S(S(S(S(S(S(S(S(S(S(S(S(S(S Z))))))))))))))" by normalization
 | 
|  |     66 | lemma "exp (S(S Z)) (S(S(S(S Z)))) = exp (S(S(S(S Z)))) (S(S Z))" by normalization
 | 
|  |     67 | 
 | 
|  |     68 | lemma "(let ((x,y),(u,v)) = ((Z,Z),(Z,Z)) in add (add x y) (add u v)) = Z" by normalization
 | 
|  |     69 | lemma "(%((x,y),(u,v)). add (add x y) (add u v)) ((Z,Z),(Z,Z)) = Z" by normalization
 | 
|  |     70 | 
 | 
|  |     71 | lemma "case Z of Z \<Rightarrow> True | S x \<Rightarrow> False" by normalization
 | 
| 19829 |     72 | 
 | 
|  |     73 | normal_form "[] @ []"
 | 
|  |     74 | normal_form "[] @ xs"
 | 
|  |     75 | normal_form "[a::'d,b,c] @ xs"
 | 
|  |     76 | normal_form "[%a::'x. a, %b. b, c] @ xs"
 | 
|  |     77 | normal_form "[%a::'x. a, %b. b, c] @ [u,v]"
 | 
|  |     78 | normal_form "map f (xs::'c list)"
 | 
|  |     79 | normal_form "map f [x,y,z::'x]"
 | 
|  |     80 | normal_form "map (%f. f True) [id,g,Not]"
 | 
|  |     81 | normal_form "map (%f. f True) ([id,g,Not] @ fs)"
 | 
|  |     82 | normal_form "rev[a,b,c]"
 | 
|  |     83 | normal_form "rev(a#b#cs)"
 | 
|  |     84 | normal_form "map map [f,g,h]"
 | 
|  |     85 | normal_form "map (%F. F [a,b,c::'x]) (map map [f,g,h])"
 | 
|  |     86 | normal_form "map (%F. F ([a,b,c] @ ds)) (map map ([f,g,h]@fs))"
 | 
|  |     87 | normal_form "map (%F. F [Z,S Z,S(S Z)]) (map map [S,add (S Z),mul (S(S Z)),id])"
 | 
|  |     88 | normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()]"
 | 
|  |     89 | normal_form "case xs of [] \<Rightarrow> True | x#xs \<Rightarrow> False"
 | 
|  |     90 | normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) xs"
 | 
|  |     91 | normal_form "let x = y::'x in [x,x]"
 | 
|  |     92 | normal_form "Let y (%x. [x,x])"
 | 
|  |     93 | normal_form "case n of Z \<Rightarrow> True | S x \<Rightarrow> False"
 | 
|  |     94 | normal_form "(%(x,y). add x y) (S z,S z)"
 | 
|  |     95 | normal_form "filter (%x. x) ([True,False,x]@xs)"
 | 
|  |     96 | normal_form "filter Not ([True,False,x]@xs)"
 | 
|  |     97 | 
 | 
|  |     98 | normal_form "[x,y,z] @ [a,b,c]"
 | 
|  |     99 | normal_form "%(xs, ys). xs @ ys"
 | 
|  |    100 | normal_form "(%(xs, ys). xs @ ys) ([a, b, c], [d, e, f])"
 | 
|  |    101 | normal_form "%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True"
 | 
|  |    102 | normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()]"
 | 
|  |    103 | 
 | 
|  |    104 | normal_form "last[a,b,c]"
 | 
|  |    105 | normal_form "last([a,b,c]@xs)"
 | 
|  |    106 | 
 | 
| 20191 |    107 | normal_form "max 0 x"
 | 
| 19829 |    108 | 
 | 
|  |    109 | text {*
 | 
|  |    110 |   Numerals still take their time\<dots>
 | 
|  |    111 | *}
 | 
|  |    112 | 
 | 
|  |    113 | end
 |