| author | paulson <lp15@cam.ac.uk> | 
| Fri, 18 Feb 2022 21:40:01 +0000 | |
| changeset 75101 | f0e2023f361a | 
| parent 66345 | 882abe912da9 | 
| child 82444 | 7a9164068583 | 
| permissions | -rw-r--r-- | 
| 30946 | 1 | (* Authors: Klaus Aehlig, Tobias Nipkow *) | 
| 19829 | 2 | |
| 61343 | 3 | section \<open>Testing implementation of normalization by evaluation\<close> | 
| 19829 | 4 | |
| 39395 | 5 | theory Normalization_by_Evaluation | 
| 35372 | 6 | imports Complex_Main | 
| 19829 | 7 | begin | 
| 8 | ||
| 21117 | 9 | lemma "True" by normalization | 
| 19971 | 10 | lemma "p \<longrightarrow> True" by normalization | 
| 28350 | 11 | declare disj_assoc [code nbe] | 
| 12 | lemma "((P | Q) | R) = (P | (Q | R))" by normalization | |
| 13 | lemma "0 + (n::nat) = n" by normalization | |
| 14 | lemma "0 + Suc n = Suc n" by normalization | |
| 15 | lemma "Suc n + Suc m = n + Suc (Suc m)" by normalization | |
| 19971 | 16 | lemma "~((0::nat) < (0::nat))" by normalization | 
| 17 | ||
| 58310 | 18 | datatype n = Z | S n | 
| 28350 | 19 | |
| 30946 | 20 | primrec add :: "n \<Rightarrow> n \<Rightarrow> n" where | 
| 21 | "add Z = id" | |
| 22 | | "add (S m) = S o add m" | |
| 23 | ||
| 24 | primrec add2 :: "n \<Rightarrow> n \<Rightarrow> n" where | |
| 25 | "add2 Z n = n" | |
| 26 | | "add2 (S m) n = S(add2 m n)" | |
| 19829 | 27 | |
| 28143 | 28 | declare add2.simps [code] | 
| 28709 | 29 | lemma [code nbe]: "add2 (add2 n m) k = add2 n (add2 m k)" | 
| 28143 | 30 | by (induct n) auto | 
| 20842 | 31 | lemma [code]: "add2 n (S m) = S (add2 n m)" | 
| 32 | by(induct n) auto | |
| 19829 | 33 | lemma [code]: "add2 n Z = n" | 
| 20842 | 34 | by(induct n) auto | 
| 19971 | 35 | |
| 28350 | 36 | lemma "add2 (add2 n m) k = add2 n (add2 m k)" by normalization | 
| 37 | lemma "add2 (add2 (S n) (S m)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization | |
| 38 | lemma "add2 (add2 (S n) (add2 (S m) Z)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization | |
| 19829 | 39 | |
| 30946 | 40 | primrec mul :: "n \<Rightarrow> n \<Rightarrow> n" where | 
| 41 | "mul Z = (%n. Z)" | |
| 42 | | "mul (S m) = (%n. add (mul m n) n)" | |
| 43 | ||
| 44 | primrec mul2 :: "n \<Rightarrow> n \<Rightarrow> n" where | |
| 45 | "mul2 Z n = Z" | |
| 46 | | "mul2 (S m) n = add2 n (mul2 m n)" | |
| 47 | ||
| 48 | primrec exp :: "n \<Rightarrow> n \<Rightarrow> n" where | |
| 49 | "exp m Z = S Z" | |
| 50 | | "exp m (S n) = mul (exp m n) m" | |
| 19829 | 51 | |
| 19971 | 52 | 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 | 
| 53 | 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 | |
| 54 | lemma "exp (S(S Z)) (S(S(S(S Z)))) = exp (S(S(S(S Z)))) (S(S Z))" by normalization | |
| 55 | ||
| 56 | lemma "(let ((x,y),(u,v)) = ((Z,Z),(Z,Z)) in add (add x y) (add u v)) = Z" by normalization | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61343diff
changeset | 57 | lemma "case_prod (%x y. x) (a, b) = a" by normalization | 
| 19971 | 58 | lemma "(%((x,y),(u,v)). add (add x y) (add u v)) ((Z,Z),(Z,Z)) = Z" by normalization | 
| 59 | ||
| 60 | lemma "case Z of Z \<Rightarrow> True | S x \<Rightarrow> False" by normalization | |
| 19829 | 61 | |
| 20842 | 62 | lemma "[] @ [] = []" by normalization | 
| 28350 | 63 | lemma "map f [x,y,z::'x] = [f x, f y, f z]" by normalization | 
| 64 | lemma "[a, b, c] @ xs = a # b # c # xs" by normalization | |
| 65 | lemma "[] @ xs = xs" by normalization | |
| 66 | lemma "map (%f. f True) [id, g, Not] = [True, g True, False]" by normalization | |
| 67 | ||
| 28422 | 68 | lemma "map (%f. f True) ([id, g, Not] @ fs) = [True, g True, False] @ map (%f. f True) fs" | 
| 41037 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 69 | by normalization rule | 
| 28350 | 70 | lemma "rev [a, b, c] = [c, b, a]" by normalization | 
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 71 | value [nbe] "rev (a#b#cs) = rev cs @ [b, a]" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 72 | value [nbe] "map (%F. F [a,b,c::'x]) (map map [f,g,h])" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 73 | value [nbe] "map (%F. F ([a,b,c] @ ds)) (map map ([f,g,h]@fs))" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 74 | value [nbe] "map (%F. F [Z,S Z,S(S Z)]) (map map [S,add (S Z),mul (S(S Z)),id])" | 
| 25934 | 75 | lemma "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()] = [False, True]" | 
| 76 | by normalization | |
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 77 | value [nbe] "case xs of [] \<Rightarrow> True | x#xs \<Rightarrow> False" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 78 | value [nbe] "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) xs = P" | 
| 28350 | 79 | lemma "let x = y in [x, x] = [y, y]" by normalization | 
| 80 | lemma "Let y (%x. [x,x]) = [y, y]" by normalization | |
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 81 | value [nbe] "case n of Z \<Rightarrow> True | S x \<Rightarrow> False" | 
| 28350 | 82 | lemma "(%(x,y). add x y) (S z,S z) = S (add z (S z))" by normalization | 
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 83 | value [nbe] "filter (%x. x) ([True,False,x]@xs)" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 84 | value [nbe] "filter Not ([True,False,x]@xs)" | 
| 19829 | 85 | |
| 28350 | 86 | lemma "[x,y,z] @ [a,b,c] = [x, y, z, a, b, c]" by normalization | 
| 87 | lemma "(%(xs, ys). xs @ ys) ([a, b, c], [d, e, f]) = [a, b, c, d, e, f]" by normalization | |
| 25100 | 88 | lemma "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()] = [False, True]" by normalization | 
| 19829 | 89 | |
| 28350 | 90 | lemma "last [a, b, c] = c" by normalization | 
| 91 | lemma "last ([a, b, c] @ xs) = last (c # xs)" by normalization | |
| 19829 | 92 | |
| 28350 | 93 | lemma "(2::int) + 3 - 1 + (- k) * 2 = 4 + - k * 2" by normalization | 
| 20842 | 94 | lemma "(-4::int) * 2 = -8" by normalization | 
| 61945 | 95 | lemma "\<bar>(-4::int) + 2 * 1\<bar> = 2" by normalization | 
| 20842 | 96 | lemma "(2::int) + 3 = 5" by normalization | 
| 97 | lemma "(2::int) + 3 * (- 4) * (- 1) = 14" by normalization | |
| 98 | lemma "(2::int) + 3 * (- 4) * 1 + 0 = -10" by normalization | |
| 99 | lemma "(2::int) < 3" by normalization | |
| 100 | lemma "(2::int) <= 3" by normalization | |
| 61945 | 101 | lemma "\<bar>(-4::int) + 2 * 1\<bar> = 2" by normalization | 
| 102 | lemma "4 - 42 * \<bar>3 + (-7::int)\<bar> = -164" by normalization | |
| 61076 | 103 | lemma "(if (0::nat) \<le> (x::nat) then 0::nat else x) = 0" by normalization | 
| 22394 | 104 | lemma "4 = Suc (Suc (Suc (Suc 0)))" by normalization | 
| 105 | lemma "nat 4 = Suc (Suc (Suc (Suc 0)))" by normalization | |
| 25100 | 106 | lemma "[Suc 0, 0] = [Suc 0, 0]" by normalization | 
| 107 | lemma "max (Suc 0) 0 = Suc 0" by normalization | |
| 25187 | 108 | lemma "(42::rat) / 1704 = 1 / 284 + 3 / 142" by normalization | 
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 109 | value [nbe] "Suc 0 \<in> set ms" | 
| 20922 | 110 | |
| 40730 
2aa0390a2da7
nbe decides equality of abstractions by extensionality
 haftmann parents: 
