# HG changeset patch # User haftmann # Date 1291294416 -3600 # Node ID b37dca06477ff6e280843e619aff5a5fe56a1a3f # Parent ff53be5021337485e07b64c5d4bfc40a9d2da3b4 separate term_of function for integers -- more canonical representation of negative integers diff -r ff53be502133 -r b37dca06477f src/HOL/Code_Evaluation.thy --- 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\{semiring_div} \ 'a\{semiring_div} \ term" where - "term_of_num two = (\_. dummy_term)" +definition term_of_num_semiring :: "'a\semiring_div \ 'a \ term" where + "term_of_num_semiring two = (\_. 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 <\> term_of_num two (k div two) - else termify Int.Bit1 <\> 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 <\> term_of_num_semiring two (k div two) + else termify Int.Bit1 <\> 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 \ nat) <\> 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 \ int) <\> term_of_num (2::int) k - else termify (uminus :: int \ int) <\> (termify (number_of :: int \ int) <\> term_of_num (2::int) (- k)))" + "term_of (n::nat) = termify (number_of :: int \ nat) <\> 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 \ code_numeral) <\> term_of_num (2::code_numeral) k" by (simp only: term_of_anything) +definition term_of_num_ring :: "'a\ring_div \ 'a \ term" where + "term_of_num_ring two = (\_. 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 <\> term_of_num_ring two (k div two) + else termify Int.Bit1 <\> 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 \ int) <\> 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