# HG changeset patch # User haftmann # Date 1192177578 -7200 # Node ID 1dbe785ed529e59031501aee3e9d3ae8400571a7 # Parent a339b33adfaff6a1ff318767814967dee352241c added diff -r a339b33adfaf -r 1dbe785ed529 src/HOL/Library/Code_Char.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Char.thy Fri Oct 12 10:26:18 2007 +0200 @@ -0,0 +1,50 @@ +(* Title: HOL/Library/Code_Char.thy + ID: $Id$ + Author: Florian Haftmann +*) + +header {* Code generation of pretty characters (and strings) *} + +theory Code_Char +imports List +begin + +code_type char + (SML "char") + (OCaml "char") + (Haskell "Char") + +setup {* +let + val charr = @{const_name Char} + val nibbles = [@{const_name Nibble0}, @{const_name Nibble1}, + @{const_name Nibble2}, @{const_name Nibble3}, + @{const_name Nibble4}, @{const_name Nibble5}, + @{const_name Nibble6}, @{const_name Nibble7}, + @{const_name Nibble8}, @{const_name Nibble9}, + @{const_name NibbleA}, @{const_name NibbleB}, + @{const_name NibbleC}, @{const_name NibbleD}, + @{const_name NibbleE}, @{const_name NibbleF}]; +in + fold (fn target => CodeTarget.add_pretty_char target charr nibbles) + ["SML", "OCaml", "Haskell"] + #> CodeTarget.add_pretty_list_string "Haskell" + @{const_name Nil} @{const_name Cons} charr nibbles +end +*} + +code_instance char :: eq + (Haskell -) + +code_reserved SML + char + +code_reserved OCaml + char + +code_const "op = \ char \ char \ bool" + (SML "!((_ : char) = _)") + (OCaml "!((_ : char) = _)") + (Haskell infixl 4 "==") + +end diff -r a339b33adfaf -r 1dbe785ed529 src/HOL/Library/Code_Char_chr.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Char_chr.thy Fri Oct 12 10:26:18 2007 +0200 @@ -0,0 +1,45 @@ +(* Title: HOL/Library/Code_Char_chr.thy + ID: $Id$ + Author: Florian Haftmann +*) + +header {* Code generation of pretty characters with character codes *} + +theory Code_Char_chr +imports Char_nat Code_Char Code_Integer +begin + +definition + "int_of_char = int o nat_of_char" + +lemma [code func]: + "nat_of_char = nat o int_of_char" + unfolding int_of_char_def by (simp add: expand_fun_eq) + +definition + "char_of_int = char_of_nat o nat" + +lemma [code func]: + "char_of_nat = char_of_int o int" + unfolding char_of_int_def by (simp add: expand_fun_eq) + +lemmas [code func del] = char.recs char.cases char.size + +lemma [code func, code inline]: + "char_rec f c = split f (nibble_pair_of_nat (nat_of_char c))" + by (cases c) (auto simp add: nibble_pair_of_nat_char) + +lemma [code func, code inline]: + "char_case f c = split f (nibble_pair_of_nat (nat_of_char c))" + by (cases c) (auto simp add: nibble_pair_of_nat_char) + +lemma [code func]: + "size (c\char) = 0" + by (cases c) auto + +code_const int_of_char and char_of_int + (SML "!(IntInf.fromInt o Char.ord)" and "!(Char.chr o IntInf.toInt)") + (OCaml "Big'_int.big'_int'_of'_int (Char.code _)" and "Char.chr (Big'_int.int'_of'_big'_int _)") + (Haskell "toInteger (fromEnum (_ :: Char))" and "!(let chr k | k < 256 = toEnum k :: Char in chr . fromInteger)") + +end \ No newline at end of file diff -r a339b33adfaf -r 1dbe785ed529 src/HOL/Library/Code_Index.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Index.thy Fri Oct 12 10:26:18 2007 +0200 @@ -0,0 +1,255 @@ +(* ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* Type of indices *} + +theory Code_Index +imports PreList +begin + +text {* + Indices are isomorphic to HOL @{typ int} but + mapped to target-language builtin integers +*} + +subsection {* Datatype of indices *} + +datatype index = index_of_int int + +lemmas [code func del] = index.recs index.cases + +fun + int_of_index :: "index \ int" +where + "int_of_index (index_of_int k) = k" +lemmas [code func del] = int_of_index.simps + +lemma index_id [simp]: + "index_of_int (int_of_index k) = k" + by (cases k) simp_all + +lemma index: + "(\k\index. PROP P k) \ (\k\int. PROP P (index_of_int k))" +proof + fix k :: int + assume "\k\index. PROP P k" + then show "PROP P (index_of_int k)" . +next + fix k :: index + assume "\k\int. PROP P (index_of_int k)" + then have "PROP P (index_of_int (int_of_index k))" . + then show "PROP P k" by simp +qed + +lemma [code func]: "size (k\index) = 0" + by (cases k) simp_all + + +subsection {* Built-in integers as datatype on numerals *} + +instance index :: number + "number_of \ index_of_int" .. + +code_datatype "number_of \ int \ index" + +lemma number_of_index_id [simp]: + "number_of (int_of_index k) = k" + unfolding number_of_index_def by simp + +lemma number_of_index_shift: + "number_of k = index_of_int (number_of k)" + by (simp add: number_of_is_id number_of_index_def) + + +subsection {* Basic arithmetic *} + +instance index :: zero + [simp]: "0 \ index_of_int 0" .. +lemmas [code func del] = zero_index_def + +instance index :: one + [simp]: "1 \ index_of_int 1" .. +lemmas [code func del] = one_index_def + +instance index :: plus + [simp]: "k + l \ index_of_int (int_of_index k + int_of_index l)" .. +lemmas [code func del] = plus_index_def +lemma plus_index_code [code func]: + "index_of_int k + index_of_int l = index_of_int (k + l)" + unfolding plus_index_def by simp + +instance index :: minus + [simp]: "- k \ index_of_int (- int_of_index k)" + [simp]: "k - l \ index_of_int (int_of_index k - int_of_index l)" .. +lemmas [code func del] = uminus_index_def minus_index_def +lemma uminus_index_code [code func]: + "- index_of_int k \ index_of_int (- k)" + unfolding uminus_index_def by simp +lemma minus_index_code [code func]: + "index_of_int k - index_of_int l = index_of_int (k - l)" + unfolding minus_index_def by simp + +instance index :: times + [simp]: "k * l \ index_of_int (int_of_index k * int_of_index l)" .. +lemmas [code func del] = times_index_def +lemma times_index_code [code func]: + "index_of_int k * index_of_int l = index_of_int (k * l)" + unfolding times_index_def by simp + +instance index :: ord + [simp]: "k \ l \ int_of_index k \ int_of_index l" + [simp]: "k < l \ int_of_index k < int_of_index l" .. +lemmas [code func del] = less_eq_index_def less_index_def +lemma less_eq_index_code [code func]: + "index_of_int k \ index_of_int l \ k \ l" + unfolding less_eq_index_def by simp +lemma less_index_code [code func]: + "index_of_int k < index_of_int l \ k < l" + unfolding less_index_def by simp + +instance index :: ring_1 + by default (auto simp add: left_distrib right_distrib) + +lemma of_nat_index: "of_nat n = index_of_int (of_nat n)" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) + then have "int_of_index (index_of_int (int n)) + = int_of_index (of_nat n)" by simp + then have "int n = int_of_index (of_nat n)" by simp + then show ?case by simp +qed + +instance index :: number_ring + by default + (simp_all add: left_distrib number_of_index_def of_int_of_nat of_nat_index) + +lemma zero_index_code [code inline, code func]: + "(0\index) = Numeral0" + by simp + +lemma one_index_code [code inline, code func]: + "(1\index) = Numeral1" + by simp + +instance index :: abs + "\k\ \ if k < 0 then -k else k" .. + +lemma index_of_int [code func]: + "index_of_int k = (if k = 0 then 0 + else if k = -1 then -1 + else let (l, m) = divAlg (k, 2) in 2 * index_of_int l + + (if m = 0 then 0 else 1))" + by (simp add: number_of_index_shift Let_def split_def divAlg_mod_div) arith + + +subsection {* Conversion to and from @{typ nat} *} + +definition + nat_of_index :: "index \ nat" +where + [code func del]: "nat_of_index = nat o int_of_index" + +definition + nat_of_index_aux :: "index \ nat \ nat" where + [code func del]: "nat_of_index_aux i n = nat_of_index i + n" + +lemma nat_of_index_aux_code [code]: + "nat_of_index_aux i n = (if i \ 0 then n else nat_of_index_aux (i - 1) (Suc n))" + by (auto simp add: nat_of_index_aux_def nat_of_index_def) + +lemma nat_of_index_code [code]: + "nat_of_index i = nat_of_index_aux i 0" + by (simp add: nat_of_index_aux_def) + +definition + index_of_nat :: "nat \ index" +where + [code func del]: "index_of_nat = index_of_int o of_nat" + +lemma index_of_nat [code func]: + "index_of_nat 0 = 0" + "index_of_nat (Suc n) = index_of_nat n + 1" + unfolding index_of_nat_def by simp_all + +lemma index_nat_id [simp]: + "nat_of_index (index_of_nat n) = n" + "index_of_nat (nat_of_index i) = (if i \ 0 then 0 else i)" + unfolding index_of_nat_def nat_of_index_def by simp_all + + +subsection {* ML interface *} + +ML {* +structure Index = +struct + +fun mk k = @{term index_of_int} $ HOLogic.mk_number @{typ index} k; + +end; +*} + + +subsection {* Code serialization *} + +code_type index + (SML "int") + (OCaml "int") + (Haskell "Integer") + +code_instance index :: eq + (Haskell -) + +setup {* + fold (fn target => CodeTarget.add_pretty_numeral target true + @{const_name number_index_inst.number_of_index} + @{const_name Numeral.B0} @{const_name Numeral.B1} + @{const_name Numeral.Pls} @{const_name Numeral.Min} + @{const_name Numeral.Bit} + ) ["SML", "OCaml", "Haskell"] +*} + +code_reserved SML int +code_reserved OCaml int + +code_const "op + \ index \ index \ index" + (SML "Int.+ ((_), (_))") + (OCaml "Pervasives.+") + (Haskell infixl 6 "+") + +code_const "uminus \ index \ index" + (SML "Int.~") + (OCaml "Pervasives.~-") + (Haskell "negate") + +code_const "op - \ index \ index \ index" + (SML "Int.- ((_), (_))") + (OCaml "Pervasives.-") + (Haskell infixl 6 "-") + +code_const "op * \ index \ index \ index" + (SML "Int.* ((_), (_))") + (OCaml "Pervasives.*") + (Haskell infixl 7 "*") + +code_const "op = \ index \ index \ bool" + (SML "!((_ : Int.int) = _)") + (OCaml "!((_ : Pervasives.int) = _)") + (Haskell infixl 4 "==") + +code_const "op \ \ index \ index \ bool" + (SML "Int.<= ((_), (_))") + (OCaml "!((_ : Pervasives.int) <= _)") + (Haskell infix 4 "<=") + +code_const "op < \ index \ index \ bool" + (SML "Int.< ((_), (_))") + (OCaml "!((_ : Pervasives.int) < _)") + (Haskell infix 4 "<") + +code_reserved SML Int +code_reserved OCaml Pervasives + +end diff -r a339b33adfaf -r 1dbe785ed529 src/HOL/Library/Code_Integer.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Integer.thy Fri Oct 12 10:26:18 2007 +0200 @@ -0,0 +1,99 @@ +(* Title: HOL/Library/Code_Integer.thy + ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* Pretty integer literals for code generation *} + +theory Code_Integer +imports IntArith Code_Index +begin + +text {* + HOL numeral expressions are mapped to integer literals + in target languages, using predefined target language + operations for abstract integer operations. +*} + +code_type int + (SML "IntInf.int") + (OCaml "Big'_int.big'_int") + (Haskell "Integer") + +code_instance int :: eq + (Haskell -) + +setup {* + fold (fn target => CodeTarget.add_pretty_numeral target true + @{const_name number_int_inst.number_of_int} + @{const_name Numeral.B0} @{const_name Numeral.B1} + @{const_name Numeral.Pls} @{const_name Numeral.Min} + @{const_name Numeral.Bit} + ) ["SML", "OCaml", "Haskell"] +*} + +code_const "Numeral.Pls" and "Numeral.Min" and "Numeral.Bit" + (SML "raise/ Fail/ \"Pls\"" + and "raise/ Fail/ \"Min\"" + and "!((_);/ (_);/ raise/ Fail/ \"Bit\")") + (OCaml "failwith/ \"Pls\"" + and "failwith/ \"Min\"" + and "!((_);/ (_);/ failwith/ \"Bit\")") + (Haskell "error/ \"Pls\"" + and "error/ \"Min\"" + and "error/ \"Bit\"") + +code_const Numeral.pred + (SML "IntInf.- ((_), 1)") + (OCaml "Big'_int.pred'_big'_int") + (Haskell "!(_/ -/ 1)") + +code_const Numeral.succ + (SML "IntInf.+ ((_), 1)") + (OCaml "Big'_int.succ'_big'_int") + (Haskell "!(_/ +/ 1)") + +code_const "op + \ int \ int \ int" + (SML "IntInf.+ ((_), (_))") + (OCaml "Big'_int.add'_big'_int") + (Haskell infixl 6 "+") + +code_const "uminus \ int \ int" + (SML "IntInf.~") + (OCaml "Big'_int.minus'_big'_int") + (Haskell "negate") + +code_const "op - \ int \ int \ int" + (SML "IntInf.- ((_), (_))") + (OCaml "Big'_int.sub'_big'_int") + (Haskell infixl 6 "-") + +code_const "op * \ int \ int \ int" + (SML "IntInf.* ((_), (_))") + (OCaml "Big'_int.mult'_big'_int") + (Haskell infixl 7 "*") + +code_const "op = \ int \ int \ bool" + (SML "!((_ : IntInf.int) = _)") + (OCaml "Big'_int.eq'_big'_int") + (Haskell infixl 4 "==") + +code_const "op \ \ int \ int \ bool" + (SML "IntInf.<= ((_), (_))") + (OCaml "Big'_int.le'_big'_int") + (Haskell infix 4 "<=") + +code_const "op < \ int \ int \ bool" + (SML "IntInf.< ((_), (_))") + (OCaml "Big'_int.lt'_big'_int") + (Haskell infix 4 "<") + +code_const index_of_int and int_of_index + (SML "IntInf.toInt" and "IntInf.fromInt") + (OCaml "Big'_int.int'_of'_big'_int" and "Big'_int.big'_int'_of'_int") + (Haskell "_" and "_") + +code_reserved SML IntInf +code_reserved OCaml Big_int + +end \ No newline at end of file diff -r a339b33adfaf -r 1dbe785ed529 src/HOL/Library/Code_Message.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Message.thy Fri Oct 12 10:26:18 2007 +0200 @@ -0,0 +1,68 @@ +(* ID: $Id$ + Author: Florian Haftmann, TU Muenchen +*) + +header {* Monolithic strings (message strings) for code generation *} + +theory Code_Message +imports List +begin + +subsection {* Datatype of messages *} + +datatype message_string = STR string + +lemmas [code func del] = message_string.recs message_string.cases + +lemma [code func]: "size (s\message_string) = 0" + by (cases s) simp_all + +subsection {* ML interface *} + +ML {* +structure Message_String = +struct + +fun mk s = @{term STR} $ HOLogic.mk_string s; + +end; +*} + + +subsection {* Code serialization *} + +code_type message_string + (SML "string") + (OCaml "string") + (Haskell "String") + +setup {* +let + val charr = @{const_name Char} + val nibbles = [@{const_name Nibble0}, @{const_name Nibble1}, + @{const_name Nibble2}, @{const_name Nibble3}, + @{const_name Nibble4}, @{const_name Nibble5}, + @{const_name Nibble6}, @{const_name Nibble7}, + @{const_name Nibble8}, @{const_name Nibble9}, + @{const_name NibbleA}, @{const_name NibbleB}, + @{const_name NibbleC}, @{const_name NibbleD}, + @{const_name NibbleE}, @{const_name NibbleF}]; +in + fold (fn target => CodeTarget.add_pretty_message target + charr nibbles @{const_name Nil} @{const_name Cons} @{const_name STR}) + ["SML", "OCaml", "Haskell"] +end +*} + +code_reserved SML string +code_reserved OCaml string + +code_instance message_string :: eq + (Haskell -) + +code_const "op = \ message_string \ message_string \ bool" + (SML "!((_ : string) = _)") + (OCaml "!((_ : string) = _)") + (Haskell infixl 4 "==") + +end