src/HOL/ex/Codegenerator_Pretty.thy
author haftmann
Mon, 21 Jan 2008 08:43:29 +0100
changeset 25929 6bfef23e26be
parent 25616 28d373f1482a
child 25933 7fc0f4065251
permissions -rw-r--r--
avoiding direct references to numeral presentation

(*  Title:      HOL/ex/Codegenerator_Pretty.thy
    ID:         $Id$
    Author:     Florian Haftmann, TU Muenchen
*)

header {* Simple examples for pretty numerals and such *}

theory Codegenerator_Pretty
imports "~~/src/HOL/Real/RealDef" Efficient_Nat
begin

definition
  foo :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> rat" where
  "foo r s t = (t + s) / t"

definition
  bar :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> bool" where
  "bar r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"

definition
  "R1 = Fract 3 7"

definition
  "R2 = Fract (-7) 5"

definition
  "R3 = Fract 11 (-9)"

definition
  "foobar = (foo R1 1 R3, bar R2 0 R3, foo R1 R3 R2)"

definition
  foo' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
  "foo' r s t = (t + s) / t"

definition
  bar' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
  "bar' r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"

definition
  "R1' = real_of_rat (Fract 3 7)"

definition
  "R2' = real_of_rat (Fract (-7) 5)"

definition
  "R3' = real_of_rat (Fract 11 (-9))"

definition
  "foobar' = (foo' R1' 1 R3', bar' R2' 0 R3', foo' R1' R3' R2')"

definition
  "(doodle :: nat) = 1705 div 42 * 42 + 1705 mod 42"

export_code foobar foobar' doodle in SML module_name Foo
  in OCaml file -
  in Haskell file -
ML {* (Foo.foobar, Foo.foobar', Foo.doodle) *}

end