separate term_of function for integers -- more canonical representation of negative integers
--- a/src/HOL/Code_Evaluation.thy Thu Dec 02 08:34:23 2010 +0100
+++ b/src/HOL/Code_Evaluation.thy Thu Dec 02 13:53:36 2010 +0100
@@ -138,30 +138,39 @@
subsubsection {* Numeric types *}
-definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
- "term_of_num two = (\<lambda>_. dummy_term)"
+definition term_of_num_semiring :: "'a\<Colon>semiring_div \<Rightarrow> 'a \<Rightarrow> term" where
+ "term_of_num_semiring two = (\<lambda>_. dummy_term)"
-lemma (in term_syntax) term_of_num_code [code]:
- "term_of_num two k = (if k = 0 then termify Int.Pls
+lemma (in term_syntax) term_of_num_semiring_code [code]:
+ "term_of_num_semiring two k = (if k = 0 then termify Int.Pls
else (if k mod two = 0
- then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
- else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
- by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
+ then termify Int.Bit0 <\<cdot>> term_of_num_semiring two (k div two)
+ else termify Int.Bit1 <\<cdot>> term_of_num_semiring two (k div two)))"
+ by (auto simp add: term_of_anything Const_def App_def term_of_num_semiring_def Let_def)
lemma (in term_syntax) term_of_nat_code [code]:
- "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
- by (simp only: term_of_anything)
-
-lemma (in term_syntax) term_of_int_code [code]:
- "term_of (k::int) = (if k = 0 then termify (0 :: int)
- else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
- else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
+ "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num_semiring (2::nat) n"
by (simp only: term_of_anything)
lemma (in term_syntax) term_of_code_numeral_code [code]:
"term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
by (simp only: term_of_anything)
+definition term_of_num_ring :: "'a\<Colon>ring_div \<Rightarrow> 'a \<Rightarrow> term" where
+ "term_of_num_ring two = (\<lambda>_. dummy_term)"
+
+lemma (in term_syntax) term_of_num_ring_code [code]:
+ "term_of_num_ring two k = (if k = 0 then termify Int.Pls
+ else if k = -1 then termify Int.Min
+ else if k mod two = 0 then termify Int.Bit0 <\<cdot>> term_of_num_ring two (k div two)
+ else termify Int.Bit1 <\<cdot>> term_of_num_ring two (k div two))"
+ by (auto simp add: term_of_anything Const_def App_def term_of_num_ring_def Let_def)
+
+lemma (in term_syntax) term_of_int_code [code]:
+ "term_of (k::int) = (if k = 0 then termify (0 :: int)
+ else termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num_ring (2::int) k)"
+ by (simp only: term_of_anything)
+
subsubsection {* Obfuscation *}
@@ -188,6 +197,6 @@
hide_const dummy_term valapp
-hide_const (open) Const App Abs termify valtermify term_of term_of_num tracing
+hide_const (open) Const App Abs termify valtermify term_of term_of_num_semiring term_of_num_ring tracing
end