src/HOL/ex/Computations.thy
author wenzelm
Sun, 21 May 2017 23:10:39 +0200
changeset 65894 54f621d5fa00
parent 64992 41e2c3617582
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
merged

(*  Title:      HOL/ex/Computations.thy
    Author:     Florian Haftmann, TU Muenchen
*)

section \<open>Simple example for computations generated by the code generator\<close>

theory Computations
  imports "../Nat" "../Fun_Def" "../Num" "../Code_Numeral"
begin

fun even :: "nat \<Rightarrow> bool"
  where "even 0 \<longleftrightarrow> True"
      | "even (Suc 0) \<longleftrightarrow> False"
      | "even (Suc (Suc n)) \<longleftrightarrow> even n"
  
fun fib :: "nat \<Rightarrow> nat"
  where "fib 0 = 0"
      | "fib (Suc 0) = Suc 0"
      | "fib (Suc (Suc n)) = fib (Suc n) + fib n"

declare [[ML_source_trace]]

ML \<open>
local 

fun int_of_nat @{code "0 :: nat"} = 0
  | int_of_nat (@{code Suc} n) = int_of_nat n + 1;

in

val comp_nat = @{computation nat
  terms: "plus :: nat \<Rightarrow>_" "times :: nat \<Rightarrow> _" fib
  datatypes: nat}
  (fn post => post o HOLogic.mk_nat o int_of_nat o the);

val comp_numeral = @{computation nat
  terms: "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat"}
  (fn post => post o HOLogic.mk_nat o int_of_nat o the);

val comp_bool = @{computation bool
  terms: HOL.conj HOL.disj HOL.implies
    HOL.iff even "less_eq :: nat \<Rightarrow> _" "less :: nat \<Rightarrow> _" "HOL.eq :: nat \<Rightarrow> _"
  datatypes: bool}
  (K the);

val comp_check = @{computation_check terms: Trueprop};

val comp_dummy = @{computation "(nat \<times> unit) option"
  datatypes: "(nat \<times> unit) option"}

end
\<close>

declare [[ML_source_trace = false]]
  
ML_val \<open>
  comp_nat @{context} @{term "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0"}
  |> Syntax.string_of_term @{context}
  |> writeln
\<close>
  
ML_val \<open>
  comp_bool @{context} @{term "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0 < fib (Suc (Suc 0))"}
\<close>

ML_val \<open>
  comp_check @{context} @{cprop "fib (Suc (Suc (Suc 0)) * Suc (Suc (Suc 0))) + Suc 0 > fib (Suc (Suc 0))"}
\<close>
  
ML_val \<open>
  comp_numeral @{context} @{term "Suc 42 + 7"}
  |> Syntax.string_of_term @{context}
  |> writeln
\<close>

end