39395diff
changeset | 111 | (* non-left-linear patterns, equality by extensionality *) | 
| 
2aa0390a2da7
nbe decides equality of abstractions by extensionality
 haftmann parents: 
39395diff
changeset | 112 | |
| 28350 | 113 | lemma "f = f" by normalization | 
| 114 | lemma "f x = f x" by normalization | |
| 115 | lemma "(f o g) x = f (g x)" by normalization | |
| 116 | lemma "(f o id) x = f x" by normalization | |
| 40730 
2aa0390a2da7
nbe decides equality of abstractions by extensionality
 haftmann parents: 
39395diff
changeset | 117 | lemma "(id :: bool \<Rightarrow> bool) = id" by normalization | 
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 118 | value [nbe] "(\<lambda>x. x)" | 
| 21987 | 119 | |
| 23396 | 120 | (* Church numerals: *) | 
| 121 | ||
| 66345 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 122 | value [nbe] "(%m n f x. m f (n f x)) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 123 | value [nbe] "(%m n f x. m (n f) x) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))" | 
| 
882abe912da9
do not fall back on nbe if plain evaluation fails
 haftmann parents: 
61945diff
changeset | 124 | value [nbe] "(%m n. n m) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))" | 
| 23396 | 125 | |
| 32544 | 126 | (* handling of type classes in connection with equality *) | 
| 127 | ||
| 128 | lemma "map f [x, y] = [f x, f y]" by normalization | |
| 129 | lemma "(map f [x, y], w) = ([f x, f y], w)" by normalization | |
| 61076 | 130 | lemma "map f [x, y] = [f x :: 'a::semigroup_add, f y]" by normalization | 
| 131 | lemma "map f [x :: 'a::semigroup_add, y] = [f x, f y]" by normalization | |
| 132 | lemma "(map f [x :: 'a::semigroup_add, y], w :: 'b::finite) = ([f x, f y], w)" by normalization | |
| 32544 | 133 | |
| 19829 | 134 | end |