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