23930
|
1 |
(* ID: $Id$
|
|
2 |
Authors: Klaus Aehlig, Tobias Nipkow
|
|
3 |
*)
|
|
4 |
|
|
5 |
header {* Alternativ implementation of "normalization by evaluation" for HOL, including test examples *}
|
|
6 |
|
|
7 |
theory Nbe
|
|
8 |
imports Main
|
|
9 |
uses
|
|
10 |
"~~/src/Tools/Nbe/nbe_eval.ML"
|
|
11 |
"~~/src/Tools/Nbe/nbe_package.ML"
|
|
12 |
begin
|
|
13 |
|
23999
|
14 |
lemma [code inline]: "If b f g = bool_case f g b" by auto
|
|
15 |
lemma [code func]: "null xs \<longleftrightarrow> (case xs of [] \<Rightarrow> True | _ \<Rightarrow> False)"
|
|
16 |
by (cases xs) auto
|
|
17 |
|
|
18 |
|
23930
|
19 |
ML {* reset Toplevel.debug *}
|
|
20 |
|
|
21 |
setup Nbe_Package.setup
|
|
22 |
|
|
23 |
method_setup normalization = {*
|
|
24 |
Method.no_args (Method.SIMPLE_METHOD'
|
|
25 |
(CONVERSION (ObjectLogic.judgment_conv Nbe_Package.normalization_conv)
|
|
26 |
THEN' resolve_tac [TrueI, refl]))
|
|
27 |
*} "solve goal by normalization"
|
|
28 |
|
|
29 |
|
|
30 |
text {* lazy @{const If} *}
|
|
31 |
|
|
32 |
definition
|
|
33 |
if_delayed :: "bool \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> 'a" where
|
|
34 |
[code func del]: "if_delayed b f g = (if b then f True else g False)"
|
|
35 |
|
|
36 |
lemma [code func]:
|
|
37 |
shows "if_delayed True f g = f True"
|
|
38 |
and "if_delayed False f g = g False"
|
|
39 |
unfolding if_delayed_def by simp_all
|
|
40 |
|
|
41 |
lemma [normal pre, symmetric, normal post]:
|
|
42 |
"(if b then x else y) = if_delayed b (\<lambda>_. x) (\<lambda>_. y)"
|
|
43 |
unfolding if_delayed_def ..
|
|
44 |
|
|
45 |
hide (open) const if_delayed
|
|
46 |
|
23998
|
47 |
lemma "True"
|
|
48 |
by normalization
|
23930
|
49 |
lemma "x = x" by normalization
|
|
50 |
lemma "True \<or> False"
|
|
51 |
by normalization
|
|
52 |
lemma "True \<or> p" by normalization
|
|
53 |
lemma "p \<longrightarrow> True"
|
|
54 |
by normalization
|
|
55 |
declare disj_assoc [code func]
|
|
56 |
lemma "((P | Q) | R) = (P | (Q | R))" by normalization
|
|
57 |
declare disj_assoc [code func del]
|
|
58 |
lemma "0 + (n::nat) = n" by normalization
|
|
59 |
lemma "0 + Suc n = Suc n" by normalization
|
|
60 |
lemma "Suc n + Suc m = n + Suc (Suc m)" by normalization
|
|
61 |
lemma "~((0::nat) < (0::nat))" by normalization
|
|
62 |
|
|
63 |
datatype n = Z | S n
|
|
64 |
consts
|
|
65 |
add :: "n \<Rightarrow> n \<Rightarrow> n"
|
|
66 |
add2 :: "n \<Rightarrow> n \<Rightarrow> n"
|
|
67 |
mul :: "n \<Rightarrow> n \<Rightarrow> n"
|
|
68 |
mul2 :: "n \<Rightarrow> n \<Rightarrow> n"
|
|
69 |
exp :: "n \<Rightarrow> n \<Rightarrow> n"
|
|
70 |
primrec
|
|
71 |
"add Z = id"
|
|
72 |
"add (S m) = S o add m"
|
|
73 |
primrec
|
|
74 |
"add2 Z n = n"
|
|
75 |
"add2 (S m) n = S(add2 m n)"
|
|
76 |
|
|
77 |
lemma [code]: "add2 (add2 n m) k = add2 n (add2 m k)"
|
|
78 |
by(induct n) auto
|
|
79 |
lemma [code]: "add2 n (S m) = S (add2 n m)"
|
|
80 |
by(induct n) auto
|
|
81 |
lemma [code]: "add2 n Z = n"
|
|
82 |
by(induct n) auto
|
|
83 |
|
|
84 |
lemma "add2 (add2 n m) k = add2 n (add2 m k)" by normalization
|
|
85 |
lemma "add2 (add2 (S n) (S m)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization
|
|
86 |
lemma "add2 (add2 (S n) (add2 (S m) Z)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization
|
|
87 |
|
|
88 |
primrec
|
|
89 |
"mul Z = (%n. Z)"
|
|
90 |
"mul (S m) = (%n. add (mul m n) n)"
|
|
91 |
primrec
|
|
92 |
"mul2 Z n = Z"
|
|
93 |
"mul2 (S m) n = add2 n (mul2 m n)"
|
|
94 |
primrec
|
|
95 |
"exp m Z = S Z"
|
|
96 |
"exp m (S n) = mul (exp m n) m"
|
|
97 |
|
|
98 |
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
|
|
99 |
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
|
|
100 |
lemma "exp (S(S Z)) (S(S(S(S Z)))) = exp (S(S(S(S Z)))) (S(S Z))" by normalization
|
|
101 |
|
23998
|
102 |
normal_form "f"
|
|
103 |
normal_form "f x"
|
|
104 |
normal_form "(f o g) x"
|
|
105 |
normal_form "(f o id) x"
|
|
106 |
normal_form "id"
|
|
107 |
normal_form "\<lambda>x. x"
|
23930
|
108 |
|
|
109 |
lemma "[] @ [] = []" by normalization
|
|
110 |
lemma "[] @ xs = xs" by normalization
|
|
111 |
normal_form "[a, b, c] @ xs = a # b # c # xs"
|
|
112 |
normal_form "map f [x,y,z::'x] = [f x, f y, f z]"
|
|
113 |
normal_form "map (%f. f True) [id, g, Not] = [True, g True, False]"
|
|
114 |
normal_form "map (%f. f True) ([id, g, Not] @ fs) = [True, g True, False] @ map (%f. f True) fs"
|
|
115 |
normal_form "rev [a, b, c] = [c, b, a]"
|
|
116 |
normal_form "rev (a#b#cs) = rev cs @ [b, a]"
|
|
117 |
normal_form "map (%F. F [a,b,c::'x]) (map map [f,g,h])"
|
|
118 |
normal_form "map (%F. F ([a,b,c] @ ds)) (map map ([f,g,h]@fs))"
|
|
119 |
normal_form "map (%F. F [Z,S Z,S(S Z)]) (map map [S,add (S Z),mul (S(S Z)),id])"
|
|
120 |
normal_form "filter (%x. x) ([True,False,x]@xs)"
|
|
121 |
normal_form "filter Not ([True,False,x]@xs)"
|
|
122 |
|
|
123 |
normal_form "[x,y,z] @ [a,b,c] = [x, y, z, a, b ,c]"
|
|
124 |
|
|
125 |
lemma "(2::int) + 3 - 1 + (- k) * 2 = 4 + - k * 2" by normalization
|
|
126 |
lemma "(-4::int) * 2 = -8" by normalization
|
|
127 |
lemma "abs ((-4::int) + 2 * 1) = 2" by normalization
|
|
128 |
lemma "(2::int) + 3 = 5" by normalization
|
|
129 |
lemma "(2::int) + 3 * (- 4) * (- 1) = 14" by normalization
|
|
130 |
lemma "(2::int) + 3 * (- 4) * 1 + 0 = -10" by normalization
|
|
131 |
lemma "(2::int) < 3" by normalization
|
|
132 |
lemma "(2::int) <= 3" by normalization
|
|
133 |
lemma "abs ((-4::int) + 2 * 1) = 2" by normalization
|
|
134 |
lemma "4 - 42 * abs (3 + (-7\<Colon>int)) = -164" by normalization
|
|
135 |
lemma "(if (0\<Colon>nat) \<le> (x\<Colon>nat) then 0\<Colon>nat else x) = 0" by normalization
|
23998
|
136 |
|
|
137 |
lemma "last [a, b, c] = c"
|
|
138 |
by normalization
|
|
139 |
lemma "last ([a, b, c] @ xs) = (if null xs then c else last xs)"
|
|
140 |
by normalization
|
|
141 |
|
|
142 |
lemma "(%((x,y),(u,v)). add (add x y) (add u v)) ((Z,Z),(Z,Z)) = Z" by normalization
|
|
143 |
lemma "split (%x y. x) (a, b) = a" by normalization
|
|
144 |
lemma "case Z of Z \<Rightarrow> True | S x \<Rightarrow> False" by normalization
|
|
145 |
lemma "(let ((x,y),(u,v)) = ((Z,Z),(Z,Z)) in add (add x y) (add u v)) = Z"
|
|
146 |
by normalization
|
|
147 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()]"
|
|
148 |
normal_form "case xs of [] \<Rightarrow> True | x#xs \<Rightarrow> False"
|
|
149 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) xs"
|
|
150 |
normal_form "let x = y::'x in [x,x]"
|
|
151 |
normal_form "Let y (%x. [x,x])"
|
|
152 |
normal_form "case n of Z \<Rightarrow> True | S x \<Rightarrow> False"
|
|
153 |
normal_form "(%(x,y). add x y) (S z,S z)"
|
|
154 |
normal_form "(%(xs, ys). xs @ ys) ([a, b, c], [d, e, f]) = [a, b, c, d, e, f]"
|
|
155 |
normal_form "map (%x. case x of None \<Rightarrow> False | Some y \<Rightarrow> True) [None, Some ()] = [False, True]"
|
23930
|
156 |
|
|
157 |
normal_form "Suc 0 \<in> set ms"
|
|
158 |
|
|
159 |
(* Church numerals: *)
|
|
160 |
|
|
161 |
normal_form "(%m n f x. m f (n f x)) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))"
|
|
162 |
normal_form "(%m n f x. m (n f) x) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))"
|
|
163 |
normal_form "(%m n. n m) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))"
|
23998
|
164 |
|
|
165 |
|
|
166 |
lemma "nat 4 = Suc (Suc (Suc (Suc 0)))" by normalization
|
|
167 |
lemma "4 = Suc (Suc (Suc (Suc 0)))" by normalization
|
23999
|
168 |
lemma "null [x] = False" by normalization |