# HG changeset patch # User haftmann # Date 1242654334 -7200 # Node ID a324d214009cdc69ef64f6d68bb1d074abe148a2 # Parent 7733125bac3c4b4ebdcac643d5f944eb9be8fc60 added Code_Index.int_of operation diff -r 7733125bac3c -r a324d214009c src/HOL/Code_Index.thy --- a/src/HOL/Code_Index.thy Mon May 18 15:45:32 2009 +0200 +++ b/src/HOL/Code_Index.thy Mon May 18 15:45:34 2009 +0200 @@ -256,18 +256,20 @@ "n mod m = snd (div_mod_index n m)" unfolding div_mod_index_def by simp -hide (open) const of_nat nat_of - -subsection {* ML interface *} +definition int_of :: "index \ int" where + "int_of = Nat.of_nat o nat_of" -ML {* -structure Index = -struct +lemma int_of_code [code]: + "int_of k = (if k = 0 then 0 + else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))" + by (auto simp add: int_of_def mod_div_equality') -fun mk k = HOLogic.mk_number @{typ index} k; +lemma (in term_syntax) term_of_index_code [code]: + "Code_Eval.term_of k = + Code_Eval.termify (number_of :: int \ int) <\> Code_Eval.term_of_num (2::index) k" + by (simp only: term_of_anything) -end; -*} +hide (open) const of_nat nat_of int_of subsection {* Code generator setup *} @@ -325,12 +327,4 @@ (OCaml "!((_ : int) < _)") (Haskell infix 4 "<") -text {* Evaluation *} - -lemma [code, code del]: - "(Code_Eval.term_of \ index \ term) = Code_Eval.term_of" .. - -code_const "Code_Eval.term_of \ index \ term" - (SML "HOLogic.mk'_number/ HOLogic.indexT/ (IntInf.fromInt/ _)") - end diff -r 7733125bac3c -r a324d214009c src/HOL/Library/Code_Integer.thy --- a/src/HOL/Library/Code_Integer.thy Mon May 18 15:45:32 2009 +0200 +++ b/src/HOL/Library/Code_Integer.thy Mon May 18 15:45:34 2009 +0200 @@ -5,7 +5,7 @@ header {* Pretty integer literals for code generation *} theory Code_Integer -imports Main +imports Main Code_Index begin text {* @@ -91,15 +91,17 @@ (OCaml "Big'_int.lt'_big'_int") (Haskell infix 4 "<") +code_const Code_Index.int_of + (SML "IntInf.fromInt") + (OCaml "Big'_int.big'_int'_of'_int") + (Haskell "toEnum") + code_reserved SML IntInf code_reserved OCaml Big_int text {* Evaluation *} -lemma [code, code del]: - "(Code_Eval.term_of \ int \ term) = Code_Eval.term_of" .. - code_const "Code_Eval.term_of \ int \ term" - (SML "HOLogic.mk'_number/ HOLogic.intT") + (Eval "HOLogic.mk'_number/ HOLogic.intT") end \ No newline at end of file