author | nipkow |
Mon, 09 Oct 2006 12:20:39 +0200 | |
changeset 20921 | 24b8536dcf93 |
parent 20842 | f5f69a1059f4 |
child 20922 | 14873e42659c |
permissions | -rw-r--r-- |
19829 | 1 |
(* ID: $Id$ |
2 |
Authors: Klaus Aehlig, Tobias Nipkow |
|
20807 | 3 |
*) |
19829 | 4 |
|
20807 | 5 |
header "Test of normalization function" |
19829 | 6 |
|
7 |
theory NormalForm |
|
8 |
imports Main |
|
9 |
begin |
|
10 |
||
19971 | 11 |
lemma "p \<longrightarrow> True" by normalization |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20352
diff
changeset
|
12 |
declare disj_assoc [code func] |
20595 | 13 |
lemma "((P | Q) | R) = (P | (Q | R))" by normalization |
19971 | 14 |
lemma "0 + (n::nat) = n" by normalization |
20595 | 15 |
lemma "0 + Suc n = Suc n" by normalization |
16 |
lemma "Suc n + Suc m = n + Suc (Suc m)" by normalization |
|
19971 | 17 |
lemma "~((0::nat) < (0::nat))" by normalization |
18 |
||
19829 | 19 |
datatype n = Z | S n |
20 |
consts |
|
20842 | 21 |
add :: "n \<Rightarrow> n \<Rightarrow> n" |
22 |
add2 :: "n \<Rightarrow> n \<Rightarrow> n" |
|
23 |
mul :: "n \<Rightarrow> n \<Rightarrow> n" |
|
24 |
mul2 :: "n \<Rightarrow> n \<Rightarrow> n" |
|
25 |
exp :: "n \<Rightarrow> n \<Rightarrow> n" |
|
19829 | 26 |
primrec |
20842 | 27 |
"add Z = id" |
28 |
"add (S m) = S o add m" |
|
19829 | 29 |
primrec |
20842 | 30 |
"add2 Z n = n" |
31 |
"add2 (S m) n = S(add2 m n)" |
|
19829 | 32 |
|
33 |
lemma [code]: "add2 (add2 n m) k = add2 n (add2 m k)" |
|
20842 | 34 |
by(induct n) auto |
35 |
lemma [code]: "add2 n (S m) = S (add2 n m)" |
|
36 |
by(induct n) auto |
|
19829 | 37 |
lemma [code]: "add2 n Z = n" |
20842 | 38 |
by(induct n) auto |
19971 | 39 |
|
40 |
lemma "add2 (add2 n m) k = add2 n (add2 m k)" by normalization |
|
41 |
lemma "add2 (add2 (S n) (S m)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization |
|
42 |
lemma "add2 (add2 (S n) (add2 (S m) Z)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization |
|
19829 | 43 |
|
44 |
primrec |
|
20842 | 45 |
"mul Z = (%n. Z)" |
46 |
"mul (S m) = (%n. add (mul m n) n)" |
|
19829 | 47 |
primrec |
20842 | 48 |
"mul2 Z n = Z" |
49 |
"mul2 (S m) n = add2 n (mul2 m n)" |
|
19829 | 50 |
primrec |
20842 | 51 |
"exp m Z = S Z" |
52 |
"exp m (S n) = mul (exp m n) m" |
|
19829 | 53 |
|
19971 | 54 |
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 |
55 |
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 |
|
56 |
lemma "exp (S(S Z)) (S(S(S(S Z)))) = exp (S(S(S(S Z)))) (S(S Z))" by normalization |
|
57 |
||
58 |
lemma "(let ((x,y),(u,v)) = ((Z,Z),(Z,Z)) in add (add x y) (add u v)) = Z" by normalization |
|
20842 | 59 |
lemma "split (%x y. x) (a, b) = a" by normalization |
19971 | 60 |
lemma "(%((x,y),(u,v)). add (add x y) (add u v)) ((Z,Z),(Z,Z)) = Z" by normalization |
61 |
||
62 |
lemma "case Z of Z \<Rightarrow> True | S x \<Rightarrow> False" by normalization |
|
19829 | 63 |
|
20842 | 64 |
lemma "[] @ [] = []" by normalization |
65 |
lemma "[] @ xs = xs" by normalization |
|
66 |
lemma "[a \<Colon> 'd, b, c] @ xs = a # b # c # xs" by normalization |
|
67 |
lemma "[%a::'x. a, %b. b, c] @ xs = (%x. x) # (%x. x) # c # xs" by normalization |
|
68 |
lemma "[%a::'x. a, %b. b, c] @ [u,v] = [%x. x, %x. x, c, u, v]" by normalization |
|
69 |
lemma "map f [x,y,z::'x] = [f x, f y, f z]" by normalization |
|
19829 | 70 |
normal_form "map (%f. f True) [id,g,Not]" |
71 |
normal_form "map (%f. f True) ([id,g,Not] @ fs)" |
|
20842 | 72 |
lemma "rev[a,b,c] = [c, b, a]" by normalization |
19829 | 73 |
normal_form "rev(a#b#cs)" |
20842 | 74 |
lemma "map map [f,g,h] = [map f, map g, map h]" by normalization |
19829 | 75 |
normal_form "map (%F. F [a,b,c::'x]) (map map [f,g,h])" |
76 |
normal_form "map (%F. F ([a,b,c] @ ds)) (map map ([f,g,h]@fs))" |
|
77 |
normal_form "map (%F. F [Z,S Z,S(S Z)]) (map map [S,add (S Z),mul (S(S Z)),id])" |
|
78 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()]" |
|
79 |
normal_form "case xs of [] \<Rightarrow> True | x#xs \<Rightarrow> False" |
|
80 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) xs" |
|
81 |
normal_form "let x = y::'x in [x,x]" |
|
82 |
normal_form "Let y (%x. [x,x])" |
|
83 |
normal_form "case n of Z \<Rightarrow> True | S x \<Rightarrow> False" |
|
84 |
normal_form "(%(x,y). add x y) (S z,S z)" |
|
85 |
normal_form "filter (%x. x) ([True,False,x]@xs)" |
|
86 |
normal_form "filter Not ([True,False,x]@xs)" |
|
87 |
||
20842 | 88 |
lemma "[x,y,z] @ [a,b,c] = [x, y, z, a, b ,c]" by normalization |
19829 | 89 |
normal_form "%(xs, ys). xs @ ys" |
90 |
normal_form "(%(xs, ys). xs @ ys) ([a, b, c], [d, e, f])" |
|
91 |
normal_form "%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True" |
|
92 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()]" |
|
93 |
||
20842 | 94 |
lemma "last [a, b, c] = c" |
95 |
by normalization |
|
96 |
lemma "last ([a, b, c] @ xs) = (if null xs then c else last xs)" |
|
97 |
by normalization |
|
19829 | 98 |
|
20842 | 99 |
lemma "(2::int) + 3 - 1 + (- k) * 2 = 4 + - k * 2" by normalization |
100 |
lemma "(-4::int) * 2 = -8" by normalization |
|
101 |
lemma "abs ((-4::int) + 2 * 1) = 2" by normalization |
|
102 |
lemma "(2::int) + 3 = 5" by normalization |
|
103 |
lemma "(2::int) + 3 * (- 4) * (- 1) = 14" by normalization |
|
104 |
lemma "(2::int) + 3 * (- 4) * 1 + 0 = -10" by normalization |
|
105 |
lemma "(2::int) < 3" by normalization |
|
106 |
lemma "(2::int) <= 3" by normalization |
|
107 |
lemma "abs ((-4::int) + 2 * 1) = 2" by normalization |
|
108 |
lemma "4 - 42 * abs (3 + (-7\<Colon>int)) = -164" by normalization |
|
20352 | 109 |
normal_form "min 0 x" |
110 |
normal_form "min 0 (x::nat)" |
|
20842 | 111 |
lemma "(if (0\<Colon>nat) \<le> (x\<Colon>nat) then 0\<Colon>nat else x) = 0" by normalization |
19829 | 112 |
|
20921 | 113 |
(* Delaying if: FIXME move to HOL.thy(?) *) |
114 |
||
115 |
definition delayed_if :: "bool \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
116 |
"delayed_if b f g = (if b then f() else g())" |
|
117 |
||
118 |
lemma [normal_pre]: "(if b then x else y) == delayed_if b (%u. x) (%u. y)" |
|
119 |
unfolding delayed_if_def by simp |
|
120 |
||
121 |
lemma [code func]: |
|
122 |
shows "delayed_if True f g = f()" and "delayed_if False f g = g()" |
|
123 |
by (auto simp:delayed_if_def) |
|
124 |
||
125 |
hide (open) const delayed_if |
|
126 |
||
127 |
normal_form "OperationalEquality.eq [2..<4] [2,3]" |
|
128 |
(*lemma "OperationalEquality.eq [2..<4] [2,3]" by normalization*) |
|
129 |
||
19829 | 130 |
end |